session containing computational algebra
authorhaftmann
Thu, 06 Apr 2017 21:37:13 +0200
changeset 65417 fc41a5650fb1
parent 65416 f707dbcf11e3
child 65418 c821f1f3d92d
child 65419 457e4fbed731
session containing computational algebra
NEWS
src/HOL/Algebra/Exponent.thy
src/HOL/Algebra/IntRing.thy
src/HOL/Analysis/Arcwise_Connected.thy
src/HOL/Codegenerator_Test/Candidates.thy
src/HOL/Computational_Algebra/Computational_Algebra.thy
src/HOL/Computational_Algebra/Euclidean_Algorithm.thy
src/HOL/Computational_Algebra/Factorial_Ring.thy
src/HOL/Computational_Algebra/Field_as_Ring.thy
src/HOL/Computational_Algebra/Formal_Power_Series.thy
src/HOL/Computational_Algebra/Fraction_Field.thy
src/HOL/Computational_Algebra/Fundamental_Theorem_Algebra.thy
src/HOL/Computational_Algebra/Normalized_Fraction.thy
src/HOL/Computational_Algebra/Polynomial.thy
src/HOL/Computational_Algebra/Polynomial_FPS.thy
src/HOL/Computational_Algebra/Polynomial_Factorial.thy
src/HOL/Computational_Algebra/Primes.thy
src/HOL/Isar_Examples/Fibonacci.thy
src/HOL/Library/Field_as_Ring.thy
src/HOL/Library/Formal_Power_Series.thy
src/HOL/Library/Fraction_Field.thy
src/HOL/Library/Fundamental_Theorem_Algebra.thy
src/HOL/Library/Library.thy
src/HOL/Library/Normalized_Fraction.thy
src/HOL/Library/Polynomial.thy
src/HOL/Library/Polynomial_FPS.thy
src/HOL/Library/Polynomial_Factorial.thy
src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy
src/HOL/Number_Theory/Cong.thy
src/HOL/Number_Theory/Eratosthenes.thy
src/HOL/Number_Theory/Euclidean_Algorithm.thy
src/HOL/Number_Theory/Factorial_Ring.thy
src/HOL/Number_Theory/Primes.thy
src/HOL/Proofs/Extraction/Euclid.thy
src/HOL/ROOT
src/HOL/ex/Sqrt.thy
src/HOL/ex/Sqrt_Script.thy
--- a/NEWS	Thu Apr 06 08:33:37 2017 +0200
+++ b/NEWS	Thu Apr 06 21:37:13 2017 +0200
@@ -54,6 +54,11 @@
 fps_exp/fps_ln/fps_hypergeo to avoid polluting the name space.
 INCOMPATIBILITY.
 
+* Session "Computional_Algebra" covers many previously scattered
+theories, notably Euclidean_Algorithm, Factorial_Ring, Formal_Power_Series,
+Fraction_Field, Fundamental_Theorem_Algebra, Normalized_Fraction,
+Polynomial_FPS, Polynomial, Primes.  Minor INCOMPATIBILITY.
+
 * Constant "surj" is a full input/output abbreviation (again).
 Minor INCOMPATIBILITY.
 
--- a/src/HOL/Algebra/Exponent.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Algebra/Exponent.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -6,7 +6,7 @@
 *)
 
 theory Exponent
-imports Main "~~/src/HOL/Number_Theory/Primes"
+imports Main "~~/src/HOL/Computational_Algebra/Primes"
 begin
 
 section \<open>Sylow's Theorem\<close>
--- a/src/HOL/Algebra/IntRing.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Algebra/IntRing.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -4,7 +4,7 @@
 *)
 
 theory IntRing
-imports "~~/src/HOL/Number_Theory/Primes" QuotRing Lattice Int
+imports "~~/src/HOL/Computational_Algebra/Primes" QuotRing Lattice Int
 begin
 
 section \<open>The Ring of Integers\<close>
--- a/src/HOL/Analysis/Arcwise_Connected.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Analysis/Arcwise_Connected.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -5,7 +5,7 @@
 section \<open>Arcwise-connected sets\<close>
 
 theory Arcwise_Connected
-  imports Path_Connected Ordered_Euclidean_Space "~~/src/HOL/Number_Theory/Primes"
+  imports Path_Connected Ordered_Euclidean_Space "~~/src/HOL/Computational_Algebra/Primes"
 
 begin
 
--- a/src/HOL/Codegenerator_Test/Candidates.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Codegenerator_Test/Candidates.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -8,10 +8,10 @@
   Complex_Main
   "~~/src/HOL/Library/Library"
   "~~/src/HOL/Library/Sublist_Order"
-  "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
-  "~~/src/HOL/Library/Polynomial_Factorial"
   "~~/src/HOL/Data_Structures/Tree_Map"
   "~~/src/HOL/Data_Structures/Tree_Set"
+  "~~/src/HOL/Computational_Algebra/Computational_Algebra"
+  "~~/src/HOL/Computational_Algebra/Polynomial_Factorial"
   "~~/src/HOL/Number_Theory/Eratosthenes"
   "~~/src/HOL/ex/Records"
 begin
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Computational_Algebra.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,18 @@
+
+section \<open>Pieces of computational Algebra\<close>
+
+theory Computational_Algebra
+imports
+  Euclidean_Algorithm
+  Factorial_Ring
+  Formal_Power_Series
+  Fraction_Field
+  Fundamental_Theorem_Algebra
+  Normalized_Fraction
+  Polynomial_FPS
+  Polynomial
+  Primes
+begin
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,631 @@
+(*  Title:      HOL/Number_Theory/Euclidean_Algorithm.thy
+    Author:     Manuel Eberl, TU Muenchen
+*)
+
+section \<open>Abstract euclidean algorithm in euclidean (semi)rings\<close>
+
+theory Euclidean_Algorithm
+  imports Factorial_Ring
+begin
+
+subsection \<open>Generic construction of the (simple) euclidean algorithm\<close>
+  
+context euclidean_semiring
+begin
+
+context
+begin
+
+qualified function gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  where "gcd a b = (if b = 0 then normalize a else gcd b (a mod b))"
+  by pat_completeness simp
+termination
+  by (relation "measure (euclidean_size \<circ> snd)") (simp_all add: mod_size_less)
+
+declare gcd.simps [simp del]
+
+lemma eucl_induct [case_names zero mod]:
+  assumes H1: "\<And>b. P b 0"
+  and H2: "\<And>a b. b \<noteq> 0 \<Longrightarrow> P b (a mod b) \<Longrightarrow> P a b"
+  shows "P a b"
+proof (induct a b rule: gcd.induct)
+  case (1 a b)
+  show ?case
+  proof (cases "b = 0")
+    case True then show "P a b" by simp (rule H1)
+  next
+    case False
+    then have "P b (a mod b)"
+      by (rule "1.hyps")
+    with \<open>b \<noteq> 0\<close> show "P a b"
+      by (blast intro: H2)
+  qed
+qed
+  
+qualified lemma gcd_0:
+  "gcd a 0 = normalize a"
+  by (simp add: gcd.simps [of a 0])
+  
+qualified lemma gcd_mod:
+  "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd b a"
+  by (simp add: gcd.simps [of b 0] gcd.simps [of b a])
+
+qualified definition lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  where "lcm a b = normalize (a * b) div gcd a b"
+
+qualified definition Lcm :: "'a set \<Rightarrow> 'a" \<comment>
+    \<open>Somewhat complicated definition of Lcm that has the advantage of working
+    for infinite sets as well\<close>
+  where
+  [code del]: "Lcm A = (if \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) then
+     let l = SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l =
+       (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)
+       in normalize l 
+      else 0)"
+
+qualified definition Gcd :: "'a set \<Rightarrow> 'a"
+  where [code del]: "Gcd A = Lcm {d. \<forall>a\<in>A. d dvd a}"
+
+end    
+
+lemma semiring_gcd:
+  "class.semiring_gcd one zero times gcd lcm
+    divide plus minus unit_factor normalize"
+proof
+  show "gcd a b dvd a"
+    and "gcd a b dvd b" for a b
+    by (induct a b rule: eucl_induct)
+      (simp_all add: local.gcd_0 local.gcd_mod dvd_mod_iff)
+next
+  show "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> c dvd gcd a b" for a b c
+  proof (induct a b rule: eucl_induct)
+    case (zero a) from \<open>c dvd a\<close> show ?case
+      by (rule dvd_trans) (simp add: local.gcd_0)
+  next
+    case (mod a b)
+    then show ?case
+      by (simp add: local.gcd_mod dvd_mod_iff)
+  qed
+next
+  show "normalize (gcd a b) = gcd a b" for a b
+    by (induct a b rule: eucl_induct)
+      (simp_all add: local.gcd_0 local.gcd_mod)
+next
+  show "lcm a b = normalize (a * b) div gcd a b" for a b
+    by (fact local.lcm_def)
+qed
+
+interpretation semiring_gcd one zero times gcd lcm
+  divide plus minus unit_factor normalize
+  by (fact semiring_gcd)
+  
+lemma semiring_Gcd:
+  "class.semiring_Gcd one zero times gcd lcm Gcd Lcm
+    divide plus minus unit_factor normalize"
+proof -
+  show ?thesis
+  proof
+    have "(\<forall>a\<in>A. a dvd Lcm A) \<and> (\<forall>b. (\<forall>a\<in>A. a dvd b) \<longrightarrow> Lcm A dvd b)" for A
+    proof (cases "\<exists>l. l \<noteq>  0 \<and> (\<forall>a\<in>A. a dvd l)")
+      case False
+      then have "Lcm A = 0"
+        by (auto simp add: local.Lcm_def)
+      with False show ?thesis
+        by auto
+    next
+      case True
+      then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0" "\<forall>a\<in>A. a dvd l\<^sub>0" by blast
+      define n where "n = (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
+      define l where "l = (SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
+      have "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n"
+        apply (subst n_def)
+        apply (rule LeastI [of _ "euclidean_size l\<^sub>0"])
+        apply (rule exI [of _ l\<^sub>0])
+        apply (simp add: l\<^sub>0_props)
+        done
+      from someI_ex [OF this] have "l \<noteq> 0" and "\<forall>a\<in>A. a dvd l"
+        and "euclidean_size l = n" 
+        unfolding l_def by simp_all
+      {
+        fix l' assume "\<forall>a\<in>A. a dvd l'"
+        with \<open>\<forall>a\<in>A. a dvd l\<close> have "\<forall>a\<in>A. a dvd gcd l l'"
+          by (auto intro: gcd_greatest)
+        moreover from \<open>l \<noteq> 0\<close> have "gcd l l' \<noteq> 0"
+          by simp
+        ultimately have "\<exists>b. b \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd b) \<and> 
+          euclidean_size b = euclidean_size (gcd l l')"
+          by (intro exI [of _ "gcd l l'"], auto)
+        then have "euclidean_size (gcd l l') \<ge> n"
+          by (subst n_def) (rule Least_le)
+        moreover have "euclidean_size (gcd l l') \<le> n"
+        proof -
+          have "gcd l l' dvd l"
+            by simp
+          then obtain a where "l = gcd l l' * a" ..
+          with \<open>l \<noteq> 0\<close> have "a \<noteq> 0"
+            by auto
+          hence "euclidean_size (gcd l l') \<le> euclidean_size (gcd l l' * a)"
+            by (rule size_mult_mono)
+          also have "gcd l l' * a = l" using \<open>l = gcd l l' * a\<close> ..
+          also note \<open>euclidean_size l = n\<close>
+          finally show "euclidean_size (gcd l l') \<le> n" .
+        qed
+        ultimately have *: "euclidean_size l = euclidean_size (gcd l l')" 
+          by (intro le_antisym, simp_all add: \<open>euclidean_size l = n\<close>)
+        from \<open>l \<noteq> 0\<close> have "l dvd gcd l l'"
+          by (rule dvd_euclidean_size_eq_imp_dvd) (auto simp add: *)
+        hence "l dvd l'" by (rule dvd_trans [OF _ gcd_dvd2])
+      }
+      with \<open>\<forall>a\<in>A. a dvd l\<close> and \<open>l \<noteq> 0\<close>
+        have "(\<forall>a\<in>A. a dvd normalize l) \<and> 
+          (\<forall>l'. (\<forall>a\<in>A. a dvd l') \<longrightarrow> normalize l dvd l')"
+        by auto
+      also from True have "normalize l = Lcm A"
+        by (simp add: local.Lcm_def Let_def n_def l_def)
+      finally show ?thesis .
+    qed
+    then show dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A"
+      and Lcm_least: "(\<And>a. a \<in> A \<Longrightarrow> a dvd b) \<Longrightarrow> Lcm A dvd b" for A and a b
+      by auto
+    show "a \<in> A \<Longrightarrow> Gcd A dvd a" for A and a
+      by (auto simp add: local.Gcd_def intro: Lcm_least)
+    show "(\<And>a. a \<in> A \<Longrightarrow> b dvd a) \<Longrightarrow> b dvd Gcd A" for A and b
+      by (auto simp add: local.Gcd_def intro: dvd_Lcm)
+    show [simp]: "normalize (Lcm A) = Lcm A" for A
+      by (simp add: local.Lcm_def)
+    show "normalize (Gcd A) = Gcd A" for A
+      by (simp add: local.Gcd_def)
+  qed
+qed
+
+interpretation semiring_Gcd one zero times gcd lcm Gcd Lcm
+    divide plus minus unit_factor normalize
+  by (fact semiring_Gcd)
+
+subclass factorial_semiring
+proof -
+  show "class.factorial_semiring divide plus minus zero times one
+     unit_factor normalize"
+  proof (standard, rule factorial_semiring_altI_aux) \<comment> \<open>FIXME rule\<close>
+    fix x assume "x \<noteq> 0"
+    thus "finite {p. p dvd x \<and> normalize p = p}"
+    proof (induction "euclidean_size x" arbitrary: x rule: less_induct)
+      case (less x)
+      show ?case
+      proof (cases "\<exists>y. y dvd x \<and> \<not>x dvd y \<and> \<not>is_unit y")
+        case False
+        have "{p. p dvd x \<and> normalize p = p} \<subseteq> {1, normalize x}"
+        proof
+          fix p assume p: "p \<in> {p. p dvd x \<and> normalize p = p}"
+          with False have "is_unit p \<or> x dvd p" by blast
+          thus "p \<in> {1, normalize x}"
+          proof (elim disjE)
+            assume "is_unit p"
+            hence "normalize p = 1" by (simp add: is_unit_normalize)
+            with p show ?thesis by simp
+          next
+            assume "x dvd p"
+            with p have "normalize p = normalize x" by (intro associatedI) simp_all
+            with p show ?thesis by simp
+          qed
+        qed
+        moreover have "finite \<dots>" by simp
+        ultimately show ?thesis by (rule finite_subset)
+      next
+        case True
+        then obtain y where y: "y dvd x" "\<not>x dvd y" "\<not>is_unit y" by blast
+        define z where "z = x div y"
+        let ?fctrs = "\<lambda>x. {p. p dvd x \<and> normalize p = p}"
+        from y have x: "x = y * z" by (simp add: z_def)
+        with less.prems have "y \<noteq> 0" "z \<noteq> 0" by auto
+        have normalized_factors_product:
+          "{p. p dvd a * b \<and> normalize p = p} = 
+             (\<lambda>(x,y). x * y) ` ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})" for a b
+        proof safe
+          fix p assume p: "p dvd a * b" "normalize p = p"
+          from dvd_productE[OF p(1)] guess x y . note xy = this
+          define x' y' where "x' = normalize x" and "y' = normalize y"
+          have "p = x' * y'"
+            by (subst p(2) [symmetric]) (simp add: xy x'_def y'_def normalize_mult)
+          moreover from xy have "normalize x' = x'" "normalize y' = y'" "x' dvd a" "y' dvd b" 
+            by (simp_all add: x'_def y'_def)
+          ultimately show "p \<in> (\<lambda>(x, y). x * y) ` 
+            ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})"
+            by blast
+        qed (auto simp: normalize_mult mult_dvd_mono)
+        from x y have "\<not>is_unit z" by (auto simp: mult_unit_dvd_iff)
+        have "?fctrs x = (\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z)"
+          by (subst x) (rule normalized_factors_product)
+        also have "\<not>y * z dvd y * 1" "\<not>y * z dvd 1 * z"
+          by (subst dvd_times_left_cancel_iff dvd_times_right_cancel_iff; fact)+
+        hence "finite ((\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z))"
+          by (intro finite_imageI finite_cartesian_product less dvd_proper_imp_size_less)
+             (auto simp: x)
+        finally show ?thesis .
+      qed
+    qed
+  next
+    fix p
+    assume "irreducible p"
+    then show "prime_elem p"
+      by (rule irreducible_imp_prime_elem_gcd)
+  qed
+qed
+
+lemma Gcd_eucl_set [code]:
+  "Gcd (set xs) = fold gcd xs 0"
+  by (fact Gcd_set_eq_fold)
+
+lemma Lcm_eucl_set [code]:
+  "Lcm (set xs) = fold lcm xs 1"
+  by (fact Lcm_set_eq_fold)
+ 
+end
+
+hide_const (open) gcd lcm Gcd Lcm
+
+lemma prime_elem_int_abs_iff [simp]:
+  fixes p :: int
+  shows "prime_elem \<bar>p\<bar> \<longleftrightarrow> prime_elem p"
+  using prime_elem_normalize_iff [of p] by simp
+  
+lemma prime_elem_int_minus_iff [simp]:
+  fixes p :: int
+  shows "prime_elem (- p) \<longleftrightarrow> prime_elem p"
+  using prime_elem_normalize_iff [of "- p"] by simp
+
+lemma prime_int_iff:
+  fixes p :: int
+  shows "prime p \<longleftrightarrow> p > 0 \<and> prime_elem p"
+  by (auto simp add: prime_def dest: prime_elem_not_zeroI)
+  
+  
+subsection \<open>The (simple) euclidean algorithm as gcd computation\<close>
+  
+class euclidean_semiring_gcd = euclidean_semiring + gcd + Gcd +
+  assumes gcd_eucl: "Euclidean_Algorithm.gcd = GCD.gcd"
+    and lcm_eucl: "Euclidean_Algorithm.lcm = GCD.lcm"
+  assumes Gcd_eucl: "Euclidean_Algorithm.Gcd = GCD.Gcd"
+    and Lcm_eucl: "Euclidean_Algorithm.Lcm = GCD.Lcm"
+begin
+
+subclass semiring_gcd
+  unfolding gcd_eucl [symmetric] lcm_eucl [symmetric]
+  by (fact semiring_gcd)
+
+subclass semiring_Gcd
+  unfolding  gcd_eucl [symmetric] lcm_eucl [symmetric]
+    Gcd_eucl [symmetric] Lcm_eucl [symmetric]
+  by (fact semiring_Gcd)
+
+subclass factorial_semiring_gcd
+proof
+  show "gcd a b = gcd_factorial a b" for a b
+    apply (rule sym)
+    apply (rule gcdI)
+       apply (fact gcd_lcm_factorial)+
+    done
+  then show "lcm a b = lcm_factorial a b" for a b
+    by (simp add: lcm_factorial_gcd_factorial lcm_gcd)
+  show "Gcd A = Gcd_factorial A" for A
+    apply (rule sym)
+    apply (rule GcdI)
+       apply (fact gcd_lcm_factorial)+
+    done
+  show "Lcm A = Lcm_factorial A" for A
+    apply (rule sym)
+    apply (rule LcmI)
+       apply (fact gcd_lcm_factorial)+
+    done
+qed
+
+lemma gcd_mod_right [simp]:
+  "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd a b"
+  unfolding gcd.commute [of a b]
+  by (simp add: gcd_eucl [symmetric] local.gcd_mod)
+
+lemma gcd_mod_left [simp]:
+  "b \<noteq> 0 \<Longrightarrow> gcd (a mod b) b = gcd a b"
+  by (drule gcd_mod_right [of _ a]) (simp add: gcd.commute)
+
+lemma euclidean_size_gcd_le1 [simp]:
+  assumes "a \<noteq> 0"
+  shows "euclidean_size (gcd a b) \<le> euclidean_size a"
+proof -
+  from gcd_dvd1 obtain c where A: "a = gcd a b * c" ..
+  with assms have "c \<noteq> 0"
+    by auto
+  moreover from this
+  have "euclidean_size (gcd a b) \<le> euclidean_size (gcd a b * c)"
+    by (rule size_mult_mono)
+  with A show ?thesis
+    by simp
+qed
+
+lemma euclidean_size_gcd_le2 [simp]:
+  "b \<noteq> 0 \<Longrightarrow> euclidean_size (gcd a b) \<le> euclidean_size b"
+  by (subst gcd.commute, rule euclidean_size_gcd_le1)
+
+lemma euclidean_size_gcd_less1:
+  assumes "a \<noteq> 0" and "\<not> a dvd b"
+  shows "euclidean_size (gcd a b) < euclidean_size a"
+proof (rule ccontr)
+  assume "\<not>euclidean_size (gcd a b) < euclidean_size a"
+  with \<open>a \<noteq> 0\<close> have A: "euclidean_size (gcd a b) = euclidean_size a"
+    by (intro le_antisym, simp_all)
+  have "a dvd gcd a b"
+    by (rule dvd_euclidean_size_eq_imp_dvd) (simp_all add: assms A)
+  hence "a dvd b" using dvd_gcdD2 by blast
+  with \<open>\<not> a dvd b\<close> show False by contradiction
+qed
+
+lemma euclidean_size_gcd_less2:
+  assumes "b \<noteq> 0" and "\<not> b dvd a"
+  shows "euclidean_size (gcd a b) < euclidean_size b"
+  using assms by (subst gcd.commute, rule euclidean_size_gcd_less1)
+
+lemma euclidean_size_lcm_le1: 
+  assumes "a \<noteq> 0" and "b \<noteq> 0"
+  shows "euclidean_size a \<le> euclidean_size (lcm a b)"
+proof -
+  have "a dvd lcm a b" by (rule dvd_lcm1)
+  then obtain c where A: "lcm a b = a * c" ..
+  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "c \<noteq> 0" by (auto simp: lcm_eq_0_iff)
+  then show ?thesis by (subst A, intro size_mult_mono)
+qed
+
+lemma euclidean_size_lcm_le2:
+  "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> euclidean_size b \<le> euclidean_size (lcm a b)"
+  using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps)
+
+lemma euclidean_size_lcm_less1:
+  assumes "b \<noteq> 0" and "\<not> b dvd a"
+  shows "euclidean_size a < euclidean_size (lcm a b)"
+proof (rule ccontr)
+  from assms have "a \<noteq> 0" by auto
+  assume "\<not>euclidean_size a < euclidean_size (lcm a b)"
+  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "euclidean_size (lcm a b) = euclidean_size a"
+    by (intro le_antisym, simp, intro euclidean_size_lcm_le1)
+  with assms have "lcm a b dvd a" 
+    by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_eq_0_iff)
+  hence "b dvd a" by (rule lcm_dvdD2)
+  with \<open>\<not>b dvd a\<close> show False by contradiction
+qed
+
+lemma euclidean_size_lcm_less2:
+  assumes "a \<noteq> 0" and "\<not> a dvd b"
+  shows "euclidean_size b < euclidean_size (lcm a b)"
+  using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps)
+
+end
+
+lemma factorial_euclidean_semiring_gcdI:
+  "OFCLASS('a::{factorial_semiring_gcd, euclidean_semiring}, euclidean_semiring_gcd_class)"
+proof
+  interpret semiring_Gcd 1 0 times
+    Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm
+    Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm
+    divide plus minus unit_factor normalize
+    rewrites "dvd.dvd op * = Rings.dvd"
+    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
+  show [simp]: "Euclidean_Algorithm.gcd = (gcd :: 'a \<Rightarrow> _)"
+  proof (rule ext)+
+    fix a b :: 'a
+    show "Euclidean_Algorithm.gcd a b = gcd a b"
+    proof (induct a b rule: eucl_induct)
+      case zero
+      then show ?case
+        by simp
+    next
+      case (mod a b)
+      moreover have "gcd b (a mod b) = gcd b a"
+        using GCD.gcd_add_mult [of b "a div b" "a mod b", symmetric]
+          by (simp add: div_mult_mod_eq)
+      ultimately show ?case
+        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
+    qed
+  qed
+  show [simp]: "Euclidean_Algorithm.Lcm = (Lcm :: 'a set \<Rightarrow> _)"
+    by (auto intro!: Lcm_eqI GCD.dvd_Lcm GCD.Lcm_least)
+  show "Euclidean_Algorithm.lcm = (lcm :: 'a \<Rightarrow> _)"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
+  show "Euclidean_Algorithm.Gcd = (Gcd :: 'a set \<Rightarrow> _)"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
+qed
+
+
+subsection \<open>The extended euclidean algorithm\<close>
+  
+class euclidean_ring_gcd = euclidean_semiring_gcd + idom
+begin
+
+subclass euclidean_ring ..
+subclass ring_gcd ..
+subclass factorial_ring_gcd ..
+
+function euclid_ext_aux :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
+  where "euclid_ext_aux s' s t' t r' r = (
+     if r = 0 then let c = 1 div unit_factor r' in ((s' * c, t' * c), normalize r')
+     else let q = r' div r
+          in euclid_ext_aux s (s' - q * s) t (t' - q * t) r (r' mod r))"
+  by auto
+termination
+  by (relation "measure (\<lambda>(_, _, _, _, _, b). euclidean_size b)")
+    (simp_all add: mod_size_less)
+
+abbreviation (input) euclid_ext :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
+  where "euclid_ext \<equiv> euclid_ext_aux 1 0 0 1"
+    
+lemma
+  assumes "gcd r' r = gcd a b"
+  assumes "s' * a + t' * b = r'"
+  assumes "s * a + t * b = r"
+  assumes "euclid_ext_aux s' s t' t r' r = ((x, y), c)"
+  shows euclid_ext_aux_eq_gcd: "c = gcd a b"
+    and euclid_ext_aux_bezout: "x * a + y * b = gcd a b"
+proof -
+  have "case euclid_ext_aux s' s t' t r' r of ((x, y), c) \<Rightarrow> 
+    x * a + y * b = c \<and> c = gcd a b" (is "?P (euclid_ext_aux s' s t' t r' r)")
+    using assms(1-3)
+  proof (induction s' s t' t r' r rule: euclid_ext_aux.induct)
+    case (1 s' s t' t r' r)
+    show ?case
+    proof (cases "r = 0")
+      case True
+      hence "euclid_ext_aux s' s t' t r' r = 
+               ((s' div unit_factor r', t' div unit_factor r'), normalize r')"
+        by (subst euclid_ext_aux.simps) (simp add: Let_def)
+      also have "?P \<dots>"
+      proof safe
+        have "s' div unit_factor r' * a + t' div unit_factor r' * b = 
+                (s' * a + t' * b) div unit_factor r'"
+          by (cases "r' = 0") (simp_all add: unit_div_commute)
+        also have "s' * a + t' * b = r'" by fact
+        also have "\<dots> div unit_factor r' = normalize r'" by simp
+        finally show "s' div unit_factor r' * a + t' div unit_factor r' * b = normalize r'" .
+      next
+        from "1.prems" True show "normalize r' = gcd a b"
+          by simp
+      qed
+      finally show ?thesis .
+    next
+      case False
+      hence "euclid_ext_aux s' s t' t r' r = 
+             euclid_ext_aux s (s' - r' div r * s) t (t' - r' div r * t) r (r' mod r)"
+        by (subst euclid_ext_aux.simps) (simp add: Let_def)
+      also from "1.prems" False have "?P \<dots>"
+      proof (intro "1.IH")
+        have "(s' - r' div r * s) * a + (t' - r' div r * t) * b =
+              (s' * a + t' * b) - r' div r * (s * a + t * b)" by (simp add: algebra_simps)
+        also have "s' * a + t' * b = r'" by fact
+        also have "s * a + t * b = r" by fact
+        also have "r' - r' div r * r = r' mod r" using div_mult_mod_eq [of r' r]
+          by (simp add: algebra_simps)
+        finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" .
+      qed (auto simp: gcd_mod_right algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute)
+      finally show ?thesis .
+    qed
+  qed
+  with assms(4) show "c = gcd a b" "x * a + y * b = gcd a b"
+    by simp_all
+qed
+
+declare euclid_ext_aux.simps [simp del]
+
+definition bezout_coefficients :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<times> 'a"
+  where [code]: "bezout_coefficients a b = fst (euclid_ext a b)"
+
+lemma bezout_coefficients_0: 
+  "bezout_coefficients a 0 = (1 div unit_factor a, 0)"
+  by (simp add: bezout_coefficients_def euclid_ext_aux.simps)
+
+lemma bezout_coefficients_left_0: 
+  "bezout_coefficients 0 a = (0, 1 div unit_factor a)"
+  by (simp add: bezout_coefficients_def euclid_ext_aux.simps)
+
+lemma bezout_coefficients:
+  assumes "bezout_coefficients a b = (x, y)"
+  shows "x * a + y * b = gcd a b"
+  using assms by (simp add: bezout_coefficients_def
+    euclid_ext_aux_bezout [of a b a b 1 0 0 1 x y] prod_eq_iff)
+
+lemma bezout_coefficients_fst_snd:
+  "fst (bezout_coefficients a b) * a + snd (bezout_coefficients a b) * b = gcd a b"
+  by (rule bezout_coefficients) simp
+
+lemma euclid_ext_eq [simp]:
+  "euclid_ext a b = (bezout_coefficients a b, gcd a b)" (is "?p = ?q")
+proof
+  show "fst ?p = fst ?q"
+    by (simp add: bezout_coefficients_def)
+  have "snd (euclid_ext_aux 1 0 0 1 a b) = gcd a b"
+    by (rule euclid_ext_aux_eq_gcd [of a b a b 1 0 0 1])
+      (simp_all add: prod_eq_iff)
+  then show "snd ?p = snd ?q"
+    by simp
+qed
+
+declare euclid_ext_eq [symmetric, code_unfold]
+
+end
+
+
+subsection \<open>Typical instances\<close>
+
+instance nat :: euclidean_semiring_gcd
+proof
+  interpret semiring_Gcd 1 0 times
+    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
+    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
+    divide plus minus unit_factor normalize
+    rewrites "dvd.dvd op * = Rings.dvd"
+    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
+  show [simp]: "(Euclidean_Algorithm.gcd :: nat \<Rightarrow> _) = gcd"
+  proof (rule ext)+
+    fix m n :: nat
+    show "Euclidean_Algorithm.gcd m n = gcd m n"
+    proof (induct m n rule: eucl_induct)
+      case zero
+      then show ?case
+        by simp
+    next
+      case (mod m n)
+      then have "gcd n (m mod n) = gcd n m"
+        using gcd_nat.simps [of m n] by (simp add: ac_simps)
+      with mod show ?case
+        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
+    qed
+  qed
+  show [simp]: "(Euclidean_Algorithm.Lcm :: nat set \<Rightarrow> _) = Lcm"
+    by (auto intro!: ext Lcm_eqI)
+  show "(Euclidean_Algorithm.lcm :: nat \<Rightarrow> _) = lcm"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
+  show "(Euclidean_Algorithm.Gcd :: nat set \<Rightarrow> _) = Gcd"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
+qed
+
+instance int :: euclidean_ring_gcd
+proof
+  interpret semiring_Gcd 1 0 times
+    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
+    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
+    divide plus minus unit_factor normalize
+    rewrites "dvd.dvd op * = Rings.dvd"
+    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
+  show [simp]: "(Euclidean_Algorithm.gcd :: int \<Rightarrow> _) = gcd"
+  proof (rule ext)+
+    fix k l :: int
+    show "Euclidean_Algorithm.gcd k l = gcd k l"
+    proof (induct k l rule: eucl_induct)
+      case zero
+      then show ?case
+        by simp
+    next
+      case (mod k l)
+      have "gcd l (k mod l) = gcd l k"
+      proof (cases l "0::int" rule: linorder_cases)
+        case less
+        then show ?thesis
+          using gcd_non_0_int [of "- l" "- k"] by (simp add: ac_simps)
+      next
+        case equal
+        with mod show ?thesis
+          by simp
+      next
+        case greater
+        then show ?thesis
+          using gcd_non_0_int [of l k] by (simp add: ac_simps)
+      qed
+      with mod show ?case
+        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
+    qed
+  qed
+  show [simp]: "(Euclidean_Algorithm.Lcm :: int set \<Rightarrow> _) = Lcm"
+    by (auto intro!: ext Lcm_eqI)
+  show "(Euclidean_Algorithm.lcm :: int \<Rightarrow> _) = lcm"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
+  show "(Euclidean_Algorithm.Gcd :: int set \<Rightarrow> _) = Gcd"
+    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Factorial_Ring.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,1818 @@
+(*  Title:      HOL/Number_Theory/Factorial_Ring.thy
+    Author:     Manuel Eberl, TU Muenchen
+    Author:     Florian Haftmann, TU Muenchen
+*)
+
+section \<open>Factorial (semi)rings\<close>
+
+theory Factorial_Ring
+imports 
+  Main
+  "../GCD"
+  "~~/src/HOL/Library/Multiset"
+begin
+
+subsection \<open>Irreducible and prime elements\<close>
+
+context comm_semiring_1
+begin
+
+definition irreducible :: "'a \<Rightarrow> bool" where
+  "irreducible p \<longleftrightarrow> p \<noteq> 0 \<and> \<not>p dvd 1 \<and> (\<forall>a b. p = a * b \<longrightarrow> a dvd 1 \<or> b dvd 1)"
+
+lemma not_irreducible_zero [simp]: "\<not>irreducible 0"
+  by (simp add: irreducible_def)
+
+lemma irreducible_not_unit: "irreducible p \<Longrightarrow> \<not>p dvd 1"
+  by (simp add: irreducible_def)
+
+lemma not_irreducible_one [simp]: "\<not>irreducible 1"
+  by (simp add: irreducible_def)
+
+lemma irreducibleI:
+  "p \<noteq> 0 \<Longrightarrow> \<not>p dvd 1 \<Longrightarrow> (\<And>a b. p = a * b \<Longrightarrow> a dvd 1 \<or> b dvd 1) \<Longrightarrow> irreducible p"
+  by (simp add: irreducible_def)
+
+lemma irreducibleD: "irreducible p \<Longrightarrow> p = a * b \<Longrightarrow> a dvd 1 \<or> b dvd 1"
+  by (simp add: irreducible_def)
+
+definition prime_elem :: "'a \<Rightarrow> bool" where
+  "prime_elem p \<longleftrightarrow> p \<noteq> 0 \<and> \<not>p dvd 1 \<and> (\<forall>a b. p dvd (a * b) \<longrightarrow> p dvd a \<or> p dvd b)"
+
+lemma not_prime_elem_zero [simp]: "\<not>prime_elem 0"
+  by (simp add: prime_elem_def)
+
+lemma prime_elem_not_unit: "prime_elem p \<Longrightarrow> \<not>p dvd 1"
+  by (simp add: prime_elem_def)
+
+lemma prime_elemI:
+    "p \<noteq> 0 \<Longrightarrow> \<not>p dvd 1 \<Longrightarrow> (\<And>a b. p dvd (a * b) \<Longrightarrow> p dvd a \<or> p dvd b) \<Longrightarrow> prime_elem p"
+  by (simp add: prime_elem_def)
+
+lemma prime_elem_dvd_multD:
+    "prime_elem p \<Longrightarrow> p dvd (a * b) \<Longrightarrow> p dvd a \<or> p dvd b"
+  by (simp add: prime_elem_def)
+
+lemma prime_elem_dvd_mult_iff:
+  "prime_elem p \<Longrightarrow> p dvd (a * b) \<longleftrightarrow> p dvd a \<or> p dvd b"
+  by (auto simp: prime_elem_def)
+
+lemma not_prime_elem_one [simp]:
+  "\<not> prime_elem 1"
+  by (auto dest: prime_elem_not_unit)
+
+lemma prime_elem_not_zeroI:
+  assumes "prime_elem p"
+  shows "p \<noteq> 0"
+  using assms by (auto intro: ccontr)
+
+lemma prime_elem_dvd_power: 
+  "prime_elem p \<Longrightarrow> p dvd x ^ n \<Longrightarrow> p dvd x"
+  by (induction n) (auto dest: prime_elem_dvd_multD intro: dvd_trans[of _ 1])
+
+lemma prime_elem_dvd_power_iff:
+  "prime_elem p \<Longrightarrow> n > 0 \<Longrightarrow> p dvd x ^ n \<longleftrightarrow> p dvd x"
+  by (auto dest: prime_elem_dvd_power intro: dvd_trans)
+
+lemma prime_elem_imp_nonzero [simp]:
+  "ASSUMPTION (prime_elem x) \<Longrightarrow> x \<noteq> 0"
+  unfolding ASSUMPTION_def by (rule prime_elem_not_zeroI)
+
+lemma prime_elem_imp_not_one [simp]:
+  "ASSUMPTION (prime_elem x) \<Longrightarrow> x \<noteq> 1"
+  unfolding ASSUMPTION_def by auto
+
+end
+
+context algebraic_semidom
+begin
+
+lemma prime_elem_imp_irreducible:
+  assumes "prime_elem p"
+  shows   "irreducible p"
+proof (rule irreducibleI)
+  fix a b
+  assume p_eq: "p = a * b"
+  with assms have nz: "a \<noteq> 0" "b \<noteq> 0" by auto
+  from p_eq have "p dvd a * b" by simp
+  with \<open>prime_elem p\<close> have "p dvd a \<or> p dvd b" by (rule prime_elem_dvd_multD)
+  with \<open>p = a * b\<close> have "a * b dvd 1 * b \<or> a * b dvd a * 1" by auto
+  thus "a dvd 1 \<or> b dvd 1"
+    by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)])
+qed (insert assms, simp_all add: prime_elem_def)
+
+lemma (in algebraic_semidom) unit_imp_no_irreducible_divisors:
+  assumes "is_unit x" "irreducible p"
+  shows   "\<not>p dvd x"
+proof (rule notI)
+  assume "p dvd x"
+  with \<open>is_unit x\<close> have "is_unit p"
+    by (auto intro: dvd_trans)
+  with \<open>irreducible p\<close> show False
+    by (simp add: irreducible_not_unit)
+qed
+   
+lemma unit_imp_no_prime_divisors:
+  assumes "is_unit x" "prime_elem p"
+  shows   "\<not>p dvd x"
+  using unit_imp_no_irreducible_divisors[OF assms(1) prime_elem_imp_irreducible[OF assms(2)]] .
+
+lemma prime_elem_mono:
+  assumes "prime_elem p" "\<not>q dvd 1" "q dvd p"
+  shows   "prime_elem q"
+proof -
+  from \<open>q dvd p\<close> obtain r where r: "p = q * r" by (elim dvdE)
+  hence "p dvd q * r" by simp
+  with \<open>prime_elem p\<close> have "p dvd q \<or> p dvd r" by (rule prime_elem_dvd_multD)
+  hence "p dvd q"
+  proof
+    assume "p dvd r"
+    then obtain s where s: "r = p * s" by (elim dvdE)
+    from r have "p * 1 = p * (q * s)" by (subst (asm) s) (simp add: mult_ac)
+    with \<open>prime_elem p\<close> have "q dvd 1"
+      by (subst (asm) mult_cancel_left) auto
+    with \<open>\<not>q dvd 1\<close> show ?thesis by contradiction
+  qed
+
+  show ?thesis
+  proof (rule prime_elemI)
+    fix a b assume "q dvd (a * b)"
+    with \<open>p dvd q\<close> have "p dvd (a * b)" by (rule dvd_trans)
+    with \<open>prime_elem p\<close> have "p dvd a \<or> p dvd b" by (rule prime_elem_dvd_multD)
+    with \<open>q dvd p\<close> show "q dvd a \<or> q dvd b" by (blast intro: dvd_trans)
+  qed (insert assms, auto)
+qed
+
+lemma irreducibleD':
+  assumes "irreducible a" "b dvd a"
+  shows   "a dvd b \<or> is_unit b"
+proof -
+  from assms obtain c where c: "a = b * c" by (elim dvdE)
+  from irreducibleD[OF assms(1) this] have "is_unit b \<or> is_unit c" .
+  thus ?thesis by (auto simp: c mult_unit_dvd_iff)
+qed
+
+lemma irreducibleI':
+  assumes "a \<noteq> 0" "\<not>is_unit a" "\<And>b. b dvd a \<Longrightarrow> a dvd b \<or> is_unit b"
+  shows   "irreducible a"
+proof (rule irreducibleI)
+  fix b c assume a_eq: "a = b * c"
+  hence "a dvd b \<or> is_unit b" by (intro assms) simp_all
+  thus "is_unit b \<or> is_unit c"
+  proof
+    assume "a dvd b"
+    hence "b * c dvd b * 1" by (simp add: a_eq)
+    moreover from \<open>a \<noteq> 0\<close> a_eq have "b \<noteq> 0" by auto
+    ultimately show ?thesis by (subst (asm) dvd_times_left_cancel_iff) auto
+  qed blast
+qed (simp_all add: assms(1,2))
+
+lemma irreducible_altdef:
+  "irreducible x \<longleftrightarrow> x \<noteq> 0 \<and> \<not>is_unit x \<and> (\<forall>b. b dvd x \<longrightarrow> x dvd b \<or> is_unit b)"
+  using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto
+
+lemma prime_elem_multD:
+  assumes "prime_elem (a * b)"
+  shows "is_unit a \<or> is_unit b"
+proof -
+  from assms have "a \<noteq> 0" "b \<noteq> 0" by (auto dest!: prime_elem_not_zeroI)
+  moreover from assms prime_elem_dvd_multD [of "a * b"] have "a * b dvd a \<or> a * b dvd b"
+    by auto
+  ultimately show ?thesis
+    using dvd_times_left_cancel_iff [of a b 1]
+      dvd_times_right_cancel_iff [of b a 1]
+    by auto
+qed
+
+lemma prime_elemD2:
+  assumes "prime_elem p" and "a dvd p" and "\<not> is_unit a"
+  shows "p dvd a"
+proof -
+  from \<open>a dvd p\<close> obtain b where "p = a * b" ..
+  with \<open>prime_elem p\<close> prime_elem_multD \<open>\<not> is_unit a\<close> have "is_unit b" by auto
+  with \<open>p = a * b\<close> show ?thesis
+    by (auto simp add: mult_unit_dvd_iff)
+qed
+
+lemma prime_elem_dvd_prod_msetE:
+  assumes "prime_elem p"
+  assumes dvd: "p dvd prod_mset A"
+  obtains a where "a \<in># A" and "p dvd a"
+proof -
+  from dvd have "\<exists>a. a \<in># A \<and> p dvd a"
+  proof (induct A)
+    case empty then show ?case
+    using \<open>prime_elem p\<close> by (simp add: prime_elem_not_unit)
+  next
+    case (add a A)
+    then have "p dvd a * prod_mset A" by simp
+    with \<open>prime_elem p\<close> consider (A) "p dvd prod_mset A" | (B) "p dvd a"
+      by (blast dest: prime_elem_dvd_multD)
+    then show ?case proof cases
+      case B then show ?thesis by auto
+    next
+      case A
+      with add.hyps obtain b where "b \<in># A" "p dvd b"
+        by auto
+      then show ?thesis by auto
+    qed
+  qed
+  with that show thesis by blast
+qed
+
+context
+begin
+
+private lemma prime_elem_powerD:
+  assumes "prime_elem (p ^ n)"
+  shows   "prime_elem p \<and> n = 1"
+proof (cases n)
+  case (Suc m)
+  note assms
+  also from Suc have "p ^ n = p * p^m" by simp
+  finally have "is_unit p \<or> is_unit (p^m)" by (rule prime_elem_multD)
+  moreover from assms have "\<not>is_unit p" by (simp add: prime_elem_def is_unit_power_iff)
+  ultimately have "is_unit (p ^ m)" by simp
+  with \<open>\<not>is_unit p\<close> have "m = 0" by (simp add: is_unit_power_iff)
+  with Suc assms show ?thesis by simp
+qed (insert assms, simp_all)
+
+lemma prime_elem_power_iff:
+  "prime_elem (p ^ n) \<longleftrightarrow> prime_elem p \<and> n = 1"
+  by (auto dest: prime_elem_powerD)
+
+end
+
+lemma irreducible_mult_unit_left:
+  "is_unit a \<Longrightarrow> irreducible (a * p) \<longleftrightarrow> irreducible p"
+  by (auto simp: irreducible_altdef mult.commute[of a] is_unit_mult_iff
+        mult_unit_dvd_iff dvd_mult_unit_iff)
+
+lemma prime_elem_mult_unit_left:
+  "is_unit a \<Longrightarrow> prime_elem (a * p) \<longleftrightarrow> prime_elem p"
+  by (auto simp: prime_elem_def mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff)
+
+lemma prime_elem_dvd_cases:
+  assumes pk: "p*k dvd m*n" and p: "prime_elem p"
+  shows "(\<exists>x. k dvd x*n \<and> m = p*x) \<or> (\<exists>y. k dvd m*y \<and> n = p*y)"
+proof -
+  have "p dvd m*n" using dvd_mult_left pk by blast
+  then consider "p dvd m" | "p dvd n"
+    using p prime_elem_dvd_mult_iff by blast
+  then show ?thesis
+  proof cases
+    case 1 then obtain a where "m = p * a" by (metis dvd_mult_div_cancel) 
+      then have "\<exists>x. k dvd x * n \<and> m = p * x"
+        using p pk by (auto simp: mult.assoc)
+    then show ?thesis ..
+  next
+    case 2 then obtain b where "n = p * b" by (metis dvd_mult_div_cancel) 
+    with p pk have "\<exists>y. k dvd m*y \<and> n = p*y" 
+      by (metis dvd_mult_right dvd_times_left_cancel_iff mult.left_commute mult_zero_left)
+    then show ?thesis ..
+  qed
+qed
+
+lemma prime_elem_power_dvd_prod:
+  assumes pc: "p^c dvd m*n" and p: "prime_elem p"
+  shows "\<exists>a b. a+b = c \<and> p^a dvd m \<and> p^b dvd n"
+using pc
+proof (induct c arbitrary: m n)
+  case 0 show ?case by simp
+next
+  case (Suc c)
+  consider x where "p^c dvd x*n" "m = p*x" | y where "p^c dvd m*y" "n = p*y"
+    using prime_elem_dvd_cases [of _ "p^c", OF _ p] Suc.prems by force
+  then show ?case
+  proof cases
+    case (1 x) 
+    with Suc.hyps[of x n] obtain a b where "a + b = c \<and> p ^ a dvd x \<and> p ^ b dvd n" by blast
+    with 1 have "Suc a + b = Suc c \<and> p ^ Suc a dvd m \<and> p ^ b dvd n"
+      by (auto intro: mult_dvd_mono)
+    thus ?thesis by blast
+  next
+    case (2 y) 
+    with Suc.hyps[of m y] obtain a b where "a + b = c \<and> p ^ a dvd m \<and> p ^ b dvd y" by blast
+    with 2 have "a + Suc b = Suc c \<and> p ^ a dvd m \<and> p ^ Suc b dvd n"
+      by (auto intro: mult_dvd_mono)
+    with Suc.hyps [of m y] show "\<exists>a b. a + b = Suc c \<and> p ^ a dvd m \<and> p ^ b dvd n"
+      by force
+  qed
+qed
+
+lemma prime_elem_power_dvd_cases:
+  assumes "p ^ c dvd m * n" and "a + b = Suc c" and "prime_elem p"
+  shows "p ^ a dvd m \<or> p ^ b dvd n"
+proof -
+  from assms obtain r s
+    where "r + s = c \<and> p ^ r dvd m \<and> p ^ s dvd n"
+    by (blast dest: prime_elem_power_dvd_prod)
+  moreover with assms have
+    "a \<le> r \<or> b \<le> s" by arith
+  ultimately show ?thesis by (auto intro: power_le_dvd)
+qed
+
+lemma prime_elem_not_unit' [simp]:
+  "ASSUMPTION (prime_elem x) \<Longrightarrow> \<not>is_unit x"
+  unfolding ASSUMPTION_def by (rule prime_elem_not_unit)
+
+lemma prime_elem_dvd_power_iff:
+  assumes "prime_elem p"
+  shows "p dvd a ^ n \<longleftrightarrow> p dvd a \<and> n > 0"
+  using assms by (induct n) (auto dest: prime_elem_not_unit prime_elem_dvd_multD)
+
+lemma prime_power_dvd_multD:
+  assumes "prime_elem p"
+  assumes "p ^ n dvd a * b" and "n > 0" and "\<not> p dvd a"
+  shows "p ^ n dvd b"
+  using \<open>p ^ n dvd a * b\<close> and \<open>n > 0\<close> 
+proof (induct n arbitrary: b)
+  case 0 then show ?case by simp
+next
+  case (Suc n) show ?case
+  proof (cases "n = 0")
+    case True with Suc \<open>prime_elem p\<close> \<open>\<not> p dvd a\<close> show ?thesis
+      by (simp add: prime_elem_dvd_mult_iff)
+  next
+    case False then have "n > 0" by simp
+    from \<open>prime_elem p\<close> have "p \<noteq> 0" by auto
+    from Suc.prems have *: "p * p ^ n dvd a * b"
+      by simp
+    then have "p dvd a * b"
+      by (rule dvd_mult_left)
+    with Suc \<open>prime_elem p\<close> \<open>\<not> p dvd a\<close> have "p dvd b"
+      by (simp add: prime_elem_dvd_mult_iff)
+    moreover define c where "c = b div p"
+    ultimately have b: "b = p * c" by simp
+    with * have "p * p ^ n dvd p * (a * c)"
+      by (simp add: ac_simps)
+    with \<open>p \<noteq> 0\<close> have "p ^ n dvd a * c"
+      by simp
+    with Suc.hyps \<open>n > 0\<close> have "p ^ n dvd c"
+      by blast
+    with \<open>p \<noteq> 0\<close> show ?thesis
+      by (simp add: b)
+  qed
+qed
+
+end
+
+
+subsection \<open>Generalized primes: normalized prime elements\<close>
+
+context normalization_semidom
+begin
+
+lemma irreducible_normalized_divisors:
+  assumes "irreducible x" "y dvd x" "normalize y = y"
+  shows   "y = 1 \<or> y = normalize x"
+proof -
+  from assms have "is_unit y \<or> x dvd y" by (auto simp: irreducible_altdef)
+  thus ?thesis
+  proof (elim disjE)
+    assume "is_unit y"
+    hence "normalize y = 1" by (simp add: is_unit_normalize)
+    with assms show ?thesis by simp
+  next
+    assume "x dvd y"
+    with \<open>y dvd x\<close> have "normalize y = normalize x" by (rule associatedI)
+    with assms show ?thesis by simp
+  qed
+qed
+
+lemma irreducible_normalize_iff [simp]: "irreducible (normalize x) = irreducible x"
+  using irreducible_mult_unit_left[of "1 div unit_factor x" x]
+  by (cases "x = 0") (simp_all add: unit_div_commute)
+
+lemma prime_elem_normalize_iff [simp]: "prime_elem (normalize x) = prime_elem x"
+  using prime_elem_mult_unit_left[of "1 div unit_factor x" x]
+  by (cases "x = 0") (simp_all add: unit_div_commute)
+
+lemma prime_elem_associated:
+  assumes "prime_elem p" and "prime_elem q" and "q dvd p"
+  shows "normalize q = normalize p"
+using \<open>q dvd p\<close> proof (rule associatedI)
+  from \<open>prime_elem q\<close> have "\<not> is_unit q"
+    by (auto simp add: prime_elem_not_unit)
+  with \<open>prime_elem p\<close> \<open>q dvd p\<close> show "p dvd q"
+    by (blast intro: prime_elemD2)
+qed
+
+definition prime :: "'a \<Rightarrow> bool" where
+  "prime p \<longleftrightarrow> prime_elem p \<and> normalize p = p"
+
+lemma not_prime_0 [simp]: "\<not>prime 0" by (simp add: prime_def)
+
+lemma not_prime_unit: "is_unit x \<Longrightarrow> \<not>prime x"
+  using prime_elem_not_unit[of x] by (auto simp add: prime_def)
+
+lemma not_prime_1 [simp]: "\<not>prime 1" by (simp add: not_prime_unit)
+
+lemma primeI: "prime_elem x \<Longrightarrow> normalize x = x \<Longrightarrow> prime x"
+  by (simp add: prime_def)
+
+lemma prime_imp_prime_elem [dest]: "prime p \<Longrightarrow> prime_elem p"
+  by (simp add: prime_def)
+
+lemma normalize_prime: "prime p \<Longrightarrow> normalize p = p"
+  by (simp add: prime_def)
+
+lemma prime_normalize_iff [simp]: "prime (normalize p) \<longleftrightarrow> prime_elem p"
+  by (auto simp add: prime_def)
+
+lemma prime_power_iff:
+  "prime (p ^ n) \<longleftrightarrow> prime p \<and> n = 1"
+  by (auto simp: prime_def prime_elem_power_iff)
+
+lemma prime_imp_nonzero [simp]:
+  "ASSUMPTION (prime x) \<Longrightarrow> x \<noteq> 0"
+  unfolding ASSUMPTION_def prime_def by auto
+
+lemma prime_imp_not_one [simp]:
+  "ASSUMPTION (prime x) \<Longrightarrow> x \<noteq> 1"
+  unfolding ASSUMPTION_def by auto
+
+lemma prime_not_unit' [simp]:
+  "ASSUMPTION (prime x) \<Longrightarrow> \<not>is_unit x"
+  unfolding ASSUMPTION_def prime_def by auto
+
+lemma prime_normalize' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> normalize x = x"
+  unfolding ASSUMPTION_def prime_def by simp
+
+lemma unit_factor_prime: "prime x \<Longrightarrow> unit_factor x = 1"
+  using unit_factor_normalize[of x] unfolding prime_def by auto
+
+lemma unit_factor_prime' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> unit_factor x = 1"
+  unfolding ASSUMPTION_def by (rule unit_factor_prime)
+
+lemma prime_imp_prime_elem' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> prime_elem x"
+  by (simp add: prime_def ASSUMPTION_def)
+
+lemma prime_dvd_multD: "prime p \<Longrightarrow> p dvd a * b \<Longrightarrow> p dvd a \<or> p dvd b"
+  by (intro prime_elem_dvd_multD) simp_all
+
+lemma prime_dvd_mult_iff: "prime p \<Longrightarrow> p dvd a * b \<longleftrightarrow> p dvd a \<or> p dvd b"
+  by (auto dest: prime_dvd_multD)
+
+lemma prime_dvd_power: 
+  "prime p \<Longrightarrow> p dvd x ^ n \<Longrightarrow> p dvd x"
+  by (auto dest!: prime_elem_dvd_power simp: prime_def)
+
+lemma prime_dvd_power_iff:
+  "prime p \<Longrightarrow> n > 0 \<Longrightarrow> p dvd x ^ n \<longleftrightarrow> p dvd x"
+  by (subst prime_elem_dvd_power_iff) simp_all
+
+lemma prime_dvd_prod_mset_iff: "prime p \<Longrightarrow> p dvd prod_mset A \<longleftrightarrow> (\<exists>x. x \<in># A \<and> p dvd x)"
+  by (induction A) (simp_all add: prime_elem_dvd_mult_iff prime_imp_prime_elem, blast+)
+
+lemma primes_dvd_imp_eq:
+  assumes "prime p" "prime q" "p dvd q"
+  shows   "p = q"
+proof -
+  from assms have "irreducible q" by (simp add: prime_elem_imp_irreducible prime_def)
+  from irreducibleD'[OF this \<open>p dvd q\<close>] assms have "q dvd p" by simp
+  with \<open>p dvd q\<close> have "normalize p = normalize q" by (rule associatedI)
+  with assms show "p = q" by simp
+qed
+
+lemma prime_dvd_prod_mset_primes_iff:
+  assumes "prime p" "\<And>q. q \<in># A \<Longrightarrow> prime q"
+  shows   "p dvd prod_mset A \<longleftrightarrow> p \<in># A"
+proof -
+  from assms(1) have "p dvd prod_mset A \<longleftrightarrow> (\<exists>x. x \<in># A \<and> p dvd x)" by (rule prime_dvd_prod_mset_iff)
+  also from assms have "\<dots> \<longleftrightarrow> p \<in># A" by (auto dest: primes_dvd_imp_eq)
+  finally show ?thesis .
+qed
+
+lemma prod_mset_primes_dvd_imp_subset:
+  assumes "prod_mset A dvd prod_mset B" "\<And>p. p \<in># A \<Longrightarrow> prime p" "\<And>p. p \<in># B \<Longrightarrow> prime p"
+  shows   "A \<subseteq># B"
+using assms
+proof (induction A arbitrary: B)
+  case empty
+  thus ?case by simp
+next
+  case (add p A B)
+  hence p: "prime p" by simp
+  define B' where "B' = B - {#p#}"
+  from add.prems have "p dvd prod_mset B" by (simp add: dvd_mult_left)
+  with add.prems have "p \<in># B"
+    by (subst (asm) (2) prime_dvd_prod_mset_primes_iff) simp_all
+  hence B: "B = B' + {#p#}" by (simp add: B'_def)
+  from add.prems p have "A \<subseteq># B'" by (intro add.IH) (simp_all add: B)
+  thus ?case by (simp add: B)
+qed
+
+lemma normalize_prod_mset_primes:
+  "(\<And>p. p \<in># A \<Longrightarrow> prime p) \<Longrightarrow> normalize (prod_mset A) = prod_mset A"
+proof (induction A)
+  case (add p A)
+  hence "prime p" by simp
+  hence "normalize p = p" by simp
+  with add show ?case by (simp add: normalize_mult)
+qed simp_all
+
+lemma prod_mset_dvd_prod_mset_primes_iff:
+  assumes "\<And>x. x \<in># A \<Longrightarrow> prime x" "\<And>x. x \<in># B \<Longrightarrow> prime x"
+  shows   "prod_mset A dvd prod_mset B \<longleftrightarrow> A \<subseteq># B"
+  using assms by (auto intro: prod_mset_subset_imp_dvd prod_mset_primes_dvd_imp_subset)
+
+lemma is_unit_prod_mset_primes_iff:
+  assumes "\<And>x. x \<in># A \<Longrightarrow> prime x"
+  shows   "is_unit (prod_mset A) \<longleftrightarrow> A = {#}"
+  by (auto simp add: is_unit_prod_mset_iff)
+    (meson all_not_in_conv assms not_prime_unit set_mset_eq_empty_iff)
+
+lemma prod_mset_primes_irreducible_imp_prime:
+  assumes irred: "irreducible (prod_mset A)"
+  assumes A: "\<And>x. x \<in># A \<Longrightarrow> prime x"
+  assumes B: "\<And>x. x \<in># B \<Longrightarrow> prime x"
+  assumes C: "\<And>x. x \<in># C \<Longrightarrow> prime x"
+  assumes dvd: "prod_mset A dvd prod_mset B * prod_mset C"
+  shows   "prod_mset A dvd prod_mset B \<or> prod_mset A dvd prod_mset C"
+proof -
+  from dvd have "prod_mset A dvd prod_mset (B + C)"
+    by simp
+  with A B C have subset: "A \<subseteq># B + C"
+    by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto
+  define A1 and A2 where "A1 = A \<inter># B" and "A2 = A - A1"
+  have "A = A1 + A2" unfolding A1_def A2_def
+    by (rule sym, intro subset_mset.add_diff_inverse) simp_all
+  from subset have "A1 \<subseteq># B" "A2 \<subseteq># C"
+    by (auto simp: A1_def A2_def Multiset.subset_eq_diff_conv Multiset.union_commute)
+  from \<open>A = A1 + A2\<close> have "prod_mset A = prod_mset A1 * prod_mset A2" by simp
+  from irred and this have "is_unit (prod_mset A1) \<or> is_unit (prod_mset A2)"
+    by (rule irreducibleD)
+  with A have "A1 = {#} \<or> A2 = {#}" unfolding A1_def A2_def
+    by (subst (asm) (1 2) is_unit_prod_mset_primes_iff) (auto dest: Multiset.in_diffD)
+  with dvd \<open>A = A1 + A2\<close> \<open>A1 \<subseteq># B\<close> \<open>A2 \<subseteq># C\<close> show ?thesis
+    by (auto intro: prod_mset_subset_imp_dvd)
+qed
+
+lemma prod_mset_primes_finite_divisor_powers:
+  assumes A: "\<And>x. x \<in># A \<Longrightarrow> prime x"
+  assumes B: "\<And>x. x \<in># B \<Longrightarrow> prime x"
+  assumes "A \<noteq> {#}"
+  shows   "finite {n. prod_mset A ^ n dvd prod_mset B}"
+proof -
+  from \<open>A \<noteq> {#}\<close> obtain x where x: "x \<in># A" by blast
+  define m where "m = count B x"
+  have "{n. prod_mset A ^ n dvd prod_mset B} \<subseteq> {..m}"
+  proof safe
+    fix n assume dvd: "prod_mset A ^ n dvd prod_mset B"
+    from x have "x ^ n dvd prod_mset A ^ n" by (intro dvd_power_same dvd_prod_mset)
+    also note dvd
+    also have "x ^ n = prod_mset (replicate_mset n x)" by simp
+    finally have "replicate_mset n x \<subseteq># B"
+      by (rule prod_mset_primes_dvd_imp_subset) (insert A B x, simp_all split: if_splits)
+    thus "n \<le> m" by (simp add: count_le_replicate_mset_subset_eq m_def)
+  qed
+  moreover have "finite {..m}" by simp
+  ultimately show ?thesis by (rule finite_subset)
+qed
+
+end
+
+
+subsection \<open>In a semiring with GCD, each irreducible element is a prime elements\<close>
+
+context semiring_gcd
+begin
+
+lemma irreducible_imp_prime_elem_gcd:
+  assumes "irreducible x"
+  shows   "prime_elem x"
+proof (rule prime_elemI)
+  fix a b assume "x dvd a * b"
+  from dvd_productE[OF this] obtain y z where yz: "x = y * z" "y dvd a" "z dvd b" .
+  from \<open>irreducible x\<close> and \<open>x = y * z\<close> have "is_unit y \<or> is_unit z" by (rule irreducibleD)
+  with yz show "x dvd a \<or> x dvd b"
+    by (auto simp: mult_unit_dvd_iff mult_unit_dvd_iff')
+qed (insert assms, auto simp: irreducible_not_unit)
+
+lemma prime_elem_imp_coprime:
+  assumes "prime_elem p" "\<not>p dvd n"
+  shows   "coprime p n"
+proof (rule coprimeI)
+  fix d assume "d dvd p" "d dvd n"
+  show "is_unit d"
+  proof (rule ccontr)
+    assume "\<not>is_unit d"
+    from \<open>prime_elem p\<close> and \<open>d dvd p\<close> and this have "p dvd d"
+      by (rule prime_elemD2)
+    from this and \<open>d dvd n\<close> have "p dvd n" by (rule dvd_trans)
+    with \<open>\<not>p dvd n\<close> show False by contradiction
+  qed
+qed
+
+lemma prime_imp_coprime:
+  assumes "prime p" "\<not>p dvd n"
+  shows   "coprime p n"
+  using assms by (simp add: prime_elem_imp_coprime)
+
+lemma prime_elem_imp_power_coprime: 
+  "prime_elem p \<Longrightarrow> \<not>p dvd a \<Longrightarrow> coprime a (p ^ m)"
+  by (auto intro!: coprime_exp dest: prime_elem_imp_coprime simp: gcd.commute)
+
+lemma prime_imp_power_coprime: 
+  "prime p \<Longrightarrow> \<not>p dvd a \<Longrightarrow> coprime a (p ^ m)"
+  by (simp add: prime_elem_imp_power_coprime)
+
+lemma prime_elem_divprod_pow:
+  assumes p: "prime_elem p" and ab: "coprime a b" and pab: "p^n dvd a * b"
+  shows   "p^n dvd a \<or> p^n dvd b"
+  using assms
+proof -
+  from ab p have "\<not>p dvd a \<or> \<not>p dvd b"
+    by (auto simp: coprime prime_elem_def)
+  with p have "coprime (p^n) a \<or> coprime (p^n) b" 
+    by (auto intro: prime_elem_imp_coprime coprime_exp_left)
+  with pab show ?thesis by (auto intro: coprime_dvd_mult simp: mult_ac)
+qed
+
+lemma primes_coprime: 
+  "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+  using prime_imp_coprime primes_dvd_imp_eq by blast
+
+end
+
+
+subsection \<open>Factorial semirings: algebraic structures with unique prime factorizations\<close>
+
+class factorial_semiring = normalization_semidom +
+  assumes prime_factorization_exists:
+    "x \<noteq> 0 \<Longrightarrow> \<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize x"
+
+text \<open>Alternative characterization\<close>
+  
+lemma (in normalization_semidom) factorial_semiring_altI_aux:
+  assumes finite_divisors: "\<And>x. x \<noteq> 0 \<Longrightarrow> finite {y. y dvd x \<and> normalize y = y}"
+  assumes irreducible_imp_prime_elem: "\<And>x. irreducible x \<Longrightarrow> prime_elem x"
+  assumes "x \<noteq> 0"
+  shows   "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize x"
+using \<open>x \<noteq> 0\<close>
+proof (induction "card {b. b dvd x \<and> normalize b = b}" arbitrary: x rule: less_induct)
+  case (less a)
+  let ?fctrs = "\<lambda>a. {b. b dvd a \<and> normalize b = b}"
+  show ?case
+  proof (cases "is_unit a")
+    case True
+    thus ?thesis by (intro exI[of _ "{#}"]) (auto simp: is_unit_normalize)
+  next
+    case False
+    show ?thesis
+    proof (cases "\<exists>b. b dvd a \<and> \<not>is_unit b \<and> \<not>a dvd b")
+      case False
+      with \<open>\<not>is_unit a\<close> less.prems have "irreducible a" by (auto simp: irreducible_altdef)
+      hence "prime_elem a" by (rule irreducible_imp_prime_elem)
+      thus ?thesis by (intro exI[of _ "{#normalize a#}"]) auto
+    next
+      case True
+      then guess b by (elim exE conjE) note b = this
+
+      from b have "?fctrs b \<subseteq> ?fctrs a" by (auto intro: dvd_trans)
+      moreover from b have "normalize a \<notin> ?fctrs b" "normalize a \<in> ?fctrs a" by simp_all
+      hence "?fctrs b \<noteq> ?fctrs a" by blast
+      ultimately have "?fctrs b \<subset> ?fctrs a" by (subst subset_not_subset_eq) blast
+      with finite_divisors[OF \<open>a \<noteq> 0\<close>] have "card (?fctrs b) < card (?fctrs a)"
+        by (rule psubset_card_mono)
+      moreover from \<open>a \<noteq> 0\<close> b have "b \<noteq> 0" by auto
+      ultimately have "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize b"
+        by (intro less) auto
+      then guess A .. note A = this
+
+      define c where "c = a div b"
+      from b have c: "a = b * c" by (simp add: c_def)
+      from less.prems c have "c \<noteq> 0" by auto
+      from b c have "?fctrs c \<subseteq> ?fctrs a" by (auto intro: dvd_trans)
+      moreover have "normalize a \<notin> ?fctrs c"
+      proof safe
+        assume "normalize a dvd c"
+        hence "b * c dvd 1 * c" by (simp add: c)
+        hence "b dvd 1" by (subst (asm) dvd_times_right_cancel_iff) fact+
+        with b show False by simp
+      qed
+      with \<open>normalize a \<in> ?fctrs a\<close> have "?fctrs a \<noteq> ?fctrs c" by blast
+      ultimately have "?fctrs c \<subset> ?fctrs a" by (subst subset_not_subset_eq) blast
+      with finite_divisors[OF \<open>a \<noteq> 0\<close>] have "card (?fctrs c) < card (?fctrs a)"
+        by (rule psubset_card_mono)
+      with \<open>c \<noteq> 0\<close> have "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize c"
+        by (intro less) auto
+      then guess B .. note B = this
+
+      from A B show ?thesis by (intro exI[of _ "A + B"]) (auto simp: c normalize_mult)
+    qed
+  qed
+qed 
+
+lemma factorial_semiring_altI:
+  assumes finite_divisors: "\<And>x::'a. x \<noteq> 0 \<Longrightarrow> finite {y. y dvd x \<and> normalize y = y}"
+  assumes irreducible_imp_prime: "\<And>x::'a. irreducible x \<Longrightarrow> prime_elem x"
+  shows   "OFCLASS('a :: normalization_semidom, factorial_semiring_class)"
+  by intro_classes (rule factorial_semiring_altI_aux[OF assms])
+  
+text \<open>Properties\<close>
+
+context factorial_semiring
+begin
+
+lemma prime_factorization_exists':
+  assumes "x \<noteq> 0"
+  obtains A where "\<And>x. x \<in># A \<Longrightarrow> prime x" "prod_mset A = normalize x"
+proof -
+  from prime_factorization_exists[OF assms] obtain A
+    where A: "\<And>x. x \<in># A \<Longrightarrow> prime_elem x" "prod_mset A = normalize x" by blast
+  define A' where "A' = image_mset normalize A"
+  have "prod_mset A' = normalize (prod_mset A)"
+    by (simp add: A'_def normalize_prod_mset)
+  also note A(2)
+  finally have "prod_mset A' = normalize x" by simp
+  moreover from A(1) have "\<forall>x. x \<in># A' \<longrightarrow> prime x" by (auto simp: prime_def A'_def)
+  ultimately show ?thesis by (intro that[of A']) blast
+qed
+
+lemma irreducible_imp_prime_elem:
+  assumes "irreducible x"
+  shows   "prime_elem x"
+proof (rule prime_elemI)
+  fix a b assume dvd: "x dvd a * b"
+  from assms have "x \<noteq> 0" by auto
+  show "x dvd a \<or> x dvd b"
+  proof (cases "a = 0 \<or> b = 0")
+    case False
+    hence "a \<noteq> 0" "b \<noteq> 0" by blast+
+    note nz = \<open>x \<noteq> 0\<close> this
+    from nz[THEN prime_factorization_exists'] guess A B C . note ABC = this
+    from assms ABC have "irreducible (prod_mset A)" by simp
+    from dvd prod_mset_primes_irreducible_imp_prime[of A B C, OF this ABC(1,3,5)] ABC(2,4,6)
+      show ?thesis by (simp add: normalize_mult [symmetric])
+  qed auto
+qed (insert assms, simp_all add: irreducible_def)
+
+lemma finite_divisor_powers:
+  assumes "y \<noteq> 0" "\<not>is_unit x"
+  shows   "finite {n. x ^ n dvd y}"
+proof (cases "x = 0")
+  case True
+  with assms have "{n. x ^ n dvd y} = {0}" by (auto simp: power_0_left)
+  thus ?thesis by simp
+next
+  case False
+  note nz = this \<open>y \<noteq> 0\<close>
+  from nz[THEN prime_factorization_exists'] guess A B . note AB = this
+  from AB assms have "A \<noteq> {#}" by (auto simp: normalize_1_iff)
+  from AB(2,4) prod_mset_primes_finite_divisor_powers [of A B, OF AB(1,3) this]
+    show ?thesis by (simp add: normalize_power [symmetric])
+qed
+
+lemma finite_prime_divisors:
+  assumes "x \<noteq> 0"
+  shows   "finite {p. prime p \<and> p dvd x}"
+proof -
+  from prime_factorization_exists'[OF assms] guess A . note A = this
+  have "{p. prime p \<and> p dvd x} \<subseteq> set_mset A"
+  proof safe
+    fix p assume p: "prime p" and dvd: "p dvd x"
+    from dvd have "p dvd normalize x" by simp
+    also from A have "normalize x = prod_mset A" by simp
+    finally show "p \<in># A" using p A by (subst (asm) prime_dvd_prod_mset_primes_iff)
+  qed
+  moreover have "finite (set_mset A)" by simp
+  ultimately show ?thesis by (rule finite_subset)
+qed
+
+lemma prime_elem_iff_irreducible: "prime_elem x \<longleftrightarrow> irreducible x"
+  by (blast intro: irreducible_imp_prime_elem prime_elem_imp_irreducible)
+
+lemma prime_divisor_exists:
+  assumes "a \<noteq> 0" "\<not>is_unit a"
+  shows   "\<exists>b. b dvd a \<and> prime b"
+proof -
+  from prime_factorization_exists'[OF assms(1)] guess A . note A = this
+  moreover from A and assms have "A \<noteq> {#}" by auto
+  then obtain x where "x \<in># A" by blast
+  with A(1) have *: "x dvd prod_mset A" "prime x" by (auto simp: dvd_prod_mset)
+  with A have "x dvd a" by simp
+  with * show ?thesis by blast
+qed
+
+lemma prime_divisors_induct [case_names zero unit factor]:
+  assumes "P 0" "\<And>x. is_unit x \<Longrightarrow> P x" "\<And>p x. prime p \<Longrightarrow> P x \<Longrightarrow> P (p * x)"
+  shows   "P x"
+proof (cases "x = 0")
+  case False
+  from prime_factorization_exists'[OF this] guess A . note A = this
+  from A(1) have "P (unit_factor x * prod_mset A)"
+  proof (induction A)
+    case (add p A)
+    from add.prems have "prime p" by simp
+    moreover from add.prems have "P (unit_factor x * prod_mset A)" by (intro add.IH) simp_all
+    ultimately have "P (p * (unit_factor x * prod_mset A))" by (rule assms(3))
+    thus ?case by (simp add: mult_ac)
+  qed (simp_all add: assms False)
+  with A show ?thesis by simp
+qed (simp_all add: assms(1))
+
+lemma no_prime_divisors_imp_unit:
+  assumes "a \<noteq> 0" "\<And>b. b dvd a \<Longrightarrow> normalize b = b \<Longrightarrow> \<not> prime_elem b"
+  shows "is_unit a"
+proof (rule ccontr)
+  assume "\<not>is_unit a"
+  from prime_divisor_exists[OF assms(1) this] guess b by (elim exE conjE)
+  with assms(2)[of b] show False by (simp add: prime_def)
+qed
+
+lemma prime_divisorE:
+  assumes "a \<noteq> 0" and "\<not> is_unit a"
+  obtains p where "prime p" and "p dvd a"
+  using assms no_prime_divisors_imp_unit unfolding prime_def by blast
+
+definition multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
+  "multiplicity p x = (if finite {n. p ^ n dvd x} then Max {n. p ^ n dvd x} else 0)"
+
+lemma multiplicity_dvd: "p ^ multiplicity p x dvd x"
+proof (cases "finite {n. p ^ n dvd x}")
+  case True
+  hence "multiplicity p x = Max {n. p ^ n dvd x}"
+    by (simp add: multiplicity_def)
+  also have "\<dots> \<in> {n. p ^ n dvd x}"
+    by (rule Max_in) (auto intro!: True exI[of _ "0::nat"])
+  finally show ?thesis by simp
+qed (simp add: multiplicity_def)
+
+lemma multiplicity_dvd': "n \<le> multiplicity p x \<Longrightarrow> p ^ n dvd x"
+  by (rule dvd_trans[OF le_imp_power_dvd multiplicity_dvd])
+
+context
+  fixes x p :: 'a
+  assumes xp: "x \<noteq> 0" "\<not>is_unit p"
+begin
+
+lemma multiplicity_eq_Max: "multiplicity p x = Max {n. p ^ n dvd x}"
+  using finite_divisor_powers[OF xp] by (simp add: multiplicity_def)
+
+lemma multiplicity_geI:
+  assumes "p ^ n dvd x"
+  shows   "multiplicity p x \<ge> n"
+proof -
+  from assms have "n \<le> Max {n. p ^ n dvd x}"
+    by (intro Max_ge finite_divisor_powers xp) simp_all
+  thus ?thesis by (subst multiplicity_eq_Max)
+qed
+
+lemma multiplicity_lessI:
+  assumes "\<not>p ^ n dvd x"
+  shows   "multiplicity p x < n"
+proof (rule ccontr)
+  assume "\<not>(n > multiplicity p x)"
+  hence "p ^ n dvd x" by (intro multiplicity_dvd') simp
+  with assms show False by contradiction
+qed
+
+lemma power_dvd_iff_le_multiplicity:
+  "p ^ n dvd x \<longleftrightarrow> n \<le> multiplicity p x"
+  using multiplicity_geI[of n] multiplicity_lessI[of n] by (cases "p ^ n dvd x") auto
+
+lemma multiplicity_eq_zero_iff:
+  shows   "multiplicity p x = 0 \<longleftrightarrow> \<not>p dvd x"
+  using power_dvd_iff_le_multiplicity[of 1] by auto
+
+lemma multiplicity_gt_zero_iff:
+  shows   "multiplicity p x > 0 \<longleftrightarrow> p dvd x"
+  using power_dvd_iff_le_multiplicity[of 1] by auto
+
+lemma multiplicity_decompose:
+  "\<not>p dvd (x div p ^ multiplicity p x)"
+proof
+  assume *: "p dvd x div p ^ multiplicity p x"
+  have "x = x div p ^ multiplicity p x * (p ^ multiplicity p x)"
+    using multiplicity_dvd[of p x] by simp
+  also from * have "x div p ^ multiplicity p x = (x div p ^ multiplicity p x div p) * p" by simp
+  also have "x div p ^ multiplicity p x div p * p * p ^ multiplicity p x =
+               x div p ^ multiplicity p x div p * p ^ Suc (multiplicity p x)"
+    by (simp add: mult_assoc)
+  also have "p ^ Suc (multiplicity p x) dvd \<dots>" by (rule dvd_triv_right)
+  finally show False by (subst (asm) power_dvd_iff_le_multiplicity) simp
+qed
+
+lemma multiplicity_decompose':
+  obtains y where "x = p ^ multiplicity p x * y" "\<not>p dvd y"
+  using that[of "x div p ^ multiplicity p x"]
+  by (simp add: multiplicity_decompose multiplicity_dvd)
+
+end
+
+lemma multiplicity_zero [simp]: "multiplicity p 0 = 0"
+  by (simp add: multiplicity_def)
+
+lemma prime_elem_multiplicity_eq_zero_iff:
+  "prime_elem p \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> multiplicity p x = 0 \<longleftrightarrow> \<not>p dvd x"
+  by (rule multiplicity_eq_zero_iff) simp_all
+
+lemma prime_multiplicity_other:
+  assumes "prime p" "prime q" "p \<noteq> q"
+  shows   "multiplicity p q = 0"
+  using assms by (subst prime_elem_multiplicity_eq_zero_iff) (auto dest: primes_dvd_imp_eq)  
+
+lemma prime_multiplicity_gt_zero_iff:
+  "prime_elem p \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> multiplicity p x > 0 \<longleftrightarrow> p dvd x"
+  by (rule multiplicity_gt_zero_iff) simp_all
+
+lemma multiplicity_unit_left: "is_unit p \<Longrightarrow> multiplicity p x = 0"
+  by (simp add: multiplicity_def is_unit_power_iff unit_imp_dvd)
+
+lemma multiplicity_unit_right:
+  assumes "is_unit x"
+  shows   "multiplicity p x = 0"
+proof (cases "is_unit p \<or> x = 0")
+  case False
+  with multiplicity_lessI[of x p 1] this assms
+    show ?thesis by (auto dest: dvd_unit_imp_unit)
+qed (auto simp: multiplicity_unit_left)
+
+lemma multiplicity_one [simp]: "multiplicity p 1 = 0"
+  by (rule multiplicity_unit_right) simp_all
+
+lemma multiplicity_eqI:
+  assumes "p ^ n dvd x" "\<not>p ^ Suc n dvd x"
+  shows   "multiplicity p x = n"
+proof -
+  consider "x = 0" | "is_unit p" | "x \<noteq> 0" "\<not>is_unit p" by blast
+  thus ?thesis
+  proof cases
+    assume xp: "x \<noteq> 0" "\<not>is_unit p"
+    from xp assms(1) have "multiplicity p x \<ge> n" by (intro multiplicity_geI)
+    moreover from assms(2) xp have "multiplicity p x < Suc n" by (intro multiplicity_lessI)
+    ultimately show ?thesis by simp
+  next
+    assume "is_unit p"
+    hence "is_unit (p ^ Suc n)" by (simp add: is_unit_power_iff del: power_Suc)
+    hence "p ^ Suc n dvd x" by (rule unit_imp_dvd)
+    with \<open>\<not>p ^ Suc n dvd x\<close> show ?thesis by contradiction
+  qed (insert assms, simp_all)
+qed
+
+
+context
+  fixes x p :: 'a
+  assumes xp: "x \<noteq> 0" "\<not>is_unit p"
+begin
+
+lemma multiplicity_times_same:
+  assumes "p \<noteq> 0"
+  shows   "multiplicity p (p * x) = Suc (multiplicity p x)"
+proof (rule multiplicity_eqI)
+  show "p ^ Suc (multiplicity p x) dvd p * x"
+    by (auto intro!: mult_dvd_mono multiplicity_dvd)
+  from xp assms show "\<not> p ^ Suc (Suc (multiplicity p x)) dvd p * x"
+    using power_dvd_iff_le_multiplicity[OF xp, of "Suc (multiplicity p x)"] by simp
+qed
+
+end
+
+lemma multiplicity_same_power': "multiplicity p (p ^ n) = (if p = 0 \<or> is_unit p then 0 else n)"
+proof -
+  consider "p = 0" | "is_unit p" |"p \<noteq> 0" "\<not>is_unit p" by blast
+  thus ?thesis
+  proof cases
+    assume "p \<noteq> 0" "\<not>is_unit p"
+    thus ?thesis by (induction n) (simp_all add: multiplicity_times_same)
+  qed (simp_all add: power_0_left multiplicity_unit_left)
+qed
+
+lemma multiplicity_same_power:
+  "p \<noteq> 0 \<Longrightarrow> \<not>is_unit p \<Longrightarrow> multiplicity p (p ^ n) = n"
+  by (simp add: multiplicity_same_power')
+
+lemma multiplicity_prime_elem_times_other:
+  assumes "prime_elem p" "\<not>p dvd q"
+  shows   "multiplicity p (q * x) = multiplicity p x"
+proof (cases "x = 0")
+  case False
+  show ?thesis
+  proof (rule multiplicity_eqI)
+    have "1 * p ^ multiplicity p x dvd q * x"
+      by (intro mult_dvd_mono multiplicity_dvd) simp_all
+    thus "p ^ multiplicity p x dvd q * x" by simp
+  next
+    define n where "n = multiplicity p x"
+    from assms have "\<not>is_unit p" by simp
+    from multiplicity_decompose'[OF False this] guess y . note y = this [folded n_def]
+    from y have "p ^ Suc n dvd q * x \<longleftrightarrow> p ^ n * p dvd p ^ n * (q * y)" by (simp add: mult_ac)
+    also from assms have "\<dots> \<longleftrightarrow> p dvd q * y" by simp
+    also have "\<dots> \<longleftrightarrow> p dvd q \<or> p dvd y" by (rule prime_elem_dvd_mult_iff) fact+
+    also from assms y have "\<dots> \<longleftrightarrow> False" by simp
+    finally show "\<not>(p ^ Suc n dvd q * x)" by blast
+  qed
+qed simp_all
+
+lemma multiplicity_self:
+  assumes "p \<noteq> 0" "\<not>is_unit p"
+  shows   "multiplicity p p = 1"
+proof -
+  from assms have "multiplicity p p = Max {n. p ^ n dvd p}"
+    by (simp add: multiplicity_eq_Max)
+  also from assms have "p ^ n dvd p \<longleftrightarrow> n \<le> 1" for n
+    using dvd_power_iff[of p n 1] by auto
+  hence "{n. p ^ n dvd p} = {..1}" by auto
+  also have "\<dots> = {0,1}" by auto
+  finally show ?thesis by simp
+qed
+
+lemma multiplicity_times_unit_left:
+  assumes "is_unit c"
+  shows   "multiplicity (c * p) x = multiplicity p x"
+proof -
+  from assms have "{n. (c * p) ^ n dvd x} = {n. p ^ n dvd x}"
+    by (subst mult.commute) (simp add: mult_unit_dvd_iff power_mult_distrib is_unit_power_iff)
+  thus ?thesis by (simp add: multiplicity_def)
+qed
+
+lemma multiplicity_times_unit_right:
+  assumes "is_unit c"
+  shows   "multiplicity p (c * x) = multiplicity p x"
+proof -
+  from assms have "{n. p ^ n dvd c * x} = {n. p ^ n dvd x}"
+    by (subst mult.commute) (simp add: dvd_mult_unit_iff)
+  thus ?thesis by (simp add: multiplicity_def)
+qed
+
+lemma multiplicity_normalize_left [simp]:
+  "multiplicity (normalize p) x = multiplicity p x"
+proof (cases "p = 0")
+  case [simp]: False
+  have "normalize p = (1 div unit_factor p) * p"
+    by (simp add: unit_div_commute is_unit_unit_factor)
+  also have "multiplicity \<dots> x = multiplicity p x"
+    by (rule multiplicity_times_unit_left) (simp add: is_unit_unit_factor)
+  finally show ?thesis .
+qed simp_all
+
+lemma multiplicity_normalize_right [simp]:
+  "multiplicity p (normalize x) = multiplicity p x"
+proof (cases "x = 0")
+  case [simp]: False
+  have "normalize x = (1 div unit_factor x) * x"
+    by (simp add: unit_div_commute is_unit_unit_factor)
+  also have "multiplicity p \<dots> = multiplicity p x"
+    by (rule multiplicity_times_unit_right) (simp add: is_unit_unit_factor)
+  finally show ?thesis .
+qed simp_all   
+
+lemma multiplicity_prime [simp]: "prime_elem p \<Longrightarrow> multiplicity p p = 1"
+  by (rule multiplicity_self) auto
+
+lemma multiplicity_prime_power [simp]: "prime_elem p \<Longrightarrow> multiplicity p (p ^ n) = n"
+  by (subst multiplicity_same_power') auto
+
+lift_definition prime_factorization :: "'a \<Rightarrow> 'a multiset" is
+  "\<lambda>x p. if prime p then multiplicity p x else 0"
+  unfolding multiset_def
+proof clarify
+  fix x :: 'a
+  show "finite {p. 0 < (if prime p then multiplicity p x else 0)}" (is "finite ?A")
+  proof (cases "x = 0")
+    case False
+    from False have "?A \<subseteq> {p. prime p \<and> p dvd x}"
+      by (auto simp: multiplicity_gt_zero_iff)
+    moreover from False have "finite {p. prime p \<and> p dvd x}"
+      by (rule finite_prime_divisors)
+    ultimately show ?thesis by (rule finite_subset)
+  qed simp_all
+qed
+
+abbreviation prime_factors :: "'a \<Rightarrow> 'a set" where
+  "prime_factors a \<equiv> set_mset (prime_factorization a)"
+
+lemma count_prime_factorization_nonprime:
+  "\<not>prime p \<Longrightarrow> count (prime_factorization x) p = 0"
+  by transfer simp
+
+lemma count_prime_factorization_prime:
+  "prime p \<Longrightarrow> count (prime_factorization x) p = multiplicity p x"
+  by transfer simp
+
+lemma count_prime_factorization:
+  "count (prime_factorization x) p = (if prime p then multiplicity p x else 0)"
+  by transfer simp
+
+lemma dvd_imp_multiplicity_le:
+  assumes "a dvd b" "b \<noteq> 0"
+  shows   "multiplicity p a \<le> multiplicity p b"
+proof (cases "is_unit p")
+  case False
+  with assms show ?thesis
+    by (intro multiplicity_geI ) (auto intro: dvd_trans[OF multiplicity_dvd' assms(1)])
+qed (insert assms, auto simp: multiplicity_unit_left)
+
+lemma prime_factorization_0 [simp]: "prime_factorization 0 = {#}"
+  by (simp add: multiset_eq_iff count_prime_factorization)
+
+lemma prime_factorization_empty_iff:
+  "prime_factorization x = {#} \<longleftrightarrow> x = 0 \<or> is_unit x"
+proof
+  assume *: "prime_factorization x = {#}"
+  {
+    assume x: "x \<noteq> 0" "\<not>is_unit x"
+    {
+      fix p assume p: "prime p"
+      have "count (prime_factorization x) p = 0" by (simp add: *)
+      also from p have "count (prime_factorization x) p = multiplicity p x"
+        by (rule count_prime_factorization_prime)
+      also from x p have "\<dots> = 0 \<longleftrightarrow> \<not>p dvd x" by (simp add: multiplicity_eq_zero_iff)
+      finally have "\<not>p dvd x" .
+    }
+    with prime_divisor_exists[OF x] have False by blast
+  }
+  thus "x = 0 \<or> is_unit x" by blast
+next
+  assume "x = 0 \<or> is_unit x"
+  thus "prime_factorization x = {#}"
+  proof
+    assume x: "is_unit x"
+    {
+      fix p assume p: "prime p"
+      from p x have "multiplicity p x = 0"
+        by (subst multiplicity_eq_zero_iff)
+           (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors)
+    }
+    thus ?thesis by (simp add: multiset_eq_iff count_prime_factorization)
+  qed simp_all
+qed
+
+lemma prime_factorization_unit:
+  assumes "is_unit x"
+  shows   "prime_factorization x = {#}"
+proof (rule multiset_eqI)
+  fix p :: 'a
+  show "count (prime_factorization x) p = count {#} p"
+  proof (cases "prime p")
+    case True
+    with assms have "multiplicity p x = 0"
+      by (subst multiplicity_eq_zero_iff)
+         (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors)
+    with True show ?thesis by (simp add: count_prime_factorization_prime)
+  qed (simp_all add: count_prime_factorization_nonprime)
+qed
+
+lemma prime_factorization_1 [simp]: "prime_factorization 1 = {#}"
+  by (simp add: prime_factorization_unit)
+
+lemma prime_factorization_times_prime:
+  assumes "x \<noteq> 0" "prime p"
+  shows   "prime_factorization (p * x) = {#p#} + prime_factorization x"
+proof (rule multiset_eqI)
+  fix q :: 'a
+  consider "\<not>prime q" | "p = q" | "prime q" "p \<noteq> q" by blast
+  thus "count (prime_factorization (p * x)) q = count ({#p#} + prime_factorization x) q"
+  proof cases
+    assume q: "prime q" "p \<noteq> q"
+    with assms primes_dvd_imp_eq[of q p] have "\<not>q dvd p" by auto
+    with q assms show ?thesis
+      by (simp add: multiplicity_prime_elem_times_other count_prime_factorization)
+  qed (insert assms, auto simp: count_prime_factorization multiplicity_times_same)
+qed
+
+lemma prod_mset_prime_factorization:
+  assumes "x \<noteq> 0"
+  shows   "prod_mset (prime_factorization x) = normalize x"
+  using assms
+  by (induction x rule: prime_divisors_induct)
+     (simp_all add: prime_factorization_unit prime_factorization_times_prime
+                    is_unit_normalize normalize_mult)
+
+lemma in_prime_factors_iff:
+  "p \<in> prime_factors x \<longleftrightarrow> x \<noteq> 0 \<and> p dvd x \<and> prime p"
+proof -
+  have "p \<in> prime_factors x \<longleftrightarrow> count (prime_factorization x) p > 0" by simp
+  also have "\<dots> \<longleftrightarrow> x \<noteq> 0 \<and> p dvd x \<and> prime p"
+   by (subst count_prime_factorization, cases "x = 0")
+      (auto simp: multiplicity_eq_zero_iff multiplicity_gt_zero_iff)
+  finally show ?thesis .
+qed
+
+lemma in_prime_factors_imp_prime [intro]:
+  "p \<in> prime_factors x \<Longrightarrow> prime p"
+  by (simp add: in_prime_factors_iff)
+
+lemma in_prime_factors_imp_dvd [dest]:
+  "p \<in> prime_factors x \<Longrightarrow> p dvd x"
+  by (simp add: in_prime_factors_iff)
+
+lemma prime_factorsI:
+  "x \<noteq> 0 \<Longrightarrow> prime p \<Longrightarrow> p dvd x \<Longrightarrow> p \<in> prime_factors x"
+  by (auto simp: in_prime_factors_iff)
+
+lemma prime_factors_dvd:
+  "x \<noteq> 0 \<Longrightarrow> prime_factors x = {p. prime p \<and> p dvd x}"
+  by (auto intro: prime_factorsI)
+
+lemma prime_factors_multiplicity:
+  "prime_factors n = {p. prime p \<and> multiplicity p n > 0}"
+  by (cases "n = 0") (auto simp add: prime_factors_dvd prime_multiplicity_gt_zero_iff)
+
+lemma prime_factorization_prime:
+  assumes "prime p"
+  shows   "prime_factorization p = {#p#}"
+proof (rule multiset_eqI)
+  fix q :: 'a
+  consider "\<not>prime q" | "q = p" | "prime q" "q \<noteq> p" by blast
+  thus "count (prime_factorization p) q = count {#p#} q"
+    by cases (insert assms, auto dest: primes_dvd_imp_eq
+                simp: count_prime_factorization multiplicity_self multiplicity_eq_zero_iff)
+qed
+
+lemma prime_factorization_prod_mset_primes:
+  assumes "\<And>p. p \<in># A \<Longrightarrow> prime p"
+  shows   "prime_factorization (prod_mset A) = A"
+  using assms
+proof (induction A)
+  case (add p A)
+  from add.prems[of 0] have "0 \<notin># A" by auto
+  hence "prod_mset A \<noteq> 0" by auto
+  with add show ?case
+    by (simp_all add: mult_ac prime_factorization_times_prime Multiset.union_commute)
+qed simp_all
+
+lemma prime_factorization_cong:
+  "normalize x = normalize y \<Longrightarrow> prime_factorization x = prime_factorization y"
+  by (simp add: multiset_eq_iff count_prime_factorization
+                multiplicity_normalize_right [of _ x, symmetric]
+                multiplicity_normalize_right [of _ y, symmetric]
+           del:  multiplicity_normalize_right)
+
+lemma prime_factorization_unique:
+  assumes "x \<noteq> 0" "y \<noteq> 0"
+  shows   "prime_factorization x = prime_factorization y \<longleftrightarrow> normalize x = normalize y"
+proof
+  assume "prime_factorization x = prime_factorization y"
+  hence "prod_mset (prime_factorization x) = prod_mset (prime_factorization y)" by simp
+  with assms show "normalize x = normalize y" by (simp add: prod_mset_prime_factorization)
+qed (rule prime_factorization_cong)
+
+lemma prime_factorization_mult:
+  assumes "x \<noteq> 0" "y \<noteq> 0"
+  shows   "prime_factorization (x * y) = prime_factorization x + prime_factorization y"
+proof -
+  have "prime_factorization (prod_mset (prime_factorization (x * y))) =
+          prime_factorization (prod_mset (prime_factorization x + prime_factorization y))"
+    by (simp add: prod_mset_prime_factorization assms normalize_mult)
+  also have "prime_factorization (prod_mset (prime_factorization (x * y))) =
+               prime_factorization (x * y)"
+    by (rule prime_factorization_prod_mset_primes) (simp_all add: in_prime_factors_imp_prime)
+  also have "prime_factorization (prod_mset (prime_factorization x + prime_factorization y)) =
+               prime_factorization x + prime_factorization y"
+    by (rule prime_factorization_prod_mset_primes) (auto simp: in_prime_factors_imp_prime)
+  finally show ?thesis .
+qed
+
+lemma prime_elem_multiplicity_mult_distrib:
+  assumes "prime_elem p" "x \<noteq> 0" "y \<noteq> 0"
+  shows   "multiplicity p (x * y) = multiplicity p x + multiplicity p y"
+proof -
+  have "multiplicity p (x * y) = count (prime_factorization (x * y)) (normalize p)"
+    by (subst count_prime_factorization_prime) (simp_all add: assms)
+  also from assms 
+    have "prime_factorization (x * y) = prime_factorization x + prime_factorization y"
+      by (intro prime_factorization_mult)
+  also have "count \<dots> (normalize p) = 
+    count (prime_factorization x) (normalize p) + count (prime_factorization y) (normalize p)"
+    by simp
+  also have "\<dots> = multiplicity p x + multiplicity p y" 
+    by (subst (1 2) count_prime_factorization_prime) (simp_all add: assms)
+  finally show ?thesis .
+qed
+
+lemma prime_elem_multiplicity_prod_mset_distrib:
+  assumes "prime_elem p" "0 \<notin># A"
+  shows   "multiplicity p (prod_mset A) = sum_mset (image_mset (multiplicity p) A)"
+  using assms by (induction A) (auto simp: prime_elem_multiplicity_mult_distrib)
+
+lemma prime_elem_multiplicity_power_distrib:
+  assumes "prime_elem p" "x \<noteq> 0"
+  shows   "multiplicity p (x ^ n) = n * multiplicity p x"
+  using assms prime_elem_multiplicity_prod_mset_distrib [of p "replicate_mset n x"]
+  by simp
+
+lemma prime_elem_multiplicity_prod_distrib:
+  assumes "prime_elem p" "0 \<notin> f ` A" "finite A"
+  shows   "multiplicity p (prod f A) = (\<Sum>x\<in>A. multiplicity p (f x))"
+proof -
+  have "multiplicity p (prod f A) = (\<Sum>x\<in>#mset_set A. multiplicity p (f x))"
+    using assms by (subst prod_unfold_prod_mset)
+                   (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset 
+                      multiset.map_comp o_def)
+  also from \<open>finite A\<close> have "\<dots> = (\<Sum>x\<in>A. multiplicity p (f x))"
+    by (induction A rule: finite_induct) simp_all
+  finally show ?thesis .
+qed
+
+lemma multiplicity_distinct_prime_power:
+  "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> multiplicity p (q ^ n) = 0"
+  by (subst prime_elem_multiplicity_power_distrib) (auto simp: prime_multiplicity_other)
+
+lemma prime_factorization_prime_power:
+  "prime p \<Longrightarrow> prime_factorization (p ^ n) = replicate_mset n p"
+  by (induction n)
+     (simp_all add: prime_factorization_mult prime_factorization_prime Multiset.union_commute)
+
+lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x"
+  by (cases "x = 0") (simp_all add: prod_mset_prime_factorization)
+
+lemma prime_factorization_subset_iff_dvd:
+  assumes [simp]: "x \<noteq> 0" "y \<noteq> 0"
+  shows   "prime_factorization x \<subseteq># prime_factorization y \<longleftrightarrow> x dvd y"
+proof -
+  have "x dvd y \<longleftrightarrow> prod_mset (prime_factorization x) dvd prod_mset (prime_factorization y)"
+    by (simp add: prod_mset_prime_factorization)
+  also have "\<dots> \<longleftrightarrow> prime_factorization x \<subseteq># prime_factorization y"
+    by (auto intro!: prod_mset_primes_dvd_imp_subset prod_mset_subset_imp_dvd)
+  finally show ?thesis ..
+qed
+
+lemma prime_factorization_subset_imp_dvd: 
+  "x \<noteq> 0 \<Longrightarrow> (prime_factorization x \<subseteq># prime_factorization y) \<Longrightarrow> x dvd y"
+  by (cases "y = 0") (simp_all add: prime_factorization_subset_iff_dvd)
+
+lemma prime_factorization_divide:
+  assumes "b dvd a"
+  shows   "prime_factorization (a div b) = prime_factorization a - prime_factorization b"
+proof (cases "a = 0")
+  case [simp]: False
+  from assms have [simp]: "b \<noteq> 0" by auto
+  have "prime_factorization ((a div b) * b) = prime_factorization (a div b) + prime_factorization b"
+    by (intro prime_factorization_mult) (insert assms, auto elim!: dvdE)
+  with assms show ?thesis by simp
+qed simp_all
+
+lemma zero_not_in_prime_factors [simp]: "0 \<notin> prime_factors x"
+  by (auto dest: in_prime_factors_imp_prime)
+
+lemma prime_prime_factors:
+  "prime p \<Longrightarrow> prime_factors p = {p}"
+  by (drule prime_factorization_prime) simp
+
+lemma prod_prime_factors:
+  assumes "x \<noteq> 0"
+  shows   "(\<Prod>p \<in> prime_factors x. p ^ multiplicity p x) = normalize x"
+proof -
+  have "normalize x = prod_mset (prime_factorization x)"
+    by (simp add: prod_mset_prime_factorization assms)
+  also have "\<dots> = (\<Prod>p \<in> prime_factors x. p ^ count (prime_factorization x) p)"
+    by (subst prod_mset_multiplicity) simp_all
+  also have "\<dots> = (\<Prod>p \<in> prime_factors x. p ^ multiplicity p x)"
+    by (intro prod.cong) 
+      (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime)
+  finally show ?thesis ..
+qed
+
+lemma prime_factorization_unique'':
+  assumes S_eq: "S = {p. 0 < f p}"
+    and "finite S"
+    and S: "\<forall>p\<in>S. prime p" "normalize n = (\<Prod>p\<in>S. p ^ f p)"
+  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
+proof
+  define A where "A = Abs_multiset f"
+  from \<open>finite S\<close> S(1) have "(\<Prod>p\<in>S. p ^ f p) \<noteq> 0" by auto
+  with S(2) have nz: "n \<noteq> 0" by auto
+  from S_eq \<open>finite S\<close> have count_A: "count A x = f x" for x
+    unfolding A_def by (subst multiset.Abs_multiset_inverse) (simp_all add: multiset_def)
+  from S_eq count_A have set_mset_A: "set_mset A = S"
+    by (simp only: set_mset_def)
+  from S(2) have "normalize n = (\<Prod>p\<in>S. p ^ f p)" .
+  also have "\<dots> = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A)
+  also from nz have "normalize n = prod_mset (prime_factorization n)" 
+    by (simp add: prod_mset_prime_factorization)
+  finally have "prime_factorization (prod_mset A) = 
+                  prime_factorization (prod_mset (prime_factorization n))" by simp
+  also from S(1) have "prime_factorization (prod_mset A) = A"
+    by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A)
+  also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n"
+    by (intro prime_factorization_prod_mset_primes) auto
+  finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric])
+  
+  show "(\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
+  proof safe
+    fix p :: 'a assume p: "prime p"
+    have "multiplicity p n = multiplicity p (normalize n)" by simp
+    also have "normalize n = prod_mset A" 
+      by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S)
+    also from p set_mset_A S(1) 
+    have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
+      by (intro prime_elem_multiplicity_prod_mset_distrib) auto
+    also from S(1) p
+    have "image_mset (multiplicity p) A = image_mset (\<lambda>q. if p = q then 1 else 0) A"
+      by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other)
+    also have "sum_mset \<dots> = f p" by (simp add: sum_mset_delta' count_A)
+    finally show "f p = multiplicity p n" ..
+  qed
+qed
+
+lemma prime_factors_product: 
+  "x \<noteq> 0 \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> prime_factors (x * y) = prime_factors x \<union> prime_factors y"
+  by (simp add: prime_factorization_mult)
+
+lemma dvd_prime_factors [intro]:
+  "y \<noteq> 0 \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x \<subseteq> prime_factors y"
+  by (intro set_mset_mono, subst prime_factorization_subset_iff_dvd) auto
+
+(* RENAMED multiplicity_dvd *)
+lemma multiplicity_le_imp_dvd:
+  assumes "x \<noteq> 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x \<le> multiplicity p y"
+  shows   "x dvd y"
+proof (cases "y = 0")
+  case False
+  from assms this have "prime_factorization x \<subseteq># prime_factorization y"
+    by (intro mset_subset_eqI) (auto simp: count_prime_factorization)
+  with assms False show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd)
+qed auto
+
+lemma dvd_multiplicity_eq:
+  "x \<noteq> 0 \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> x dvd y \<longleftrightarrow> (\<forall>p. multiplicity p x \<le> multiplicity p y)"
+  by (auto intro: dvd_imp_multiplicity_le multiplicity_le_imp_dvd)
+
+lemma multiplicity_eq_imp_eq:
+  assumes "x \<noteq> 0" "y \<noteq> 0"
+  assumes "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+  shows   "normalize x = normalize y"
+  using assms by (intro associatedI multiplicity_le_imp_dvd) simp_all
+
+lemma prime_factorization_unique':
+  assumes "\<forall>p \<in># M. prime p" "\<forall>p \<in># N. prime p" "(\<Prod>i \<in># M. i) = (\<Prod>i \<in># N. i)"
+  shows   "M = N"
+proof -
+  have "prime_factorization (\<Prod>i \<in># M. i) = prime_factorization (\<Prod>i \<in># N. i)"
+    by (simp only: assms)
+  also from assms have "prime_factorization (\<Prod>i \<in># M. i) = M"
+    by (subst prime_factorization_prod_mset_primes) simp_all
+  also from assms have "prime_factorization (\<Prod>i \<in># N. i) = N"
+    by (subst prime_factorization_prod_mset_primes) simp_all
+  finally show ?thesis .
+qed
+
+lemma multiplicity_cong:
+  "(\<And>r. p ^ r dvd a \<longleftrightarrow> p ^ r dvd b) \<Longrightarrow> multiplicity p a = multiplicity p b"
+  by (simp add: multiplicity_def)
+
+lemma not_dvd_imp_multiplicity_0: 
+  assumes "\<not>p dvd x"
+  shows   "multiplicity p x = 0"
+proof -
+  from assms have "multiplicity p x < 1"
+    by (intro multiplicity_lessI) auto
+  thus ?thesis by simp
+qed
+
+
+subsection \<open>GCD and LCM computation with unique factorizations\<close>
+
+definition "gcd_factorial a b = (if a = 0 then normalize b
+     else if b = 0 then normalize a
+     else prod_mset (prime_factorization a \<inter># prime_factorization b))"
+
+definition "lcm_factorial a b = (if a = 0 \<or> b = 0 then 0
+     else prod_mset (prime_factorization a \<union># prime_factorization b))"
+
+definition "Gcd_factorial A =
+  (if A \<subseteq> {0} then 0 else prod_mset (Inf (prime_factorization ` (A - {0}))))"
+
+definition "Lcm_factorial A =
+  (if A = {} then 1
+   else if 0 \<notin> A \<and> subset_mset.bdd_above (prime_factorization ` (A - {0})) then
+     prod_mset (Sup (prime_factorization ` A))
+   else
+     0)"
+
+lemma prime_factorization_gcd_factorial:
+  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
+  shows   "prime_factorization (gcd_factorial a b) = prime_factorization a \<inter># prime_factorization b"
+proof -
+  have "prime_factorization (gcd_factorial a b) =
+          prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
+    by (simp add: gcd_factorial_def)
+  also have "\<dots> = prime_factorization a \<inter># prime_factorization b"
+    by (subst prime_factorization_prod_mset_primes) auto
+  finally show ?thesis .
+qed
+
+lemma prime_factorization_lcm_factorial:
+  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
+  shows   "prime_factorization (lcm_factorial a b) = prime_factorization a \<union># prime_factorization b"
+proof -
+  have "prime_factorization (lcm_factorial a b) =
+          prime_factorization (prod_mset (prime_factorization a \<union># prime_factorization b))"
+    by (simp add: lcm_factorial_def)
+  also have "\<dots> = prime_factorization a \<union># prime_factorization b"
+    by (subst prime_factorization_prod_mset_primes) auto
+  finally show ?thesis .
+qed
+
+lemma prime_factorization_Gcd_factorial:
+  assumes "\<not>A \<subseteq> {0}"
+  shows   "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))"
+proof -
+  from assms obtain x where x: "x \<in> A - {0}" by auto
+  hence "Inf (prime_factorization ` (A - {0})) \<subseteq># prime_factorization x"
+    by (intro subset_mset.cInf_lower) simp_all
+  hence "\<forall>y. y \<in># Inf (prime_factorization ` (A - {0})) \<longrightarrow> y \<in> prime_factors x"
+    by (auto dest: mset_subset_eqD)
+  with in_prime_factors_imp_prime[of _ x]
+    have "\<forall>p. p \<in># Inf (prime_factorization ` (A - {0})) \<longrightarrow> prime p" by blast
+  with assms show ?thesis
+    by (simp add: Gcd_factorial_def prime_factorization_prod_mset_primes)
+qed
+
+lemma prime_factorization_Lcm_factorial:
+  assumes "0 \<notin> A" "subset_mset.bdd_above (prime_factorization ` A)"
+  shows   "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)"
+proof (cases "A = {}")
+  case True
+  hence "prime_factorization ` A = {}" by auto
+  also have "Sup \<dots> = {#}" by (simp add: Sup_multiset_empty)
+  finally show ?thesis by (simp add: Lcm_factorial_def)
+next
+  case False
+  have "\<forall>y. y \<in># Sup (prime_factorization ` A) \<longrightarrow> prime y"
+    by (auto simp: in_Sup_multiset_iff assms)
+  with assms False show ?thesis
+    by (simp add: Lcm_factorial_def prime_factorization_prod_mset_primes)
+qed
+
+lemma gcd_factorial_commute: "gcd_factorial a b = gcd_factorial b a"
+  by (simp add: gcd_factorial_def multiset_inter_commute)
+
+lemma gcd_factorial_dvd1: "gcd_factorial a b dvd a"
+proof (cases "a = 0 \<or> b = 0")
+  case False
+  hence "gcd_factorial a b \<noteq> 0" by (auto simp: gcd_factorial_def)
+  with False show ?thesis
+    by (subst prime_factorization_subset_iff_dvd [symmetric])
+       (auto simp: prime_factorization_gcd_factorial)
+qed (auto simp: gcd_factorial_def)
+
+lemma gcd_factorial_dvd2: "gcd_factorial a b dvd b"
+  by (subst gcd_factorial_commute) (rule gcd_factorial_dvd1)
+
+lemma normalize_gcd_factorial: "normalize (gcd_factorial a b) = gcd_factorial a b"
+proof -
+  have "normalize (prod_mset (prime_factorization a \<inter># prime_factorization b)) =
+          prod_mset (prime_factorization a \<inter># prime_factorization b)"
+    by (intro normalize_prod_mset_primes) auto
+  thus ?thesis by (simp add: gcd_factorial_def)
+qed
+
+lemma gcd_factorial_greatest: "c dvd gcd_factorial a b" if "c dvd a" "c dvd b" for a b c
+proof (cases "a = 0 \<or> b = 0")
+  case False
+  with that have [simp]: "c \<noteq> 0" by auto
+  let ?p = "prime_factorization"
+  from that False have "?p c \<subseteq># ?p a" "?p c \<subseteq># ?p b"
+    by (simp_all add: prime_factorization_subset_iff_dvd)
+  hence "prime_factorization c \<subseteq>#
+           prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
+    using False by (subst prime_factorization_prod_mset_primes) auto
+  with False show ?thesis
+    by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric])
+qed (auto simp: gcd_factorial_def that)
+
+lemma lcm_factorial_gcd_factorial:
+  "lcm_factorial a b = normalize (a * b) div gcd_factorial a b" for a b
+proof (cases "a = 0 \<or> b = 0")
+  case False
+  let ?p = "prime_factorization"
+  from False have "prod_mset (?p (a * b)) div gcd_factorial a b =
+                     prod_mset (?p a + ?p b - ?p a \<inter># ?p b)"
+    by (subst prod_mset_diff) (auto simp: lcm_factorial_def gcd_factorial_def
+                                prime_factorization_mult subset_mset.le_infI1)
+  also from False have "prod_mset (?p (a * b)) = normalize (a * b)"
+    by (intro prod_mset_prime_factorization) simp_all
+  also from False have "prod_mset (?p a + ?p b - ?p a \<inter># ?p b) = lcm_factorial a b"
+    by (simp add: union_diff_inter_eq_sup lcm_factorial_def)
+  finally show ?thesis ..
+qed (auto simp: lcm_factorial_def)
+
+lemma normalize_Gcd_factorial:
+  "normalize (Gcd_factorial A) = Gcd_factorial A"
+proof (cases "A \<subseteq> {0}")
+  case False
+  then obtain x where "x \<in> A" "x \<noteq> 0" by blast
+  hence "Inf (prime_factorization ` (A - {0})) \<subseteq># prime_factorization x"
+    by (intro subset_mset.cInf_lower) auto
+  hence "prime p" if "p \<in># Inf (prime_factorization ` (A - {0}))" for p
+    using that by (auto dest: mset_subset_eqD)
+  with False show ?thesis
+    by (auto simp add: Gcd_factorial_def normalize_prod_mset_primes)
+qed (simp_all add: Gcd_factorial_def)
+
+lemma Gcd_factorial_eq_0_iff:
+  "Gcd_factorial A = 0 \<longleftrightarrow> A \<subseteq> {0}"
+  by (auto simp: Gcd_factorial_def in_Inf_multiset_iff split: if_splits)
+
+lemma Gcd_factorial_dvd:
+  assumes "x \<in> A"
+  shows   "Gcd_factorial A dvd x"
+proof (cases "x = 0")
+  case False
+  with assms have "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))"
+    by (intro prime_factorization_Gcd_factorial) auto
+  also from False assms have "\<dots> \<subseteq># prime_factorization x"
+    by (intro subset_mset.cInf_lower) auto
+  finally show ?thesis
+    by (subst (asm) prime_factorization_subset_iff_dvd)
+       (insert assms False, auto simp: Gcd_factorial_eq_0_iff)
+qed simp_all
+
+lemma Gcd_factorial_greatest:
+  assumes "\<And>y. y \<in> A \<Longrightarrow> x dvd y"
+  shows   "x dvd Gcd_factorial A"
+proof (cases "A \<subseteq> {0}")
+  case False
+  from False obtain y where "y \<in> A" "y \<noteq> 0" by auto
+  with assms[of y] have nz: "x \<noteq> 0" by auto
+  from nz assms have "prime_factorization x \<subseteq># prime_factorization y" if "y \<in> A - {0}" for y
+    using that by (subst prime_factorization_subset_iff_dvd) auto
+  with False have "prime_factorization x \<subseteq># Inf (prime_factorization ` (A - {0}))"
+    by (intro subset_mset.cInf_greatest) auto
+  also from False have "\<dots> = prime_factorization (Gcd_factorial A)"
+    by (rule prime_factorization_Gcd_factorial [symmetric])
+  finally show ?thesis
+    by (subst (asm) prime_factorization_subset_iff_dvd)
+       (insert nz False, auto simp: Gcd_factorial_eq_0_iff)
+qed (simp_all add: Gcd_factorial_def)
+
+lemma normalize_Lcm_factorial:
+  "normalize (Lcm_factorial A) = Lcm_factorial A"
+proof (cases "subset_mset.bdd_above (prime_factorization ` A)")
+  case True
+  hence "normalize (prod_mset (Sup (prime_factorization ` A))) =
+           prod_mset (Sup (prime_factorization ` A))"
+    by (intro normalize_prod_mset_primes)
+       (auto simp: in_Sup_multiset_iff)
+  with True show ?thesis by (simp add: Lcm_factorial_def)
+qed (auto simp: Lcm_factorial_def)
+
+lemma Lcm_factorial_eq_0_iff:
+  "Lcm_factorial A = 0 \<longleftrightarrow> 0 \<in> A \<or> \<not>subset_mset.bdd_above (prime_factorization ` A)"
+  by (auto simp: Lcm_factorial_def in_Sup_multiset_iff)
+
+lemma dvd_Lcm_factorial:
+  assumes "x \<in> A"
+  shows   "x dvd Lcm_factorial A"
+proof (cases "0 \<notin> A \<and> subset_mset.bdd_above (prime_factorization ` A)")
+  case True
+  with assms have [simp]: "0 \<notin> A" "x \<noteq> 0" "A \<noteq> {}" by auto
+  from assms True have "prime_factorization x \<subseteq># Sup (prime_factorization ` A)"
+    by (intro subset_mset.cSup_upper) auto
+  also have "\<dots> = prime_factorization (Lcm_factorial A)"
+    by (rule prime_factorization_Lcm_factorial [symmetric]) (insert True, simp_all)
+  finally show ?thesis
+    by (subst (asm) prime_factorization_subset_iff_dvd)
+       (insert True, auto simp: Lcm_factorial_eq_0_iff)
+qed (insert assms, auto simp: Lcm_factorial_def)
+
+lemma Lcm_factorial_least:
+  assumes "\<And>y. y \<in> A \<Longrightarrow> y dvd x"
+  shows   "Lcm_factorial A dvd x"
+proof -
+  consider "A = {}" | "0 \<in> A" | "x = 0" | "A \<noteq> {}" "0 \<notin> A" "x \<noteq> 0" by blast
+  thus ?thesis
+  proof cases
+    assume *: "A \<noteq> {}" "0 \<notin> A" "x \<noteq> 0"
+    hence nz: "x \<noteq> 0" if "x \<in> A" for x using that by auto
+    from * have bdd: "subset_mset.bdd_above (prime_factorization ` A)"
+      by (intro subset_mset.bdd_aboveI[of _ "prime_factorization x"])
+         (auto simp: prime_factorization_subset_iff_dvd nz dest: assms)
+    have "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)"
+      by (rule prime_factorization_Lcm_factorial) fact+
+    also from * have "\<dots> \<subseteq># prime_factorization x"
+      by (intro subset_mset.cSup_least)
+         (auto simp: prime_factorization_subset_iff_dvd nz dest: assms)
+    finally show ?thesis
+      by (subst (asm) prime_factorization_subset_iff_dvd)
+         (insert * bdd, auto simp: Lcm_factorial_eq_0_iff)
+  qed (auto simp: Lcm_factorial_def dest: assms)
+qed
+
+lemmas gcd_lcm_factorial =
+  gcd_factorial_dvd1 gcd_factorial_dvd2 gcd_factorial_greatest
+  normalize_gcd_factorial lcm_factorial_gcd_factorial
+  normalize_Gcd_factorial Gcd_factorial_dvd Gcd_factorial_greatest
+  normalize_Lcm_factorial dvd_Lcm_factorial Lcm_factorial_least
+
+end
+
+class factorial_semiring_gcd = factorial_semiring + gcd + Gcd +
+  assumes gcd_eq_gcd_factorial: "gcd a b = gcd_factorial a b"
+  and     lcm_eq_lcm_factorial: "lcm a b = lcm_factorial a b"
+  and     Gcd_eq_Gcd_factorial: "Gcd A = Gcd_factorial A"
+  and     Lcm_eq_Lcm_factorial: "Lcm A = Lcm_factorial A"
+begin
+
+lemma prime_factorization_gcd:
+  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
+  shows   "prime_factorization (gcd a b) = prime_factorization a \<inter># prime_factorization b"
+  by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial)
+
+lemma prime_factorization_lcm:
+  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
+  shows   "prime_factorization (lcm a b) = prime_factorization a \<union># prime_factorization b"
+  by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial)
+
+lemma prime_factorization_Gcd:
+  assumes "Gcd A \<noteq> 0"
+  shows   "prime_factorization (Gcd A) = Inf (prime_factorization ` (A - {0}))"
+  using assms
+  by (simp add: prime_factorization_Gcd_factorial Gcd_eq_Gcd_factorial Gcd_factorial_eq_0_iff)
+
+lemma prime_factorization_Lcm:
+  assumes "Lcm A \<noteq> 0"
+  shows   "prime_factorization (Lcm A) = Sup (prime_factorization ` A)"
+  using assms
+  by (simp add: prime_factorization_Lcm_factorial Lcm_eq_Lcm_factorial Lcm_factorial_eq_0_iff)
+
+subclass semiring_gcd
+  by (standard, unfold gcd_eq_gcd_factorial lcm_eq_lcm_factorial)
+     (rule gcd_lcm_factorial; assumption)+
+
+subclass semiring_Gcd
+  by (standard, unfold Gcd_eq_Gcd_factorial Lcm_eq_Lcm_factorial)
+     (rule gcd_lcm_factorial; assumption)+
+
+lemma
+  assumes "x \<noteq> 0" "y \<noteq> 0"
+  shows gcd_eq_factorial': 
+          "gcd x y = (\<Prod>p \<in> prime_factors x \<inter> prime_factors y. 
+                          p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs1")
+    and lcm_eq_factorial':
+          "lcm x y = (\<Prod>p \<in> prime_factors x \<union> prime_factors y. 
+                          p ^ max (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs2")
+proof -
+  have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial)
+  also have "\<dots> = ?rhs1"
+    by (auto simp: gcd_factorial_def assms prod_mset_multiplicity
+          count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong)
+  finally show "gcd x y = ?rhs1" .
+  have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial)
+  also have "\<dots> = ?rhs2"
+    by (auto simp: lcm_factorial_def assms prod_mset_multiplicity
+          count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong)
+  finally show "lcm x y = ?rhs2" .
+qed
+
+lemma
+  assumes "x \<noteq> 0" "y \<noteq> 0" "prime p"
+  shows   multiplicity_gcd: "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)"
+    and   multiplicity_lcm: "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)"
+proof -
+  have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial)
+  also from assms have "multiplicity p \<dots> = min (multiplicity p x) (multiplicity p y)"
+    by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_gcd_factorial)
+  finally show "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" .
+  have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial)
+  also from assms have "multiplicity p \<dots> = max (multiplicity p x) (multiplicity p y)"
+    by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_lcm_factorial)
+  finally show "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" .
+qed
+
+lemma gcd_lcm_distrib:
+  "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)"
+proof (cases "x = 0 \<or> y = 0 \<or> z = 0")
+  case True
+  thus ?thesis
+    by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd)
+next
+  case False
+  hence "normalize (gcd x (lcm y z)) = normalize (lcm (gcd x y) (gcd x z))"
+    by (intro associatedI prime_factorization_subset_imp_dvd)
+       (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm 
+          subset_mset.inf_sup_distrib1)
+  thus ?thesis by simp
+qed
+
+lemma lcm_gcd_distrib:
+  "lcm x (gcd y z) = gcd (lcm x y) (lcm x z)"
+proof (cases "x = 0 \<or> y = 0 \<or> z = 0")
+  case True
+  thus ?thesis
+    by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd)
+next
+  case False
+  hence "normalize (lcm x (gcd y z)) = normalize (gcd (lcm x y) (lcm x z))"
+    by (intro associatedI prime_factorization_subset_imp_dvd)
+       (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm 
+          subset_mset.sup_inf_distrib1)
+  thus ?thesis by simp
+qed
+
+end
+
+class factorial_ring_gcd = factorial_semiring_gcd + idom
+begin
+
+subclass ring_gcd ..
+
+subclass idom_divide ..
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Field_as_Ring.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,120 @@
+(*  Title:      HOL/Library/Field_as_Ring.thy
+    Author:     Manuel Eberl
+*)
+
+theory Field_as_Ring
+imports 
+  Complex_Main
+  Euclidean_Algorithm
+begin
+
+context field
+begin
+
+subclass idom_divide ..
+
+definition normalize_field :: "'a \<Rightarrow> 'a" 
+  where [simp]: "normalize_field x = (if x = 0 then 0 else 1)"
+definition unit_factor_field :: "'a \<Rightarrow> 'a" 
+  where [simp]: "unit_factor_field x = x"
+definition euclidean_size_field :: "'a \<Rightarrow> nat" 
+  where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)"
+definition mod_field :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  where [simp]: "mod_field x y = (if y = 0 then x else 0)"
+
+end
+
+instantiation real :: unique_euclidean_ring
+begin
+
+definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)"
+definition [simp]: "unit_factor_real = (unit_factor_field :: real \<Rightarrow> _)"
+definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \<Rightarrow> _)"
+definition [simp]: "uniqueness_constraint_real = (top :: real \<Rightarrow> real \<Rightarrow> bool)"
+definition [simp]: "modulo_real = (mod_field :: real \<Rightarrow> _)"
+
+instance
+  by standard
+    (simp_all add: dvd_field_iff divide_simps split: if_splits)
+
+end
+
+instantiation real :: euclidean_ring_gcd
+begin
+
+definition gcd_real :: "real \<Rightarrow> real \<Rightarrow> real" where
+  "gcd_real = Euclidean_Algorithm.gcd"
+definition lcm_real :: "real \<Rightarrow> real \<Rightarrow> real" where
+  "lcm_real = Euclidean_Algorithm.lcm"
+definition Gcd_real :: "real set \<Rightarrow> real" where
+ "Gcd_real = Euclidean_Algorithm.Gcd"
+definition Lcm_real :: "real set \<Rightarrow> real" where
+ "Lcm_real = Euclidean_Algorithm.Lcm"
+
+instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def)
+
+end
+
+instantiation rat :: unique_euclidean_ring
+begin
+
+definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)"
+definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \<Rightarrow> _)"
+definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \<Rightarrow> _)"
+definition [simp]: "uniqueness_constraint_rat = (top :: rat \<Rightarrow> rat \<Rightarrow> bool)"
+definition [simp]: "modulo_rat = (mod_field :: rat \<Rightarrow> _)"
+
+instance
+  by standard
+    (simp_all add: dvd_field_iff divide_simps split: if_splits)
+
+end
+
+instantiation rat :: euclidean_ring_gcd
+begin
+
+definition gcd_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
+  "gcd_rat = Euclidean_Algorithm.gcd"
+definition lcm_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
+  "lcm_rat = Euclidean_Algorithm.lcm"
+definition Gcd_rat :: "rat set \<Rightarrow> rat" where
+ "Gcd_rat = Euclidean_Algorithm.Gcd"
+definition Lcm_rat :: "rat set \<Rightarrow> rat" where
+ "Lcm_rat = Euclidean_Algorithm.Lcm"
+
+instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def)
+
+end
+
+instantiation complex :: unique_euclidean_ring
+begin
+
+definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)"
+definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \<Rightarrow> _)"
+definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \<Rightarrow> _)"
+definition [simp]: "uniqueness_constraint_complex = (top :: complex \<Rightarrow> complex \<Rightarrow> bool)"
+definition [simp]: "modulo_complex = (mod_field :: complex \<Rightarrow> _)"
+
+instance
+  by standard
+    (simp_all add: dvd_field_iff divide_simps split: if_splits)
+
+end
+
+instantiation complex :: euclidean_ring_gcd
+begin
+
+definition gcd_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
+  "gcd_complex = Euclidean_Algorithm.gcd"
+definition lcm_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
+  "lcm_complex = Euclidean_Algorithm.lcm"
+definition Gcd_complex :: "complex set \<Rightarrow> complex" where
+ "Gcd_complex = Euclidean_Algorithm.Gcd"
+definition Lcm_complex :: "complex set \<Rightarrow> complex" where
+ "Lcm_complex = Euclidean_Algorithm.Lcm"
+
+instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def)
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,4785 @@
+(*  Title:      HOL/Library/Formal_Power_Series.thy
+    Author:     Amine Chaieb, University of Cambridge
+*)
+
+section \<open>A formalization of formal power series\<close>
+
+theory Formal_Power_Series
+imports
+  Complex_Main
+  Euclidean_Algorithm
+begin
+
+
+subsection \<open>The type of formal power series\<close>
+
+typedef 'a fps = "{f :: nat \<Rightarrow> 'a. True}"
+  morphisms fps_nth Abs_fps
+  by simp
+
+notation fps_nth (infixl "$" 75)
+
+lemma expand_fps_eq: "p = q \<longleftrightarrow> (\<forall>n. p $ n = q $ n)"
+  by (simp add: fps_nth_inject [symmetric] fun_eq_iff)
+
+lemma fps_ext: "(\<And>n. p $ n = q $ n) \<Longrightarrow> p = q"
+  by (simp add: expand_fps_eq)
+
+lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n"
+  by (simp add: Abs_fps_inverse)
+
+text \<open>Definition of the basic elements 0 and 1 and the basic operations of addition,
+  negation and multiplication.\<close>
+
+instantiation fps :: (zero) zero
+begin
+  definition fps_zero_def: "0 = Abs_fps (\<lambda>n. 0)"
+  instance ..
+end
+
+lemma fps_zero_nth [simp]: "0 $ n = 0"
+  unfolding fps_zero_def by simp
+
+instantiation fps :: ("{one, zero}") one
+begin
+  definition fps_one_def: "1 = Abs_fps (\<lambda>n. if n = 0 then 1 else 0)"
+  instance ..
+end
+
+lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)"
+  unfolding fps_one_def by simp
+
+instantiation fps :: (plus) plus
+begin
+  definition fps_plus_def: "op + = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n + g $ n))"
+  instance ..
+end
+
+lemma fps_add_nth [simp]: "(f + g) $ n = f $ n + g $ n"
+  unfolding fps_plus_def by simp
+
+instantiation fps :: (minus) minus
+begin
+  definition fps_minus_def: "op - = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n - g $ n))"
+  instance ..
+end
+
+lemma fps_sub_nth [simp]: "(f - g) $ n = f $ n - g $ n"
+  unfolding fps_minus_def by simp
+
+instantiation fps :: (uminus) uminus
+begin
+  definition fps_uminus_def: "uminus = (\<lambda>f. Abs_fps (\<lambda>n. - (f $ n)))"
+  instance ..
+end
+
+lemma fps_neg_nth [simp]: "(- f) $ n = - (f $ n)"
+  unfolding fps_uminus_def by simp
+
+instantiation fps :: ("{comm_monoid_add, times}") times
+begin
+  definition fps_times_def: "op * = (\<lambda>f g. Abs_fps (\<lambda>n. \<Sum>i=0..n. f $ i * g $ (n - i)))"
+  instance ..
+end
+
+lemma fps_mult_nth: "(f * g) $ n = (\<Sum>i=0..n. f$i * g$(n - i))"
+  unfolding fps_times_def by simp
+
+lemma fps_mult_nth_0 [simp]: "(f * g) $ 0 = f $ 0 * g $ 0"
+  unfolding fps_times_def by simp
+
+declare atLeastAtMost_iff [presburger]
+declare Bex_def [presburger]
+declare Ball_def [presburger]
+
+lemma mult_delta_left:
+  fixes x y :: "'a::mult_zero"
+  shows "(if b then x else 0) * y = (if b then x * y else 0)"
+  by simp
+
+lemma mult_delta_right:
+  fixes x y :: "'a::mult_zero"
+  shows "x * (if b then y else 0) = (if b then x * y else 0)"
+  by simp
+
+lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
+  by auto
+
+lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
+  by auto
+
+
+subsection \<open>Formal power series form a commutative ring with unity, if the range of sequences
+  they represent is a commutative ring with unity\<close>
+
+instance fps :: (semigroup_add) semigroup_add
+proof
+  fix a b c :: "'a fps"
+  show "a + b + c = a + (b + c)"
+    by (simp add: fps_ext add.assoc)
+qed
+
+instance fps :: (ab_semigroup_add) ab_semigroup_add
+proof
+  fix a b :: "'a fps"
+  show "a + b = b + a"
+    by (simp add: fps_ext add.commute)
+qed
+
+lemma fps_mult_assoc_lemma:
+  fixes k :: nat
+    and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
+  shows "(\<Sum>j=0..k. \<Sum>i=0..j. f i (j - i) (n - j)) =
+         (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
+  by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc)
+
+instance fps :: (semiring_0) semigroup_mult
+proof
+  fix a b c :: "'a fps"
+  show "(a * b) * c = a * (b * c)"
+  proof (rule fps_ext)
+    fix n :: nat
+    have "(\<Sum>j=0..n. \<Sum>i=0..j. a$i * b$(j - i) * c$(n - j)) =
+          (\<Sum>j=0..n. \<Sum>i=0..n - j. a$j * b$i * c$(n - j - i))"
+      by (rule fps_mult_assoc_lemma)
+    then show "((a * b) * c) $ n = (a * (b * c)) $ n"
+      by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc)
+  qed
+qed
+
+lemma fps_mult_commute_lemma:
+  fixes n :: nat
+    and f :: "nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
+  shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
+  by (rule sum.reindex_bij_witness[where i="op - n" and j="op - n"]) auto
+
+instance fps :: (comm_semiring_0) ab_semigroup_mult
+proof
+  fix a b :: "'a fps"
+  show "a * b = b * a"
+  proof (rule fps_ext)
+    fix n :: nat
+    have "(\<Sum>i=0..n. a$i * b$(n - i)) = (\<Sum>i=0..n. a$(n - i) * b$i)"
+      by (rule fps_mult_commute_lemma)
+    then show "(a * b) $ n = (b * a) $ n"
+      by (simp add: fps_mult_nth mult.commute)
+  qed
+qed
+
+instance fps :: (monoid_add) monoid_add
+proof
+  fix a :: "'a fps"
+  show "0 + a = a" by (simp add: fps_ext)
+  show "a + 0 = a" by (simp add: fps_ext)
+qed
+
+instance fps :: (comm_monoid_add) comm_monoid_add
+proof
+  fix a :: "'a fps"
+  show "0 + a = a" by (simp add: fps_ext)
+qed
+
+instance fps :: (semiring_1) monoid_mult
+proof
+  fix a :: "'a fps"
+  show "1 * a = a"
+    by (simp add: fps_ext fps_mult_nth mult_delta_left sum.delta)
+  show "a * 1 = a"
+    by (simp add: fps_ext fps_mult_nth mult_delta_right sum.delta')
+qed
+
+instance fps :: (cancel_semigroup_add) cancel_semigroup_add
+proof
+  fix a b c :: "'a fps"
+  show "b = c" if "a + b = a + c"
+    using that by (simp add: expand_fps_eq)
+  show "b = c" if "b + a = c + a"
+    using that by (simp add: expand_fps_eq)
+qed
+
+instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
+proof
+  fix a b c :: "'a fps"
+  show "a + b - a = b"
+    by (simp add: expand_fps_eq)
+  show "a - b - c = a - (b + c)"
+    by (simp add: expand_fps_eq diff_diff_eq)
+qed
+
+instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
+
+instance fps :: (group_add) group_add
+proof
+  fix a b :: "'a fps"
+  show "- a + a = 0" by (simp add: fps_ext)
+  show "a + - b = a - b" by (simp add: fps_ext)
+qed
+
+instance fps :: (ab_group_add) ab_group_add
+proof
+  fix a b :: "'a fps"
+  show "- a + a = 0" by (simp add: fps_ext)
+  show "a - b = a + - b" by (simp add: fps_ext)
+qed
+
+instance fps :: (zero_neq_one) zero_neq_one
+  by standard (simp add: expand_fps_eq)
+
+instance fps :: (semiring_0) semiring
+proof
+  fix a b c :: "'a fps"
+  show "(a + b) * c = a * c + b * c"
+    by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib)
+  show "a * (b + c) = a * b + a * c"
+    by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib)
+qed
+
+instance fps :: (semiring_0) semiring_0
+proof
+  fix a :: "'a fps"
+  show "0 * a = 0"
+    by (simp add: fps_ext fps_mult_nth)
+  show "a * 0 = 0"
+    by (simp add: fps_ext fps_mult_nth)
+qed
+
+instance fps :: (semiring_0_cancel) semiring_0_cancel ..
+
+instance fps :: (semiring_1) semiring_1 ..
+
+
+subsection \<open>Selection of the nth power of the implicit variable in the infinite sum\<close>
+
+lemma fps_square_nth: "(f^2) $ n = (\<Sum>k\<le>n. f $ k * f $ (n - k))"
+  by (simp add: power2_eq_square fps_mult_nth atLeast0AtMost)
+
+lemma fps_nonzero_nth: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0)"
+  by (simp add: expand_fps_eq)
+
+lemma fps_nonzero_nth_minimal: "f \<noteq> 0 \<longleftrightarrow> (\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m < n. f $ m = 0))"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  let ?n = "LEAST n. f $ n \<noteq> 0"
+  show ?rhs if ?lhs
+  proof -
+    from that have "\<exists>n. f $ n \<noteq> 0"
+      by (simp add: fps_nonzero_nth)
+    then have "f $ ?n \<noteq> 0"
+      by (rule LeastI_ex)
+    moreover have "\<forall>m<?n. f $ m = 0"
+      by (auto dest: not_less_Least)
+    ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
+    then show ?thesis ..
+  qed
+  show ?lhs if ?rhs
+    using that by (auto simp add: expand_fps_eq)
+qed
+
+lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
+  by (rule expand_fps_eq)
+
+lemma fps_sum_nth: "sum f S $ n = sum (\<lambda>k. (f k) $ n) S"
+proof (cases "finite S")
+  case True
+  then show ?thesis by (induct set: finite) auto
+next
+  case False
+  then show ?thesis by simp
+qed
+
+
+subsection \<open>Injection of the basic ring elements and multiplication by scalars\<close>
+
+definition "fps_const c = Abs_fps (\<lambda>n. if n = 0 then c else 0)"
+
+lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)"
+  unfolding fps_const_def by simp
+
+lemma fps_const_0_eq_0 [simp]: "fps_const 0 = 0"
+  by (simp add: fps_ext)
+
+lemma fps_const_1_eq_1 [simp]: "fps_const 1 = 1"
+  by (simp add: fps_ext)
+
+lemma fps_const_neg [simp]: "- (fps_const (c::'a::ring)) = fps_const (- c)"
+  by (simp add: fps_ext)
+
+lemma fps_const_add [simp]: "fps_const (c::'a::monoid_add) + fps_const d = fps_const (c + d)"
+  by (simp add: fps_ext)
+
+lemma fps_const_sub [simp]: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
+  by (simp add: fps_ext)
+
+lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
+  by (simp add: fps_eq_iff fps_mult_nth sum.neutral)
+
+lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
+    Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
+  by (simp add: fps_ext)
+
+lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) =
+    Abs_fps (\<lambda>n. if n = 0 then f$0 + c else f$n)"
+  by (simp add: fps_ext)
+
+lemma fps_const_mult_left: "fps_const (c::'a::semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
+  unfolding fps_eq_iff fps_mult_nth
+  by (simp add: fps_const_def mult_delta_left sum.delta)
+
+lemma fps_const_mult_right: "f * fps_const (c::'a::semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
+  unfolding fps_eq_iff fps_mult_nth
+  by (simp add: fps_const_def mult_delta_right sum.delta')
+
+lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
+  by (simp add: fps_mult_nth mult_delta_left sum.delta)
+
+lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
+  by (simp add: fps_mult_nth mult_delta_right sum.delta')
+
+
+subsection \<open>Formal power series form an integral domain\<close>
+
+instance fps :: (ring) ring ..
+
+instance fps :: (ring_1) ring_1
+  by (intro_classes, auto simp add: distrib_right)
+
+instance fps :: (comm_ring_1) comm_ring_1
+  by (intro_classes, auto simp add: distrib_right)
+
+instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
+proof
+  fix a b :: "'a fps"
+  assume "a \<noteq> 0" and "b \<noteq> 0"
+  then obtain i j where i: "a $ i \<noteq> 0" "\<forall>k<i. a $ k = 0" and j: "b $ j \<noteq> 0" "\<forall>k<j. b $ k =0"
+    unfolding fps_nonzero_nth_minimal
+    by blast+
+  have "(a * b) $ (i + j) = (\<Sum>k=0..i+j. a $ k * b $ (i + j - k))"
+    by (rule fps_mult_nth)
+  also have "\<dots> = (a $ i * b $ (i + j - i)) + (\<Sum>k\<in>{0..i+j} - {i}. a $ k * b $ (i + j - k))"
+    by (rule sum.remove) simp_all
+  also have "(\<Sum>k\<in>{0..i+j}-{i}. a $ k * b $ (i + j - k)) = 0"
+  proof (rule sum.neutral [rule_format])
+    fix k assume "k \<in> {0..i+j} - {i}"
+    then have "k < i \<or> i+j-k < j"
+      by auto
+    then show "a $ k * b $ (i + j - k) = 0"
+      using i j by auto
+  qed
+  also have "a $ i * b $ (i + j - i) + 0 = a $ i * b $ j"
+    by simp
+  also have "a $ i * b $ j \<noteq> 0"
+    using i j by simp
+  finally have "(a*b) $ (i+j) \<noteq> 0" .
+  then show "a * b \<noteq> 0"
+    unfolding fps_nonzero_nth by blast
+qed
+
+instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
+
+instance fps :: (idom) idom ..
+
+lemma numeral_fps_const: "numeral k = fps_const (numeral k)"
+  by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1
+    fps_const_add [symmetric])
+
+lemma neg_numeral_fps_const:
+  "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)"
+  by (simp add: numeral_fps_const)
+
+lemma fps_numeral_nth: "numeral n $ i = (if i = 0 then numeral n else 0)"
+  by (simp add: numeral_fps_const)
+
+lemma fps_numeral_nth_0 [simp]: "numeral n $ 0 = numeral n"
+  by (simp add: numeral_fps_const)
+
+lemma fps_of_nat: "fps_const (of_nat c) = of_nat c"
+  by (induction c) (simp_all add: fps_const_add [symmetric] del: fps_const_add)
+
+lemma numeral_neq_fps_zero [simp]: "(numeral f :: 'a :: field_char_0 fps) \<noteq> 0"
+proof
+  assume "numeral f = (0 :: 'a fps)"
+  from arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] show False by simp
+qed 
+
+
+subsection \<open>The eXtractor series X\<close>
+
+lemma minus_one_power_iff: "(- (1::'a::comm_ring_1)) ^ n = (if even n then 1 else - 1)"
+  by (induct n) auto
+
+definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
+
+lemma X_mult_nth [simp]:
+  "(X * (f :: 'a::semiring_1 fps)) $n = (if n = 0 then 0 else f $ (n - 1))"
+proof (cases "n = 0")
+  case False
+  have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
+    by (simp add: fps_mult_nth)
+  also have "\<dots> = f $ (n - 1)"
+    using False by (simp add: X_def mult_delta_left sum.delta)
+  finally show ?thesis
+    using False by simp
+next
+  case True
+  then show ?thesis
+    by (simp add: fps_mult_nth X_def)
+qed
+
+lemma X_mult_right_nth[simp]:
+  "((a::'a::semiring_1 fps) * X) $ n = (if n = 0 then 0 else a $ (n - 1))"
+proof -
+  have "(a * X) $ n = (\<Sum>i = 0..n. a $ i * (if n - i = Suc 0 then 1 else 0))"
+    by (simp add: fps_times_def X_def)
+  also have "\<dots> = (\<Sum>i = 0..n. if i = n - 1 then if n = 0 then 0 else a $ i else 0)"
+    by (intro sum.cong) auto
+  also have "\<dots> = (if n = 0 then 0 else a $ (n - 1))" by (simp add: sum.delta)
+  finally show ?thesis .
+qed
+
+lemma fps_mult_X_commute: "X * (a :: 'a :: semiring_1 fps) = a * X" 
+  by (simp add: fps_eq_iff)
+
+lemma X_power_iff: "X^k = Abs_fps (\<lambda>n. if n = k then 1::'a::comm_ring_1 else 0)"
+proof (induct k)
+  case 0
+  then show ?case by (simp add: X_def fps_eq_iff)
+next
+  case (Suc k)
+  have "(X^Suc k) $ m = (if m = Suc k then 1::'a else 0)" for m
+  proof -
+    have "(X^Suc k) $ m = (if m = 0 then 0 else (X^k) $ (m - 1))"
+      by (simp del: One_nat_def)
+    then show ?thesis
+      using Suc.hyps by (auto cong del: if_weak_cong)
+  qed
+  then show ?case
+    by (simp add: fps_eq_iff)
+qed
+
+lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)"
+  by (simp add: X_def)
+
+lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else 0::'a::comm_ring_1)"
+  by (simp add: X_power_iff)
+
+lemma X_power_mult_nth: "(X^k * (f :: 'a::comm_ring_1 fps)) $n = (if n < k then 0 else f $ (n - k))"
+  apply (induct k arbitrary: n)
+  apply simp
+  unfolding power_Suc mult.assoc
+  apply (case_tac n)
+  apply auto
+  done
+
+lemma X_power_mult_right_nth:
+    "((f :: 'a::comm_ring_1 fps) * X^k) $n = (if n < k then 0 else f $ (n - k))"
+  by (metis X_power_mult_nth mult.commute)
+
+
+lemma X_neq_fps_const [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> fps_const c"
+proof
+  assume "(X::'a fps) = fps_const (c::'a)"
+  hence "X$1 = (fps_const (c::'a))$1" by (simp only:)
+  thus False by auto
+qed
+
+lemma X_neq_zero [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 0"
+  by (simp only: fps_const_0_eq_0[symmetric] X_neq_fps_const) simp
+
+lemma X_neq_one [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 1"
+  by (simp only: fps_const_1_eq_1[symmetric] X_neq_fps_const) simp
+
+lemma X_neq_numeral [simp]: "(X :: 'a :: {semiring_1,zero_neq_one} fps) \<noteq> numeral c"
+  by (simp only: numeral_fps_const X_neq_fps_const) simp
+
+lemma X_pow_eq_X_pow_iff [simp]:
+  "(X :: ('a :: {comm_ring_1}) fps) ^ m = X ^ n \<longleftrightarrow> m = n"
+proof
+  assume "(X :: 'a fps) ^ m = X ^ n"
+  hence "(X :: 'a fps) ^ m $ m = X ^ n $ m" by (simp only:)
+  thus "m = n" by (simp split: if_split_asm)
+qed simp_all
+
+
+subsection \<open>Subdegrees\<close>
+
+definition subdegree :: "('a::zero) fps \<Rightarrow> nat" where
+  "subdegree f = (if f = 0 then 0 else LEAST n. f$n \<noteq> 0)"
+
+lemma subdegreeI:
+  assumes "f $ d \<noteq> 0" and "\<And>i. i < d \<Longrightarrow> f $ i = 0"
+  shows   "subdegree f = d"
+proof-
+  from assms(1) have "f \<noteq> 0" by auto
+  moreover from assms(1) have "(LEAST i. f $ i \<noteq> 0) = d"
+  proof (rule Least_equality)
+    fix e assume "f $ e \<noteq> 0"
+    with assms(2) have "\<not>(e < d)" by blast
+    thus "e \<ge> d" by simp
+  qed
+  ultimately show ?thesis unfolding subdegree_def by simp
+qed
+
+lemma nth_subdegree_nonzero [simp,intro]: "f \<noteq> 0 \<Longrightarrow> f $ subdegree f \<noteq> 0"
+proof-
+  assume "f \<noteq> 0"
+  hence "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
+  also from \<open>f \<noteq> 0\<close> have "\<exists>n. f$n \<noteq> 0" using fps_nonzero_nth by blast
+  from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \<noteq> 0) \<noteq> 0" .
+  finally show ?thesis .
+qed
+
+lemma nth_less_subdegree_zero [dest]: "n < subdegree f \<Longrightarrow> f $ n = 0"
+proof (cases "f = 0")
+  assume "f \<noteq> 0" and less: "n < subdegree f"
+  note less
+  also from \<open>f \<noteq> 0\<close> have "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
+  finally show "f $ n = 0" using not_less_Least by blast
+qed simp_all
+
+lemma subdegree_geI:
+  assumes "f \<noteq> 0" "\<And>i. i < n \<Longrightarrow> f$i = 0"
+  shows   "subdegree f \<ge> n"
+proof (rule ccontr)
+  assume "\<not>(subdegree f \<ge> n)"
+  with assms(2) have "f $ subdegree f = 0" by simp
+  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
+  ultimately show False by contradiction
+qed
+
+lemma subdegree_greaterI:
+  assumes "f \<noteq> 0" "\<And>i. i \<le> n \<Longrightarrow> f$i = 0"
+  shows   "subdegree f > n"
+proof (rule ccontr)
+  assume "\<not>(subdegree f > n)"
+  with assms(2) have "f $ subdegree f = 0" by simp
+  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
+  ultimately show False by contradiction
+qed
+
+lemma subdegree_leI:
+  "f $ n \<noteq> 0 \<Longrightarrow> subdegree f \<le> n"
+  by (rule leI) auto
+
+
+lemma subdegree_0 [simp]: "subdegree 0 = 0"
+  by (simp add: subdegree_def)
+
+lemma subdegree_1 [simp]: "subdegree (1 :: ('a :: zero_neq_one) fps) = 0"
+  by (auto intro!: subdegreeI)
+
+lemma subdegree_X [simp]: "subdegree (X :: ('a :: zero_neq_one) fps) = 1"
+  by (auto intro!: subdegreeI simp: X_def)
+
+lemma subdegree_fps_const [simp]: "subdegree (fps_const c) = 0"
+  by (cases "c = 0") (auto intro!: subdegreeI)
+
+lemma subdegree_numeral [simp]: "subdegree (numeral n) = 0"
+  by (simp add: numeral_fps_const)
+
+lemma subdegree_eq_0_iff: "subdegree f = 0 \<longleftrightarrow> f = 0 \<or> f $ 0 \<noteq> 0"
+proof (cases "f = 0")
+  assume "f \<noteq> 0"
+  thus ?thesis
+    using nth_subdegree_nonzero[OF \<open>f \<noteq> 0\<close>] by (fastforce intro!: subdegreeI)
+qed simp_all
+
+lemma subdegree_eq_0 [simp]: "f $ 0 \<noteq> 0 \<Longrightarrow> subdegree f = 0"
+  by (simp add: subdegree_eq_0_iff)
+
+lemma nth_subdegree_mult [simp]:
+  fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps"
+  shows "(f * g) $ (subdegree f + subdegree g) = f $ subdegree f * g $ subdegree g"
+proof-
+  let ?n = "subdegree f + subdegree g"
+  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))"
+    by (simp add: fps_mult_nth)
+  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
+  proof (intro sum.cong)
+    fix x assume x: "x \<in> {0..?n}"
+    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
+    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
+      by (elim disjE conjE) auto
+  qed auto
+  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
+  finally show ?thesis .
+qed
+
+lemma subdegree_mult [simp]:
+  assumes "f \<noteq> 0" "g \<noteq> 0"
+  shows "subdegree ((f :: ('a :: {ring_no_zero_divisors}) fps) * g) = subdegree f + subdegree g"
+proof (rule subdegreeI)
+  let ?n = "subdegree f + subdegree g"
+  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth)
+  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
+  proof (intro sum.cong)
+    fix x assume x: "x \<in> {0..?n}"
+    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
+    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
+      by (elim disjE conjE) auto
+  qed auto
+  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
+  also from assms have "... \<noteq> 0" by auto
+  finally show "(f * g) $ (subdegree f + subdegree g) \<noteq> 0" .
+next
+  fix m assume m: "m < subdegree f + subdegree g"
+  have "(f * g) $ m = (\<Sum>i=0..m. f$i * g$(m-i))" by (simp add: fps_mult_nth)
+  also have "... = (\<Sum>i=0..m. 0)"
+  proof (rule sum.cong)
+    fix i assume "i \<in> {0..m}"
+    with m have "i < subdegree f \<or> m - i < subdegree g" by auto
+    thus "f$i * g$(m-i) = 0" by (elim disjE) auto
+  qed auto
+  finally show "(f * g) $ m = 0" by simp
+qed
+
+lemma subdegree_power [simp]:
+  "subdegree ((f :: ('a :: ring_1_no_zero_divisors) fps) ^ n) = n * subdegree f"
+  by (cases "f = 0"; induction n) simp_all
+
+lemma subdegree_uminus [simp]:
+  "subdegree (-(f::('a::group_add) fps)) = subdegree f"
+  by (simp add: subdegree_def)
+
+lemma subdegree_minus_commute [simp]:
+  "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)"
+proof -
+  have "f - g = -(g - f)" by simp
+  also have "subdegree ... = subdegree (g - f)" by (simp only: subdegree_uminus)
+  finally show ?thesis .
+qed
+
+lemma subdegree_add_ge:
+  assumes "f \<noteq> -(g :: ('a :: {group_add}) fps)"
+  shows   "subdegree (f + g) \<ge> min (subdegree f) (subdegree g)"
+proof (rule subdegree_geI)
+  from assms show "f + g \<noteq> 0" by (subst (asm) eq_neg_iff_add_eq_0)
+next
+  fix i assume "i < min (subdegree f) (subdegree g)"
+  hence "f $ i = 0" and "g $ i = 0" by auto
+  thus "(f + g) $ i = 0" by force
+qed
+
+lemma subdegree_add_eq1:
+  assumes "f \<noteq> 0"
+  assumes "subdegree f < subdegree (g :: ('a :: {group_add}) fps)"
+  shows   "subdegree (f + g) = subdegree f"
+proof (rule antisym[OF subdegree_leI])
+  from assms show "subdegree (f + g) \<ge> subdegree f"
+    by (intro order.trans[OF min.boundedI subdegree_add_ge]) auto
+  from assms have "f $ subdegree f \<noteq> 0" "g $ subdegree f = 0" by auto
+  thus "(f + g) $ subdegree f \<noteq> 0" by simp
+qed
+
+lemma subdegree_add_eq2:
+  assumes "g \<noteq> 0"
+  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
+  shows   "subdegree (f + g) = subdegree g"
+  using subdegree_add_eq1[OF assms] by (simp add: add.commute)
+
+lemma subdegree_diff_eq1:
+  assumes "f \<noteq> 0"
+  assumes "subdegree f < subdegree (g :: ('a :: {ab_group_add}) fps)"
+  shows   "subdegree (f - g) = subdegree f"
+  using subdegree_add_eq1[of f "-g"] assms by (simp add: add.commute)
+
+lemma subdegree_diff_eq2:
+  assumes "g \<noteq> 0"
+  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
+  shows   "subdegree (f - g) = subdegree g"
+  using subdegree_add_eq2[of "-g" f] assms by (simp add: add.commute)
+
+lemma subdegree_diff_ge [simp]:
+  assumes "f \<noteq> (g :: ('a :: {group_add}) fps)"
+  shows   "subdegree (f - g) \<ge> min (subdegree f) (subdegree g)"
+  using assms subdegree_add_ge[of f "-g"] by simp
+
+
+
+
+subsection \<open>Shifting and slicing\<close>
+
+definition fps_shift :: "nat \<Rightarrow> 'a fps \<Rightarrow> 'a fps" where
+  "fps_shift n f = Abs_fps (\<lambda>i. f $ (i + n))"
+
+lemma fps_shift_nth [simp]: "fps_shift n f $ i = f $ (i + n)"
+  by (simp add: fps_shift_def)
+
+lemma fps_shift_0 [simp]: "fps_shift 0 f = f"
+  by (intro fps_ext) (simp add: fps_shift_def)
+
+lemma fps_shift_zero [simp]: "fps_shift n 0 = 0"
+  by (intro fps_ext) (simp add: fps_shift_def)
+
+lemma fps_shift_one: "fps_shift n 1 = (if n = 0 then 1 else 0)"
+  by (intro fps_ext) (simp add: fps_shift_def)
+
+lemma fps_shift_fps_const: "fps_shift n (fps_const c) = (if n = 0 then fps_const c else 0)"
+  by (intro fps_ext) (simp add: fps_shift_def)
+
+lemma fps_shift_numeral: "fps_shift n (numeral c) = (if n = 0 then numeral c else 0)"
+  by (simp add: numeral_fps_const fps_shift_fps_const)
+
+lemma fps_shift_X_power [simp]:
+  "n \<le> m \<Longrightarrow> fps_shift n (X ^ m) = (X ^ (m - n) ::'a::comm_ring_1 fps)"
+  by (intro fps_ext) (auto simp: fps_shift_def )
+
+lemma fps_shift_times_X_power:
+  "n \<le> subdegree f \<Longrightarrow> fps_shift n f * X ^ n = (f :: 'a :: comm_ring_1 fps)"
+  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
+
+lemma fps_shift_times_X_power' [simp]:
+  "fps_shift n (f * X^n) = (f :: 'a :: comm_ring_1 fps)"
+  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
+
+lemma fps_shift_times_X_power'':
+  "m \<le> n \<Longrightarrow> fps_shift n (f * X^m) = fps_shift (n - m) (f :: 'a :: comm_ring_1 fps)"
+  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
+
+lemma fps_shift_subdegree [simp]:
+  "n \<le> subdegree f \<Longrightarrow> subdegree (fps_shift n f) = subdegree (f :: 'a :: comm_ring_1 fps) - n"
+  by (cases "f = 0") (force intro: nth_less_subdegree_zero subdegreeI)+
+
+lemma subdegree_decompose:
+  "f = fps_shift (subdegree f) f * X ^ subdegree (f :: ('a :: comm_ring_1) fps)"
+  by (rule fps_ext) (auto simp: X_power_mult_right_nth)
+
+lemma subdegree_decompose':
+  "n \<le> subdegree (f :: ('a :: comm_ring_1) fps) \<Longrightarrow> f = fps_shift n f * X^n"
+  by (rule fps_ext) (auto simp: X_power_mult_right_nth intro!: nth_less_subdegree_zero)
+
+lemma fps_shift_fps_shift:
+  "fps_shift (m + n) f = fps_shift m (fps_shift n f)"
+  by (rule fps_ext) (simp add: add_ac)
+
+lemma fps_shift_add:
+  "fps_shift n (f + g) = fps_shift n f + fps_shift n g"
+  by (simp add: fps_eq_iff)
+
+lemma fps_shift_mult:
+  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
+  shows   "fps_shift n (h*g) = h * fps_shift n g"
+proof -
+  from assms have "g = fps_shift n g * X^n" by (rule subdegree_decompose')
+  also have "h * ... = (h * fps_shift n g) * X^n" by simp
+  also have "fps_shift n ... = h * fps_shift n g" by simp
+  finally show ?thesis .
+qed
+
+lemma fps_shift_mult_right:
+  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
+  shows   "fps_shift n (g*h) = h * fps_shift n g"
+  by (subst mult.commute, subst fps_shift_mult) (simp_all add: assms)
+
+lemma nth_subdegree_zero_iff [simp]: "f $ subdegree f = 0 \<longleftrightarrow> f = 0"
+  by (cases "f = 0") auto
+
+lemma fps_shift_subdegree_zero_iff [simp]:
+  "fps_shift (subdegree f) f = 0 \<longleftrightarrow> f = 0"
+  by (subst (1) nth_subdegree_zero_iff[symmetric], cases "f = 0")
+     (simp_all del: nth_subdegree_zero_iff)
+
+
+definition "fps_cutoff n f = Abs_fps (\<lambda>i. if i < n then f$i else 0)"
+
+lemma fps_cutoff_nth [simp]: "fps_cutoff n f $ i = (if i < n then f$i else 0)"
+  unfolding fps_cutoff_def by simp
+
+lemma fps_cutoff_zero_iff: "fps_cutoff n f = 0 \<longleftrightarrow> (f = 0 \<or> n \<le> subdegree f)"
+proof
+  assume A: "fps_cutoff n f = 0"
+  thus "f = 0 \<or> n \<le> subdegree f"
+  proof (cases "f = 0")
+    assume "f \<noteq> 0"
+    with A have "n \<le> subdegree f"
+      by (intro subdegree_geI) (auto simp: fps_eq_iff split: if_split_asm)
+    thus ?thesis ..
+  qed simp
+qed (auto simp: fps_eq_iff intro: nth_less_subdegree_zero)
+
+lemma fps_cutoff_0 [simp]: "fps_cutoff 0 f = 0"
+  by (simp add: fps_eq_iff)
+
+lemma fps_cutoff_zero [simp]: "fps_cutoff n 0 = 0"
+  by (simp add: fps_eq_iff)
+
+lemma fps_cutoff_one: "fps_cutoff n 1 = (if n = 0 then 0 else 1)"
+  by (simp add: fps_eq_iff)
+
+lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)"
+  by (simp add: fps_eq_iff)
+
+lemma fps_cutoff_numeral: "fps_cutoff n (numeral c) = (if n = 0 then 0 else numeral c)"
+  by (simp add: numeral_fps_const fps_cutoff_fps_const)
+
+lemma fps_shift_cutoff:
+  "fps_shift n (f :: ('a :: comm_ring_1) fps) * X^n + fps_cutoff n f = f"
+  by (simp add: fps_eq_iff X_power_mult_right_nth)
+
+
+subsection \<open>Formal Power series form a metric space\<close>
+
+definition (in dist) "ball x r = {y. dist y x < r}"
+
+instantiation fps :: (comm_ring_1) dist
+begin
+
+definition
+  dist_fps_def: "dist (a :: 'a fps) b = (if a = b then 0 else inverse (2 ^ subdegree (a - b)))"
+
+lemma dist_fps_ge0: "dist (a :: 'a fps) b \<ge> 0"
+  by (simp add: dist_fps_def)
+
+lemma dist_fps_sym: "dist (a :: 'a fps) b = dist b a"
+  by (simp add: dist_fps_def)
+
+instance ..
+
+end
+
+instantiation fps :: (comm_ring_1) metric_space
+begin
+
+definition uniformity_fps_def [code del]:
+  "(uniformity :: ('a fps \<times> 'a fps) filter) = (INF e:{0 <..}. principal {(x, y). dist x y < e})"
+
+definition open_fps_def' [code del]:
+  "open (U :: 'a fps set) \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)"
+
+instance
+proof
+  show th: "dist a b = 0 \<longleftrightarrow> a = b" for a b :: "'a fps"
+    by (simp add: dist_fps_def split: if_split_asm)
+  then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp
+
+  fix a b c :: "'a fps"
+  consider "a = b" | "c = a \<or> c = b" | "a \<noteq> b" "a \<noteq> c" "b \<noteq> c" by blast
+  then show "dist a b \<le> dist a c + dist b c"
+  proof cases
+    case 1
+    then show ?thesis by (simp add: dist_fps_def)
+  next
+    case 2
+    then show ?thesis
+      by (cases "c = a") (simp_all add: th dist_fps_sym)
+  next
+    case neq: 3
+    have False if "dist a b > dist a c + dist b c"
+    proof -
+      let ?n = "subdegree (a - b)"
+      from neq have "dist a b > 0" "dist b c > 0" and "dist a c > 0" by (simp_all add: dist_fps_def)
+      with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all
+      with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)"
+        by (simp_all add: dist_fps_def field_simps)
+      hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0"
+        by (simp_all only: nth_less_subdegree_zero)
+      hence "(a - b) $ ?n = 0" by simp
+      moreover from neq have "(a - b) $ ?n \<noteq> 0" by (intro nth_subdegree_nonzero) simp_all
+      ultimately show False by contradiction
+    qed
+    thus ?thesis by (auto simp add: not_le[symmetric])
+  qed
+qed (rule open_fps_def' uniformity_fps_def)+
+
+end
+
+declare uniformity_Abort[where 'a="'a :: comm_ring_1 fps", code]
+
+lemma open_fps_def: "open (S :: 'a::comm_ring_1 fps set) = (\<forall>a \<in> S. \<exists>r. r >0 \<and> ball a r \<subseteq> S)"
+  unfolding open_dist ball_def subset_eq by simp
+
+text \<open>The infinite sums and justification of the notation in textbooks.\<close>
+
+lemma reals_power_lt_ex:
+  fixes x y :: real
+  assumes xp: "x > 0"
+    and y1: "y > 1"
+  shows "\<exists>k>0. (1/y)^k < x"
+proof -
+  have yp: "y > 0"
+    using y1 by simp
+  from reals_Archimedean2[of "max 0 (- log y x) + 1"]
+  obtain k :: nat where k: "real k > max 0 (- log y x) + 1"
+    by blast
+  from k have kp: "k > 0"
+    by simp
+  from k have "real k > - log y x"
+    by simp
+  then have "ln y * real k > - ln x"
+    unfolding log_def
+    using ln_gt_zero_iff[OF yp] y1
+    by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric])
+  then have "ln y * real k + ln x > 0"
+    by simp
+  then have "exp (real k * ln y + ln x) > exp 0"
+    by (simp add: ac_simps)
+  then have "y ^ k * x > 1"
+    unfolding exp_zero exp_add exp_real_of_nat_mult exp_ln [OF xp] exp_ln [OF yp]
+    by simp
+  then have "x > (1 / y)^k" using yp
+    by (simp add: field_simps)
+  then show ?thesis
+    using kp by blast
+qed
+
+lemma fps_sum_rep_nth: "(sum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
+    (if n \<le> m then a$n else 0::'a::comm_ring_1)"
+  apply (auto simp add: fps_sum_nth cond_value_iff cong del: if_weak_cong)
+  apply (simp add: sum.delta')
+  done
+
+lemma fps_notation: "(\<lambda>n. sum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) \<longlonglongrightarrow> a"
+  (is "?s \<longlonglongrightarrow> a")
+proof -
+  have "\<exists>n0. \<forall>n \<ge> n0. dist (?s n) a < r" if "r > 0" for r
+  proof -
+    obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0"
+      using reals_power_lt_ex[OF \<open>r > 0\<close>, of 2] by auto
+    show ?thesis
+    proof -
+      have "dist (?s n) a < r" if nn0: "n \<ge> n0" for n
+      proof -
+        from that have thnn0: "(1/2)^n \<le> (1/2 :: real)^n0"
+          by (simp add: divide_simps)
+        show ?thesis
+        proof (cases "?s n = a")
+          case True
+          then show ?thesis
+            unfolding dist_eq_0_iff[of "?s n" a, symmetric]
+            using \<open>r > 0\<close> by (simp del: dist_eq_0_iff)
+        next
+          case False
+          from False have dth: "dist (?s n) a = (1/2)^subdegree (?s n - a)"
+            by (simp add: dist_fps_def field_simps)
+          from False have kn: "subdegree (?s n - a) > n"
+            by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth)
+          then have "dist (?s n) a < (1/2)^n"
+            by (simp add: field_simps dist_fps_def)
+          also have "\<dots> \<le> (1/2)^n0"
+            using nn0 by (simp add: divide_simps)
+          also have "\<dots> < r"
+            using n0 by simp
+          finally show ?thesis .
+        qed
+      qed
+      then show ?thesis by blast
+    qed
+  qed
+  then show ?thesis
+    unfolding lim_sequentially by blast
+qed
+
+
+subsection \<open>Inverses of formal power series\<close>
+
+declare sum.cong[fundef_cong]
+
+instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse
+begin
+
+fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
+where
+  "natfun_inverse f 0 = inverse (f$0)"
+| "natfun_inverse f n = - inverse (f$0) * sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
+
+definition fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))"
+
+definition fps_divide_def:
+  "f div g = (if g = 0 then 0 else
+     let n = subdegree g; h = fps_shift n g
+     in  fps_shift n (f * inverse h))"
+
+instance ..
+
+end
+
+lemma fps_inverse_zero [simp]:
+  "inverse (0 :: 'a::{comm_monoid_add,inverse,times,uminus} fps) = 0"
+  by (simp add: fps_ext fps_inverse_def)
+
+lemma fps_inverse_one [simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1"
+  apply (auto simp add: expand_fps_eq fps_inverse_def)
+  apply (case_tac n)
+  apply auto
+  done
+
+lemma inverse_mult_eq_1 [intro]:
+  assumes f0: "f$0 \<noteq> (0::'a::field)"
+  shows "inverse f * f = 1"
+proof -
+  have c: "inverse f * f = f * inverse f"
+    by (simp add: mult.commute)
+  from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n"
+    by (simp add: fps_inverse_def)
+  from f0 have th0: "(inverse f * f) $ 0 = 1"
+    by (simp add: fps_mult_nth fps_inverse_def)
+  have "(inverse f * f)$n = 0" if np: "n > 0" for n
+  proof -
+    from np have eq: "{0..n} = {0} \<union> {1 .. n}"
+      by auto
+    have d: "{0} \<inter> {1 .. n} = {}"
+      by auto
+    from f0 np have th0: "- (inverse f $ n) =
+      (sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
+      by (cases n) (simp_all add: divide_inverse fps_inverse_def)
+    from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
+    have th1: "sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n"
+      by (simp add: field_simps)
+    have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
+      unfolding fps_mult_nth ifn ..
+    also have "\<dots> = f$0 * natfun_inverse f n + (\<Sum>i = 1..n. f$i * natfun_inverse f (n-i))"
+      by (simp add: eq)
+    also have "\<dots> = 0"
+      unfolding th1 ifn by simp
+    finally show ?thesis unfolding c .
+  qed
+  with th0 show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+lemma fps_inverse_0_iff[simp]: "(inverse f) $ 0 = (0::'a::division_ring) \<longleftrightarrow> f $ 0 = 0"
+  by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero)
+
+lemma fps_inverse_nth_0 [simp]: "inverse f $ 0 = inverse (f $ 0 :: 'a :: division_ring)"
+  by (simp add: fps_inverse_def)
+
+lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::division_ring) fps) \<longleftrightarrow> f $ 0 = 0"
+proof
+  assume A: "inverse f = 0"
+  have "0 = inverse f $ 0" by (subst A) simp
+  thus "f $ 0 = 0" by simp
+qed (simp add: fps_inverse_def)
+
+lemma fps_inverse_idempotent[intro, simp]:
+  assumes f0: "f$0 \<noteq> (0::'a::field)"
+  shows "inverse (inverse f) = f"
+proof -
+  from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
+  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
+  have "inverse f * f = inverse f * inverse (inverse f)"
+    by (simp add: ac_simps)
+  then show ?thesis
+    using f0 unfolding mult_cancel_left by simp
+qed
+
+lemma fps_inverse_unique:
+  assumes fg: "(f :: 'a :: field fps) * g = 1"
+  shows   "inverse f = g"
+proof -
+  have f0: "f $ 0 \<noteq> 0"
+  proof
+    assume "f $ 0 = 0"
+    hence "0 = (f * g) $ 0" by simp
+    also from fg have "(f * g) $ 0 = 1" by simp
+    finally show False by simp
+  qed
+  from inverse_mult_eq_1[OF this] fg
+  have th0: "inverse f * f = g * f"
+    by (simp add: ac_simps)
+  then show ?thesis
+    using f0
+    unfolding mult_cancel_right
+    by (auto simp add: expand_fps_eq)
+qed
+
+lemma fps_inverse_eq_0: "f$0 = 0 \<Longrightarrow> inverse (f :: 'a :: division_ring fps) = 0"
+  by simp
+  
+lemma sum_zero_lemma:
+  fixes n::nat
+  assumes "0 < n"
+  shows "(\<Sum>i = 0..n. if n = i then 1 else if n - i = 1 then - 1 else 0) = (0::'a::field)"
+proof -
+  let ?f = "\<lambda>i. if n = i then 1 else if n - i = 1 then - 1 else 0"
+  let ?g = "\<lambda>i. if i = n then 1 else if i = n - 1 then - 1 else 0"
+  let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
+  have th1: "sum ?f {0..n} = sum ?g {0..n}"
+    by (rule sum.cong) auto
+  have th2: "sum ?g {0..n - 1} = sum ?h {0..n - 1}"
+    apply (rule sum.cong)
+    using assms
+    apply auto
+    done
+  have eq: "{0 .. n} = {0.. n - 1} \<union> {n}"
+    by auto
+  from assms have d: "{0.. n - 1} \<inter> {n} = {}"
+    by auto
+  have f: "finite {0.. n - 1}" "finite {n}"
+    by auto
+  show ?thesis
+    unfolding th1
+    apply (simp add: sum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
+    unfolding th2
+    apply (simp add: sum.delta)
+    done
+qed
+
+lemma fps_inverse_mult: "inverse (f * g :: 'a::field fps) = inverse f * inverse g"
+proof (cases "f$0 = 0 \<or> g$0 = 0")
+  assume "\<not>(f$0 = 0 \<or> g$0 = 0)"
+  hence [simp]: "f$0 \<noteq> 0" "g$0 \<noteq> 0" by simp_all
+  show ?thesis
+  proof (rule fps_inverse_unique)
+    have "f * g * (inverse f * inverse g) = (inverse f * f) * (inverse g * g)" by simp
+    also have "... = 1" by (subst (1 2) inverse_mult_eq_1) simp_all
+    finally show "f * g * (inverse f * inverse g) = 1" .
+  qed
+next
+  assume A: "f$0 = 0 \<or> g$0 = 0"
+  hence "inverse (f * g) = 0" by simp
+  also from A have "... = inverse f * inverse g" by auto
+  finally show "inverse (f * g) = inverse f * inverse g" .
+qed
+
+
+lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) =
+    Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
+  apply (rule fps_inverse_unique)
+  apply (simp_all add: fps_eq_iff fps_mult_nth sum_zero_lemma)
+  done
+
+lemma subdegree_inverse [simp]: "subdegree (inverse (f::'a::field fps)) = 0"
+proof (cases "f$0 = 0")
+  assume nz: "f$0 \<noteq> 0"
+  hence "subdegree (inverse f) + subdegree f = subdegree (inverse f * f)"
+    by (subst subdegree_mult) auto
+  also from nz have "subdegree f = 0" by (simp add: subdegree_eq_0_iff)
+  also from nz have "inverse f * f = 1" by (rule inverse_mult_eq_1)
+  finally show "subdegree (inverse f) = 0" by simp
+qed (simp_all add: fps_inverse_def)
+
+lemma fps_is_unit_iff [simp]: "(f :: 'a :: field fps) dvd 1 \<longleftrightarrow> f $ 0 \<noteq> 0"
+proof
+  assume "f dvd 1"
+  then obtain g where "1 = f * g" by (elim dvdE)
+  from this[symmetric] have "(f*g) $ 0 = 1" by simp
+  thus "f $ 0 \<noteq> 0" by auto
+next
+  assume A: "f $ 0 \<noteq> 0"
+  thus "f dvd 1" by (simp add: inverse_mult_eq_1[OF A, symmetric])
+qed
+
+lemma subdegree_eq_0' [simp]: "(f :: 'a :: field fps) dvd 1 \<Longrightarrow> subdegree f = 0"
+  by simp
+
+lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \<noteq> 0 \<Longrightarrow> f dvd g"
+  by (rule dvd_trans, subst fps_is_unit_iff) simp_all
+
+instantiation fps :: (field) normalization_semidom
+begin
+
+definition fps_unit_factor_def [simp]:
+  "unit_factor f = fps_shift (subdegree f) f"
+
+definition fps_normalize_def [simp]:
+  "normalize f = (if f = 0 then 0 else X ^ subdegree f)"
+
+instance proof
+  fix f :: "'a fps"
+  show "unit_factor f * normalize f = f"
+    by (simp add: fps_shift_times_X_power)
+next
+  fix f g :: "'a fps"
+  show "unit_factor (f * g) = unit_factor f * unit_factor g"
+  proof (cases "f = 0 \<or> g = 0")
+    assume "\<not>(f = 0 \<or> g = 0)"
+    thus "unit_factor (f * g) = unit_factor f * unit_factor g"
+    unfolding fps_unit_factor_def
+      by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
+  qed auto
+next
+  fix f g :: "'a fps"
+  assume "g \<noteq> 0"
+  then have "f * (fps_shift (subdegree g) g * inverse (fps_shift (subdegree g) g)) = f"
+    by (metis add_cancel_right_left fps_shift_nth inverse_mult_eq_1 mult.commute mult_cancel_left2 nth_subdegree_nonzero)
+  then have "fps_shift (subdegree g) (g * (f * inverse (fps_shift (subdegree g) g))) = f"
+    by (simp add: fps_shift_mult_right mult.commute)
+  with \<open>g \<noteq> 0\<close> show "f * g / g = f"
+    by (simp add: fps_divide_def Let_def ac_simps)
+qed (auto simp add: fps_divide_def Let_def)
+
+end
+
+instantiation fps :: (field) ring_div
+begin
+
+definition fps_mod_def:
+  "f mod g = (if g = 0 then f else
+     let n = subdegree g; h = fps_shift n g
+     in  fps_cutoff n (f * inverse h) * h)"
+
+lemma fps_mod_eq_zero:
+  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree g"
+  shows   "f mod g = 0"
+  using assms by (cases "f = 0") (auto simp: fps_cutoff_zero_iff fps_mod_def Let_def)
+
+lemma fps_times_divide_eq:
+  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree (g :: 'a fps)"
+  shows   "f div g * g = f"
+proof (cases "f = 0")
+  assume nz: "f \<noteq> 0"
+  define n where "n = subdegree g"
+  define h where "h = fps_shift n g"
+  from assms have [simp]: "h $ 0 \<noteq> 0" unfolding h_def by (simp add: n_def)
+
+  from assms nz have "f div g * g = fps_shift n (f * inverse h) * g"
+    by (simp add: fps_divide_def Let_def h_def n_def)
+  also have "... = fps_shift n (f * inverse h) * X^n * h" unfolding h_def n_def
+    by (subst subdegree_decompose[of g]) simp
+  also have "fps_shift n (f * inverse h) * X^n = f * inverse h"
+    by (rule fps_shift_times_X_power) (simp_all add: nz assms n_def)
+  also have "... * h = f * (inverse h * h)" by simp
+  also have "inverse h * h = 1" by (rule inverse_mult_eq_1) simp
+  finally show ?thesis by simp
+qed (simp_all add: fps_divide_def Let_def)
+
+lemma
+  assumes "g$0 \<noteq> 0"
+  shows   fps_divide_unit: "f div g = f * inverse g" and fps_mod_unit [simp]: "f mod g = 0"
+proof -
+  from assms have [simp]: "subdegree g = 0" by (simp add: subdegree_eq_0_iff)
+  from assms show "f div g = f * inverse g"
+    by (auto simp: fps_divide_def Let_def subdegree_eq_0_iff)
+  from assms show "f mod g = 0" by (intro fps_mod_eq_zero) auto
+qed
+
+context
+begin
+private lemma fps_divide_cancel_aux1:
+  assumes "h$0 \<noteq> (0 :: 'a :: field)"
+  shows   "(h * f) div (h * g) = f div g"
+proof (cases "g = 0")
+  assume "g \<noteq> 0"
+  from assms have "h \<noteq> 0" by auto
+  note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
+  from assms have [simp]: "subdegree h = 0" by (simp add: subdegree_eq_0_iff)
+
+  have "(h * f) div (h * g) =
+          fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
+    by (simp add: fps_divide_def Let_def)
+  also have "h * f * inverse (fps_shift (subdegree g) (h*g)) =
+               (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
+    by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
+  also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
+  finally show "(h * f) div (h * g) = f div g" by (simp_all add: fps_divide_def Let_def)
+qed (simp_all add: fps_divide_def)
+
+private lemma fps_divide_cancel_aux2:
+  "(f * X^m) div (g * X^m) = f div (g :: 'a :: field fps)"
+proof (cases "g = 0")
+  assume [simp]: "g \<noteq> 0"
+  have "(f * X^m) div (g * X^m) =
+          fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*X^m))*X^m)"
+    by (simp add: fps_divide_def Let_def algebra_simps)
+  also have "... = f div g"
+    by (simp add: fps_shift_times_X_power'' fps_divide_def Let_def)
+  finally show ?thesis .
+qed (simp_all add: fps_divide_def)
+
+instance proof
+  fix f g :: "'a fps"
+  define n where "n = subdegree g"
+  define h where "h = fps_shift n g"
+
+  show "f div g * g + f mod g = f"
+  proof (cases "g = 0 \<or> f = 0")
+    assume "\<not>(g = 0 \<or> f = 0)"
+    hence nz [simp]: "f \<noteq> 0" "g \<noteq> 0" by simp_all
+    show ?thesis
+    proof (rule disjE[OF le_less_linear])
+      assume "subdegree f \<ge> subdegree g"
+      with nz show ?thesis by (simp add: fps_mod_eq_zero fps_times_divide_eq)
+    next
+      assume "subdegree f < subdegree g"
+      have g_decomp: "g = h * X^n" unfolding h_def n_def by (rule subdegree_decompose)
+      have "f div g * g + f mod g =
+              fps_shift n (f * inverse h) * g + fps_cutoff n (f * inverse h) * h"
+        by (simp add: fps_mod_def fps_divide_def Let_def n_def h_def)
+      also have "... = h * (fps_shift n (f * inverse h) * X^n + fps_cutoff n (f * inverse h))"
+        by (subst g_decomp) (simp add: algebra_simps)
+      also have "... = f * (inverse h * h)"
+        by (subst fps_shift_cutoff) simp
+      also have "inverse h * h = 1" by (rule inverse_mult_eq_1) (simp add: h_def n_def)
+      finally show ?thesis by simp
+    qed
+  qed (auto simp: fps_mod_def fps_divide_def Let_def)
+next
+
+  fix f g h :: "'a fps"
+  assume "h \<noteq> 0"
+  show "(h * f) div (h * g) = f div g"
+  proof -
+    define m where "m = subdegree h"
+    define h' where "h' = fps_shift m h"
+    have h_decomp: "h = h' * X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
+    from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
+    have "(h * f) div (h * g) = (h' * f * X^m) div (h' * g * X^m)"
+      by (simp add: h_decomp algebra_simps)
+    also have "... = f div g" by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
+    finally show ?thesis .
+  qed
+
+next
+  fix f g h :: "'a fps"
+  assume [simp]: "h \<noteq> 0"
+  define n h' where dfs: "n = subdegree h" "h' = fps_shift n h"
+  have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
+    by (simp add: fps_divide_def Let_def dfs[symmetric] algebra_simps fps_shift_add)
+  also have "h * inverse h' = (inverse h' * h') * X^n"
+    by (subst subdegree_decompose) (simp_all add: dfs)
+  also have "... = X^n" by (subst inverse_mult_eq_1) (simp_all add: dfs)
+  also have "fps_shift n (g * X^n) = g" by simp
+  also have "fps_shift n (f * inverse h') = f div h"
+    by (simp add: fps_divide_def Let_def dfs)
+  finally show "(f + g * h) div h = g + f div h" by simp
+qed
+
+end
+end
+
+lemma subdegree_mod:
+  assumes "f \<noteq> 0" "subdegree f < subdegree g"
+  shows   "subdegree (f mod g) = subdegree f"
+proof (cases "f div g * g = 0")
+  assume "f div g * g \<noteq> 0"
+  hence [simp]: "f div g \<noteq> 0" "g \<noteq> 0" by auto
+  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
+  also from assms have "subdegree ... = subdegree f"
+    by (intro subdegree_diff_eq1) simp_all
+  finally show ?thesis .
+next
+  assume zero: "f div g * g = 0"
+  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
+  also note zero
+  finally show ?thesis by simp
+qed
+
+lemma fps_divide_nth_0 [simp]: "g $ 0 \<noteq> 0 \<Longrightarrow> (f div g) $ 0 = f $ 0 / (g $ 0 :: _ :: field)"
+  by (simp add: fps_divide_unit divide_inverse)
+
+
+lemma dvd_imp_subdegree_le:
+  "(f :: 'a :: idom fps) dvd g \<Longrightarrow> g \<noteq> 0 \<Longrightarrow> subdegree f \<le> subdegree g"
+  by (auto elim: dvdE)
+
+lemma fps_dvd_iff:
+  assumes "(f :: 'a :: field fps) \<noteq> 0" "g \<noteq> 0"
+  shows   "f dvd g \<longleftrightarrow> subdegree f \<le> subdegree g"
+proof
+  assume "subdegree f \<le> subdegree g"
+  with assms have "g mod f = 0"
+    by (simp add: fps_mod_def Let_def fps_cutoff_zero_iff)
+  thus "f dvd g" by (simp add: dvd_eq_mod_eq_0)
+qed (simp add: assms dvd_imp_subdegree_le)
+
+lemma fps_shift_altdef:
+  "fps_shift n f = (f :: 'a :: field fps) div X^n"
+  by (simp add: fps_divide_def)
+  
+lemma fps_div_X_power_nth: "((f :: 'a :: field fps) div X^n) $ k = f $ (k + n)"
+  by (simp add: fps_shift_altdef [symmetric])
+
+lemma fps_div_X_nth: "((f :: 'a :: field fps) div X) $ k = f $ Suc k"
+  using fps_div_X_power_nth[of f 1] by simp
+
+lemma fps_const_inverse: "inverse (fps_const (a::'a::field)) = fps_const (inverse a)"
+  by (cases "a \<noteq> 0", rule fps_inverse_unique) (auto simp: fps_eq_iff)
+
+lemma fps_const_divide: "fps_const (x :: _ :: field) / fps_const y = fps_const (x / y)"
+  by (cases "y = 0") (simp_all add: fps_divide_unit fps_const_inverse divide_inverse)
+
+lemma inverse_fps_numeral:
+  "inverse (numeral n :: ('a :: field_char_0) fps) = fps_const (inverse (numeral n))"
+  by (intro fps_inverse_unique fps_ext) (simp_all add: fps_numeral_nth)
+
+lemma fps_numeral_divide_divide:
+  "x / numeral b / numeral c = (x / numeral (b * c) :: 'a :: field fps)"
+  by (cases "numeral b = (0::'a)"; cases "numeral c = (0::'a)")
+      (simp_all add: fps_divide_unit fps_inverse_mult [symmetric] numeral_fps_const numeral_mult 
+                del: numeral_mult [symmetric])
+
+lemma fps_numeral_mult_divide:
+  "numeral b * x / numeral c = (numeral b / numeral c * x :: 'a :: field fps)"
+  by (cases "numeral c = (0::'a)") (simp_all add: fps_divide_unit numeral_fps_const)
+
+lemmas fps_numeral_simps = 
+  fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const
+
+
+subsection \<open>Formal power series form a Euclidean ring\<close>
+
+instantiation fps :: (field) euclidean_ring_cancel
+begin
+
+definition fps_euclidean_size_def:
+  "euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)"
+
+instance proof
+  fix f g :: "'a fps" assume [simp]: "g \<noteq> 0"
+  show "euclidean_size f \<le> euclidean_size (f * g)"
+    by (cases "f = 0") (auto simp: fps_euclidean_size_def)
+  show "euclidean_size (f mod g) < euclidean_size g"
+    apply (cases "f = 0", simp add: fps_euclidean_size_def)
+    apply (rule disjE[OF le_less_linear[of "subdegree g" "subdegree f"]])
+    apply (simp_all add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod)
+    done
+qed (simp_all add: fps_euclidean_size_def)
+
+end
+
+instantiation fps :: (field) euclidean_ring_gcd
+begin
+definition fps_gcd_def: "(gcd :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.gcd"
+definition fps_lcm_def: "(lcm :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.lcm"
+definition fps_Gcd_def: "(Gcd :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Gcd"
+definition fps_Lcm_def: "(Lcm :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Lcm"
+instance by standard (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def)
+end
+
+lemma fps_gcd:
+  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
+  shows   "gcd f g = X ^ min (subdegree f) (subdegree g)"
+proof -
+  let ?m = "min (subdegree f) (subdegree g)"
+  show "gcd f g = X ^ ?m"
+  proof (rule sym, rule gcdI)
+    fix d assume "d dvd f" "d dvd g"
+    thus "d dvd X ^ ?m" by (cases "d = 0") (auto simp: fps_dvd_iff)
+  qed (simp_all add: fps_dvd_iff)
+qed
+
+lemma fps_gcd_altdef: "gcd (f :: 'a :: field fps) g =
+  (if f = 0 \<and> g = 0 then 0 else
+   if f = 0 then X ^ subdegree g else
+   if g = 0 then X ^ subdegree f else
+     X ^ min (subdegree f) (subdegree g))"
+  by (simp add: fps_gcd)
+
+lemma fps_lcm:
+  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
+  shows   "lcm f g = X ^ max (subdegree f) (subdegree g)"
+proof -
+  let ?m = "max (subdegree f) (subdegree g)"
+  show "lcm f g = X ^ ?m"
+  proof (rule sym, rule lcmI)
+    fix d assume "f dvd d" "g dvd d"
+    thus "X ^ ?m dvd d" by (cases "d = 0") (auto simp: fps_dvd_iff)
+  qed (simp_all add: fps_dvd_iff)
+qed
+
+lemma fps_lcm_altdef: "lcm (f :: 'a :: field fps) g =
+  (if f = 0 \<or> g = 0 then 0 else X ^ max (subdegree f) (subdegree g))"
+  by (simp add: fps_lcm)
+
+lemma fps_Gcd:
+  assumes "A - {0} \<noteq> {}"
+  shows   "Gcd A = X ^ (INF f:A-{0}. subdegree f)"
+proof (rule sym, rule GcdI)
+  fix f assume "f \<in> A"
+  thus "X ^ (INF f:A - {0}. subdegree f) dvd f"
+    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cINF_lower)
+next
+  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> d dvd f"
+  from assms obtain f where "f \<in> A - {0}" by auto
+  with d[of f] have [simp]: "d \<noteq> 0" by auto
+  from d assms have "subdegree d \<le> (INF f:A-{0}. subdegree f)"
+    by (intro cINF_greatest) (auto simp: fps_dvd_iff[symmetric])
+  with d assms show "d dvd X ^ (INF f:A-{0}. subdegree f)" by (simp add: fps_dvd_iff)
+qed simp_all
+
+lemma fps_Gcd_altdef: "Gcd (A :: 'a :: field fps set) =
+  (if A \<subseteq> {0} then 0 else X ^ (INF f:A-{0}. subdegree f))"
+  using fps_Gcd by auto
+
+lemma fps_Lcm:
+  assumes "A \<noteq> {}" "0 \<notin> A" "bdd_above (subdegree`A)"
+  shows   "Lcm A = X ^ (SUP f:A. subdegree f)"
+proof (rule sym, rule LcmI)
+  fix f assume "f \<in> A"
+  moreover from assms(3) have "bdd_above (subdegree ` A)" by auto
+  ultimately show "f dvd X ^ (SUP f:A. subdegree f)" using assms(2)
+    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cSUP_upper)
+next
+  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> f dvd d"
+  from assms obtain f where f: "f \<in> A" "f \<noteq> 0" by auto
+  show "X ^ (SUP f:A. subdegree f) dvd d"
+  proof (cases "d = 0")
+    assume "d \<noteq> 0"
+    moreover from d have "\<And>f. f \<in> A \<Longrightarrow> f \<noteq> 0 \<Longrightarrow> f dvd d" by blast
+    ultimately have "subdegree d \<ge> (SUP f:A. subdegree f)" using assms
+      by (intro cSUP_least) (auto simp: fps_dvd_iff)
+    with \<open>d \<noteq> 0\<close> show ?thesis by (simp add: fps_dvd_iff)
+  qed simp_all
+qed simp_all
+
+lemma fps_Lcm_altdef:
+  "Lcm (A :: 'a :: field fps set) =
+     (if 0 \<in> A \<or> \<not>bdd_above (subdegree`A) then 0 else
+      if A = {} then 1 else X ^ (SUP f:A. subdegree f))"
+proof (cases "bdd_above (subdegree`A)")
+  assume unbounded: "\<not>bdd_above (subdegree`A)"
+  have "Lcm A = 0"
+  proof (rule ccontr)
+    assume "Lcm A \<noteq> 0"
+    from unbounded obtain f where f: "f \<in> A" "subdegree (Lcm A) < subdegree f"
+      unfolding bdd_above_def by (auto simp: not_le)
+    moreover from f and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
+      by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all
+    ultimately show False by simp
+  qed
+  with unbounded show ?thesis by simp
+qed (simp_all add: fps_Lcm Lcm_eq_0_I)
+
+
+
+subsection \<open>Formal Derivatives, and the MacLaurin theorem around 0\<close>
+
+definition "fps_deriv f = Abs_fps (\<lambda>n. of_nat (n + 1) * f $ (n + 1))"
+
+lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n + 1)"
+  by (simp add: fps_deriv_def)
+
+lemma fps_0th_higher_deriv: 
+  "(fps_deriv ^^ n) f $ 0 = (fact n * f $ n :: 'a :: {comm_ring_1, semiring_char_0})"
+  by (induction n arbitrary: f) (simp_all del: funpow.simps add: funpow_Suc_right algebra_simps)
+
+lemma fps_deriv_linear[simp]:
+  "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
+    fps_const a * fps_deriv f + fps_const b * fps_deriv g"
+  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
+
+lemma fps_deriv_mult[simp]:
+  fixes f :: "'a::comm_ring_1 fps"
+  shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
+proof -
+  let ?D = "fps_deriv"
+  have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" for n
+  proof -
+    let ?Zn = "{0 ..n}"
+    let ?Zn1 = "{0 .. n + 1}"
+    let ?g = "\<lambda>i. of_nat (i+1) * g $ (i+1) * f $ (n - i) +
+        of_nat (i+1)* f $ (i+1) * g $ (n - i)"
+    let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
+        of_nat i* f $ i * g $ ((n + 1) - i)"
+    have s0: "sum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
+      sum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
+       by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
+    have s1: "sum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
+      sum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
+       by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
+    have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
+      by (simp only: mult.commute)
+    also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
+      by (simp add: fps_mult_nth sum.distrib[symmetric])
+    also have "\<dots> = sum ?h {0..n+1}"
+      by (rule sum.reindex_bij_witness_not_neutral
+            [where S'="{}" and T'="{0}" and j="Suc" and i="\<lambda>i. i - 1"]) auto
+    also have "\<dots> = (fps_deriv (f * g)) $ n"
+      apply (simp only: fps_deriv_nth fps_mult_nth sum.distrib)
+      unfolding s0 s1
+      unfolding sum.distrib[symmetric] sum_distrib_left
+      apply (rule sum.cong)
+      apply (auto simp add: of_nat_diff field_simps)
+      done
+    finally show ?thesis .
+  qed
+  then show ?thesis
+    unfolding fps_eq_iff by auto
+qed
+
+lemma fps_deriv_X[simp]: "fps_deriv X = 1"
+  by (simp add: fps_deriv_def X_def fps_eq_iff)
+
+lemma fps_deriv_neg[simp]:
+  "fps_deriv (- (f:: 'a::comm_ring_1 fps)) = - (fps_deriv f)"
+  by (simp add: fps_eq_iff fps_deriv_def)
+
+lemma fps_deriv_add[simp]:
+  "fps_deriv ((f:: 'a::comm_ring_1 fps) + g) = fps_deriv f + fps_deriv g"
+  using fps_deriv_linear[of 1 f 1 g] by simp
+
+lemma fps_deriv_sub[simp]:
+  "fps_deriv ((f:: 'a::comm_ring_1 fps) - g) = fps_deriv f - fps_deriv g"
+  using fps_deriv_add [of f "- g"] by simp
+
+lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
+  by (simp add: fps_ext fps_deriv_def fps_const_def)
+
+lemma fps_deriv_of_nat [simp]: "fps_deriv (of_nat n) = 0"
+  by (simp add: fps_of_nat [symmetric])
+
+lemma fps_deriv_numeral [simp]: "fps_deriv (numeral n) = 0"
+  by (simp add: numeral_fps_const)    
+
+lemma fps_deriv_mult_const_left[simp]:
+  "fps_deriv (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_deriv f"
+  by simp
+
+lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
+  by (simp add: fps_deriv_def fps_eq_iff)
+
+lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
+  by (simp add: fps_deriv_def fps_eq_iff )
+
+lemma fps_deriv_mult_const_right[simp]:
+  "fps_deriv (f * fps_const (c::'a::comm_ring_1)) = fps_deriv f * fps_const c"
+  by simp
+
+lemma fps_deriv_sum:
+  "fps_deriv (sum f S) = sum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
+proof (cases "finite S")
+  case False
+  then show ?thesis by simp
+next
+  case True
+  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
+qed
+
+lemma fps_deriv_eq_0_iff [simp]:
+  "fps_deriv f = 0 \<longleftrightarrow> f = fps_const (f$0 :: 'a::{idom,semiring_char_0})"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show ?lhs if ?rhs
+  proof -
+    from that have "fps_deriv f = fps_deriv (fps_const (f$0))"
+      by simp
+    then show ?thesis
+      by simp
+  qed
+  show ?rhs if ?lhs
+  proof -
+    from that have "\<forall>n. (fps_deriv f)$n = 0"
+      by simp
+    then have "\<forall>n. f$(n+1) = 0"
+      by (simp del: of_nat_Suc of_nat_add One_nat_def)
+    then show ?thesis
+      apply (clarsimp simp add: fps_eq_iff fps_const_def)
+      apply (erule_tac x="n - 1" in allE)
+      apply simp
+      done
+  qed
+qed
+
+lemma fps_deriv_eq_iff:
+  fixes f :: "'a::{idom,semiring_char_0} fps"
+  shows "fps_deriv f = fps_deriv g \<longleftrightarrow> (f = fps_const(f$0 - g$0) + g)"
+proof -
+  have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0"
+    by simp
+  also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f - g) $ 0)"
+    unfolding fps_deriv_eq_0_iff ..
+  finally show ?thesis
+    by (simp add: field_simps)
+qed
+
+lemma fps_deriv_eq_iff_ex:
+  "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>c::'a::{idom,semiring_char_0}. f = fps_const c + g)"
+  by (auto simp: fps_deriv_eq_iff)
+
+
+fun fps_nth_deriv :: "nat \<Rightarrow> 'a::semiring_1 fps \<Rightarrow> 'a fps"
+where
+  "fps_nth_deriv 0 f = f"
+| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
+
+lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)"
+  by (induct n arbitrary: f) auto
+
+lemma fps_nth_deriv_linear[simp]:
+  "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
+    fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
+  by (induct n arbitrary: f g) (auto simp add: fps_nth_deriv_commute)
+
+lemma fps_nth_deriv_neg[simp]:
+  "fps_nth_deriv n (- (f :: 'a::comm_ring_1 fps)) = - (fps_nth_deriv n f)"
+  by (induct n arbitrary: f) simp_all
+
+lemma fps_nth_deriv_add[simp]:
+  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g"
+  using fps_nth_deriv_linear[of n 1 f 1 g] by simp
+
+lemma fps_nth_deriv_sub[simp]:
+  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
+  using fps_nth_deriv_add [of n f "- g"] by simp
+
+lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
+  by (induct n) simp_all
+
+lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)"
+  by (induct n) simp_all
+
+lemma fps_nth_deriv_const[simp]:
+  "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
+  by (cases n) simp_all
+
+lemma fps_nth_deriv_mult_const_left[simp]:
+  "fps_nth_deriv n (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_nth_deriv n f"
+  using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
+
+lemma fps_nth_deriv_mult_const_right[simp]:
+  "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
+  using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult.commute)
+
+lemma fps_nth_deriv_sum:
+  "fps_nth_deriv n (sum f S) = sum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
+proof (cases "finite S")
+  case True
+  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
+next
+  case False
+  then show ?thesis by simp
+qed
+
+lemma fps_deriv_maclauren_0:
+  "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k"
+  by (induct k arbitrary: f) (auto simp add: field_simps)
+
+
+subsection \<open>Powers\<close>
+
+lemma fps_power_zeroth_eq_one: "a$0 =1 \<Longrightarrow> a^n $ 0 = (1::'a::semiring_1)"
+  by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
+
+lemma fps_power_first_eq: "(a :: 'a::comm_ring_1 fps) $ 0 =1 \<Longrightarrow> a^n $ 1 = of_nat n * a$1"
+proof (induct n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  show ?case unfolding power_Suc fps_mult_nth
+    using Suc.hyps[OF \<open>a$0 = 1\<close>] \<open>a$0 = 1\<close> fps_power_zeroth_eq_one[OF \<open>a$0=1\<close>]
+    by (simp add: field_simps)
+qed
+
+lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
+  by (induct n) (auto simp add: fps_mult_nth)
+
+lemma startsby_zero_power:"a $0 = (0::'a::comm_ring_1) \<Longrightarrow> n > 0 \<Longrightarrow> a^n $0 = 0"
+  by (induct n) (auto simp add: fps_mult_nth)
+
+lemma startsby_power:"a $0 = (v::'a::comm_ring_1) \<Longrightarrow> a^n $0 = v^n"
+  by (induct n) (auto simp add: fps_mult_nth)
+
+lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::idom) \<longleftrightarrow> n \<noteq> 0 \<and> a$0 = 0"
+  apply (rule iffI)
+  apply (induct n)
+  apply (auto simp add: fps_mult_nth)
+  apply (rule startsby_zero_power, simp_all)
+  done
+
+lemma startsby_zero_power_prefix:
+  assumes a0: "a $ 0 = (0::'a::idom)"
+  shows "\<forall>n < k. a ^ k $ n = 0"
+  using a0
+proof (induct k rule: nat_less_induct)
+  fix k
+  assume H: "\<forall>m<k. a $0 =  0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $ 0 = 0"
+  show "\<forall>m<k. a ^ k $ m = 0"
+  proof (cases k)
+    case 0
+    then show ?thesis by simp
+  next
+    case (Suc l)
+    have "a^k $ m = 0" if mk: "m < k" for m
+    proof (cases "m = 0")
+      case True
+      then show ?thesis
+        using startsby_zero_power[of a k] Suc a0 by simp
+    next
+      case False
+      have "a ^k $ m = (a^l * a) $m"
+        by (simp add: Suc mult.commute)
+      also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
+        by (simp add: fps_mult_nth)
+      also have "\<dots> = 0"
+        apply (rule sum.neutral)
+        apply auto
+        apply (case_tac "x = m")
+        using a0 apply simp
+        apply (rule H[rule_format])
+        using a0 Suc mk apply auto
+        done
+      finally show ?thesis .
+    qed
+    then show ?thesis by blast
+  qed
+qed
+
+lemma startsby_zero_sum_depends:
+  assumes a0: "a $0 = (0::'a::idom)"
+    and kn: "n \<ge> k"
+  shows "sum (\<lambda>i. (a ^ i)$k) {0 .. n} = sum (\<lambda>i. (a ^ i)$k) {0 .. k}"
+  apply (rule sum.mono_neutral_right)
+  using kn
+  apply auto
+  apply (rule startsby_zero_power_prefix[rule_format, OF a0])
+  apply arith
+  done
+
+lemma startsby_zero_power_nth_same:
+  assumes a0: "a$0 = (0::'a::idom)"
+  shows "a^n $ n = (a$1) ^ n"
+proof (induct n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)"
+    by (simp add: field_simps)
+  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
+    by (simp add: fps_mult_nth)
+  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
+    apply (rule sum.mono_neutral_right)
+    apply simp
+    apply clarsimp
+    apply clarsimp
+    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
+    apply arith
+    done
+  also have "\<dots> = a^n $ n * a$1"
+    using a0 by simp
+  finally show ?case
+    using Suc.hyps by simp
+qed
+
+lemma fps_inverse_power:
+  fixes a :: "'a::field fps"
+  shows "inverse (a^n) = inverse a ^ n"
+  by (induction n) (simp_all add: fps_inverse_mult)
+
+lemma fps_deriv_power:
+  "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a::comm_ring_1) * fps_deriv a * a ^ (n - 1)"
+  apply (induct n)
+  apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add)
+  apply (case_tac n)
+  apply (auto simp add: field_simps)
+  done
+
+lemma fps_inverse_deriv:
+  fixes a :: "'a::field fps"
+  assumes a0: "a$0 \<noteq> 0"
+  shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
+proof -
+  from inverse_mult_eq_1[OF a0]
+  have "fps_deriv (inverse a * a) = 0" by simp
+  then have "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0"
+    by simp
+  then have "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0"
+    by simp
+  with inverse_mult_eq_1[OF a0]
+  have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) = 0"
+    unfolding power2_eq_square
+    apply (simp add: field_simps)
+    apply (simp add: mult.assoc[symmetric])
+    done
+  then have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * (inverse a)\<^sup>2 =
+      0 - fps_deriv a * (inverse a)\<^sup>2"
+    by simp
+  then show "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
+    by (simp add: field_simps)
+qed
+
+lemma fps_inverse_deriv':
+  fixes a :: "'a::field fps"
+  assumes a0: "a $ 0 \<noteq> 0"
+  shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2"
+  using fps_inverse_deriv[OF a0] a0
+  by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult)
+
+lemma inverse_mult_eq_1':
+  assumes f0: "f$0 \<noteq> (0::'a::field)"
+  shows "f * inverse f = 1"
+  by (metis mult.commute inverse_mult_eq_1 f0)
+
+lemma fps_inverse_minus [simp]: "inverse (-f) = -inverse (f :: 'a :: field fps)"
+  by (cases "f$0 = 0") (auto intro: fps_inverse_unique simp: inverse_mult_eq_1' fps_inverse_eq_0)
+  
+lemma divide_fps_const [simp]: "f / fps_const (c :: 'a :: field) = fps_const (inverse c) * f"
+  by (cases "c = 0") (simp_all add: fps_divide_unit fps_const_inverse)
+
+(* FIXME: The last part of this proof should go through by simp once we have a proper
+   theorem collection for simplifying division on rings *)
+lemma fps_divide_deriv:
+  assumes "b dvd (a :: 'a :: field fps)"
+  shows   "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b^2"
+proof -
+  have eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b div c" for a b c :: "'a :: field fps"
+    by (drule sym) (simp add: mult.assoc)
+  from assms have "a = a / b * b" by simp
+  also have "fps_deriv (a / b * b) = fps_deriv (a / b) * b + a / b * fps_deriv b" by simp
+  finally have "fps_deriv (a / b) * b^2 = fps_deriv a * b - a * fps_deriv b" using assms
+    by (simp add: power2_eq_square algebra_simps)
+  thus ?thesis by (cases "b = 0") (auto simp: eq_divide_imp)
+qed
+
+lemma fps_inverse_gp': "inverse (Abs_fps (\<lambda>n. 1::'a::field)) = 1 - X"
+  by (simp add: fps_inverse_gp fps_eq_iff X_def)
+
+lemma fps_one_over_one_minus_X_squared:
+  "inverse ((1 - X)^2 :: 'a :: field fps) = Abs_fps (\<lambda>n. of_nat (n+1))"
+proof -
+  have "inverse ((1 - X)^2 :: 'a fps) = fps_deriv (inverse (1 - X))"
+    by (subst fps_inverse_deriv) (simp_all add: fps_inverse_power)
+  also have "inverse (1 - X :: 'a fps) = Abs_fps (\<lambda>_. 1)"
+    by (subst fps_inverse_gp' [symmetric]) simp
+  also have "fps_deriv \<dots> = Abs_fps (\<lambda>n. of_nat (n + 1))"
+    by (simp add: fps_deriv_def)
+  finally show ?thesis .
+qed
+
+lemma fps_nth_deriv_X[simp]: "fps_nth_deriv n X = (if n = 0 then X else if n=1 then 1 else 0)"
+  by (cases n) simp_all
+
+lemma fps_inverse_X_plus1: "inverse (1 + X) = Abs_fps (\<lambda>n. (- (1::'a::field)) ^ n)"
+  (is "_ = ?r")
+proof -
+  have eq: "(1 + X) * ?r = 1"
+    unfolding minus_one_power_iff
+    by (auto simp add: field_simps fps_eq_iff)
+  show ?thesis
+    by (auto simp add: eq intro: fps_inverse_unique)
+qed
+
+
+subsection \<open>Integration\<close>
+
+definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
+  where "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
+
+lemma fps_deriv_fps_integral: "fps_deriv (fps_integral a a0) = a"
+  unfolding fps_integral_def fps_deriv_def
+  by (simp add: fps_eq_iff del: of_nat_Suc)
+
+lemma fps_integral_linear:
+  "fps_integral (fps_const a * f + fps_const b * g) (a*a0 + b*b0) =
+    fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0"
+  (is "?l = ?r")
+proof -
+  have "fps_deriv ?l = fps_deriv ?r"
+    by (simp add: fps_deriv_fps_integral)
+  moreover have "?l$0 = ?r$0"
+    by (simp add: fps_integral_def)
+  ultimately show ?thesis
+    unfolding fps_deriv_eq_iff by auto
+qed
+
+
+subsection \<open>Composition of FPSs\<close>
+
+definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps"  (infixl "oo" 55)
+  where "a oo b = Abs_fps (\<lambda>n. sum (\<lambda>i. a$i * (b^i$n)) {0..n})"
+
+lemma fps_compose_nth: "(a oo b)$n = sum (\<lambda>i. a$i * (b^i$n)) {0..n}"
+  by (simp add: fps_compose_def)
+
+lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0"
+  by (simp add: fps_compose_nth)
+
+lemma fps_compose_X[simp]: "a oo X = (a :: 'a::comm_ring_1 fps)"
+  by (simp add: fps_ext fps_compose_def mult_delta_right sum.delta')
+
+lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
+  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
+
+lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k"
+  unfolding numeral_fps_const by simp
+
+lemma neg_numeral_compose[simp]: "(- numeral k :: 'a::comm_ring_1 fps) oo b = - numeral k"
+  unfolding neg_numeral_fps_const by simp
+
+lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: 'a::comm_ring_1 fps)"
+  by (simp add: fps_eq_iff fps_compose_def mult_delta_left sum.delta not_le)
+
+
+subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
+
+subsubsection \<open>Rule 1\<close>
+  (* {a_{n+k}}_0^infty Corresponds to (f - sum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
+
+lemma fps_power_mult_eq_shift:
+  "X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
+    Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * X^i) {0 .. k}"
+  (is "?lhs = ?rhs")
+proof -
+  have "?lhs $ n = ?rhs $ n" for n :: nat
+  proof -
+    have "?lhs $ n = (if n < Suc k then 0 else a n)"
+      unfolding X_power_mult_nth by auto
+    also have "\<dots> = ?rhs $ n"
+    proof (induct k)
+      case 0
+      then show ?case
+        by (simp add: fps_sum_nth)
+    next
+      case (Suc k)
+      have "(Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
+        (Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} -
+          fps_const (a (Suc k)) * X^ Suc k) $ n"
+        by (simp add: field_simps)
+      also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
+        using Suc.hyps[symmetric] unfolding fps_sub_nth by simp
+      also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
+        unfolding X_power_mult_right_nth
+        apply (auto simp add: not_less fps_const_def)
+        apply (rule cong[of a a, OF refl])
+        apply arith
+        done
+      finally show ?case
+        by simp
+    qed
+    finally show ?thesis .
+  qed
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+
+subsubsection \<open>Rule 2\<close>
+
+  (* We can not reach the form of Wilf, but still near to it using rewrite rules*)
+  (* If f reprents {a_n} and P is a polynomial, then
+        P(xD) f represents {P(n) a_n}*)
+
+definition "XD = op * X \<circ> fps_deriv"
+
+lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: 'a::comm_ring_1 fps)"
+  by (simp add: XD_def field_simps)
+
+lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
+  by (simp add: XD_def field_simps)
+
+lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) =
+    fps_const c * XD a + fps_const d * XD (b :: 'a::comm_ring_1 fps)"
+  by simp
+
+lemma XDN_linear:
+  "(XD ^^ n) (fps_const c * a + fps_const d * b) =
+    fps_const c * (XD ^^ n) a + fps_const d * (XD ^^ n) (b :: 'a::comm_ring_1 fps)"
+  by (induct n) simp_all
+
+lemma fps_mult_X_deriv_shift: "X* fps_deriv a = Abs_fps (\<lambda>n. of_nat n* a$n)"
+  by (simp add: fps_eq_iff)
+
+lemma fps_mult_XD_shift:
+  "(XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
+  by (induct k arbitrary: a) (simp_all add: XD_def fps_eq_iff field_simps del: One_nat_def)
+
+
+subsubsection \<open>Rule 3\<close>
+
+text \<open>Rule 3 is trivial and is given by \<open>fps_times_def\<close>.\<close>
+
+
+subsubsection \<open>Rule 5 --- summation and "division" by (1 - X)\<close>
+
+lemma fps_divide_X_minus1_sum_lemma:
+  "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
+proof -
+  let ?sa = "Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
+  have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)"
+    by simp
+  have "a$n = ((1 - X) * ?sa) $ n" for n
+  proof (cases "n = 0")
+    case True
+    then show ?thesis
+      by (simp add: fps_mult_nth)
+  next
+    case False
+    then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1} \<union> {2..n} = {1..n}"
+      "{0..n - 1} \<union> {n} = {0..n}"
+      by (auto simp: set_eq_iff)
+    have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}" "{0..n - 1} \<inter> {n} = {}"
+      using False by simp_all
+    have f: "finite {0}" "finite {1}" "finite {2 .. n}"
+      "finite {0 .. n - 1}" "finite {n}" by simp_all
+    have "((1 - X) * ?sa) $ n = sum (\<lambda>i. (1 - X)$ i * ?sa $ (n - i)) {0 .. n}"
+      by (simp add: fps_mult_nth)
+    also have "\<dots> = a$n"
+      unfolding th0
+      unfolding sum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
+      unfolding sum.union_disjoint[OF f(2) f(3) d(2)]
+      apply (simp)
+      unfolding sum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
+      apply simp
+      done
+    finally show ?thesis
+      by simp
+  qed
+  then show ?thesis
+    unfolding fps_eq_iff by blast
+qed
+
+lemma fps_divide_X_minus1_sum:
+  "a /((1::'a::field fps) - X) = Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
+proof -
+  let ?X = "1 - (X::'a fps)"
+  have th0: "?X $ 0 \<noteq> 0"
+    by simp
+  have "a /?X = ?X *  Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) * inverse ?X"
+    using fps_divide_X_minus1_sum_lemma[of a, symmetric] th0
+    by (simp add: fps_divide_def mult.assoc)
+  also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) "
+    by (simp add: ac_simps)
+  finally show ?thesis
+    by (simp add: inverse_mult_eq_1[OF th0])
+qed
+
+
+subsubsection \<open>Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
+  finite product of FPS, also the relvant instance of powers of a FPS\<close>
+
+definition "natpermute n k = {l :: nat list. length l = k \<and> sum_list l = n}"
+
+lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
+  apply (auto simp add: natpermute_def)
+  apply (case_tac x)
+  apply auto
+  done
+
+lemma append_natpermute_less_eq:
+  assumes "xs @ ys \<in> natpermute n k"
+  shows "sum_list xs \<le> n"
+    and "sum_list ys \<le> n"
+proof -
+  from assms have "sum_list (xs @ ys) = n"
+    by (simp add: natpermute_def)
+  then have "sum_list xs + sum_list ys = n"
+    by simp
+  then show "sum_list xs \<le> n" and "sum_list ys \<le> n"
+    by simp_all
+qed
+
+lemma natpermute_split:
+  assumes "h \<le> k"
+  shows "natpermute n k =
+    (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})"
+  (is "?L = ?R" is "_ = (\<Union>m \<in>{0..n}. ?S m)")
+proof
+  show "?R \<subseteq> ?L"
+  proof
+    fix l
+    assume l: "l \<in> ?R"
+    from l obtain m xs ys where h: "m \<in> {0..n}"
+      and xs: "xs \<in> natpermute m h"
+      and ys: "ys \<in> natpermute (n - m) (k - h)"
+      and leq: "l = xs@ys" by blast
+    from xs have xs': "sum_list xs = m"
+      by (simp add: natpermute_def)
+    from ys have ys': "sum_list ys = n - m"
+      by (simp add: natpermute_def)
+    show "l \<in> ?L" using leq xs ys h
+      apply (clarsimp simp add: natpermute_def)
+      unfolding xs' ys'
+      using assms xs ys
+      unfolding natpermute_def
+      apply simp
+      done
+  qed
+  show "?L \<subseteq> ?R"
+  proof
+    fix l
+    assume l: "l \<in> natpermute n k"
+    let ?xs = "take h l"
+    let ?ys = "drop h l"
+    let ?m = "sum_list ?xs"
+    from l have ls: "sum_list (?xs @ ?ys) = n"
+      by (simp add: natpermute_def)
+    have xs: "?xs \<in> natpermute ?m h" using l assms
+      by (simp add: natpermute_def)
+    have l_take_drop: "sum_list l = sum_list (take h l @ drop h l)"
+      by simp
+    then have ys: "?ys \<in> natpermute (n - ?m) (k - h)"
+      using l assms ls by (auto simp add: natpermute_def simp del: append_take_drop_id)
+    from ls have m: "?m \<in> {0..n}"
+      by (simp add: l_take_drop del: append_take_drop_id)
+    from xs ys ls show "l \<in> ?R"
+      apply auto
+      apply (rule bexI [where x = "?m"])
+      apply (rule exI [where x = "?xs"])
+      apply (rule exI [where x = "?ys"])
+      using ls l
+      apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id)
+      apply simp
+      done
+  qed
+qed
+
+lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
+  by (auto simp add: natpermute_def)
+
+lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})"
+  apply (auto simp add: set_replicate_conv_if natpermute_def)
+  apply (rule nth_equalityI)
+  apply simp_all
+  done
+
+lemma natpermute_finite: "finite (natpermute n k)"
+proof (induct k arbitrary: n)
+  case 0
+  then show ?case
+    apply (subst natpermute_split[of 0 0, simplified])
+    apply (simp add: natpermute_0)
+    done
+next
+  case (Suc k)
+  then show ?case unfolding natpermute_split [of k "Suc k", simplified]
+    apply -
+    apply (rule finite_UN_I)
+    apply simp
+    unfolding One_nat_def[symmetric] natlist_trivial_1
+    apply simp
+    done
+qed
+
+lemma natpermute_contain_maximal:
+  "{xs \<in> natpermute n (k + 1). n \<in> set xs} = (\<Union>i\<in>{0 .. k}. {(replicate (k + 1) 0) [i:=n]})"
+  (is "?A = ?B")
+proof
+  show "?A \<subseteq> ?B"
+  proof
+    fix xs
+    assume "xs \<in> ?A"
+    then have H: "xs \<in> natpermute n (k + 1)" and n: "n \<in> set xs"
+      by blast+
+    then obtain i where i: "i \<in> {0.. k}" "xs!i = n"
+      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
+    have eqs: "({0..k} - {i}) \<union> {i} = {0..k}"
+      using i by auto
+    have f: "finite({0..k} - {i})" "finite {i}"
+      by auto
+    have d: "({0..k} - {i}) \<inter> {i} = {}"
+      using i by auto
+    from H have "n = sum (nth xs) {0..k}"
+      apply (simp add: natpermute_def)
+      apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth)
+      done
+    also have "\<dots> = n + sum (nth xs) ({0..k} - {i})"
+      unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp
+    finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
+      by auto
+    from H have xsl: "length xs = k+1"
+      by (simp add: natpermute_def)
+    from i have i': "i < length (replicate (k+1) 0)"   "i < k+1"
+      unfolding length_replicate by presburger+
+    have "xs = replicate (k+1) 0 [i := n]"
+      apply (rule nth_equalityI)
+      unfolding xsl length_list_update length_replicate
+      apply simp
+      apply clarify
+      unfolding nth_list_update[OF i'(1)]
+      using i zxs
+      apply (case_tac "ia = i")
+      apply (auto simp del: replicate.simps)
+      done
+    then show "xs \<in> ?B" using i by blast
+  qed
+  show "?B \<subseteq> ?A"
+  proof
+    fix xs
+    assume "xs \<in> ?B"
+    then obtain i where i: "i \<in> {0..k}" and xs: "xs = replicate (k + 1) 0 [i:=n]"
+      by auto
+    have nxs: "n \<in> set xs"
+      unfolding xs
+      apply (rule set_update_memI)
+      using i apply simp
+      done
+    have xsl: "length xs = k + 1"
+      by (simp only: xs length_replicate length_list_update)
+    have "sum_list xs = sum (nth xs) {0..<k+1}"
+      unfolding sum_list_sum_nth xsl ..
+    also have "\<dots> = sum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
+      by (rule sum.cong) (simp_all add: xs del: replicate.simps)
+    also have "\<dots> = n" using i by (simp add: sum.delta)
+    finally have "xs \<in> natpermute n (k + 1)"
+      using xsl unfolding natpermute_def mem_Collect_eq by blast
+    then show "xs \<in> ?A"
+      using nxs by blast
+  qed
+qed
+
+text \<open>The general form.\<close>
+lemma fps_prod_nth:
+  fixes m :: nat
+    and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
+  shows "(prod a {0 .. m}) $ n =
+    sum (\<lambda>v. prod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
+  (is "?P m n")
+proof (induct m arbitrary: n rule: nat_less_induct)
+  fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
+  show "?P m n"
+  proof (cases m)
+    case 0
+    then show ?thesis
+      apply simp
+      unfolding natlist_trivial_1[where n = n, unfolded One_nat_def]
+      apply simp
+      done
+  next
+    case (Suc k)
+    then have km: "k < m" by arith
+    have u0: "{0 .. k} \<union> {m} = {0..m}"
+      using Suc by (simp add: set_eq_iff) presburger
+    have f0: "finite {0 .. k}" "finite {m}" by auto
+    have d0: "{0 .. k} \<inter> {m} = {}" using Suc by auto
+    have "(prod a {0 .. m}) $ n = (prod a {0 .. k} * a m) $ n"
+      unfolding prod.union_disjoint[OF f0 d0, unfolded u0] by simp
+    also have "\<dots> = (\<Sum>i = 0..n. (\<Sum>v\<in>natpermute i (k + 1). \<Prod>j\<in>{0..k}. a j $ v ! j) * a m $ (n - i))"
+      unfolding fps_mult_nth H[rule_format, OF km] ..
+    also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
+      apply (simp add: Suc)
+      unfolding natpermute_split[of m "m + 1", simplified, of n,
+        unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
+      apply (subst sum.UNION_disjoint)
+      apply simp
+      apply simp
+      unfolding image_Collect[symmetric]
+      apply clarsimp
+      apply (rule finite_imageI)
+      apply (rule natpermute_finite)
+      apply (clarsimp simp add: set_eq_iff)
+      apply auto
+      apply (rule sum.cong)
+      apply (rule refl)
+      unfolding sum_distrib_right
+      apply (rule sym)
+      apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in sum.reindex_cong)
+      apply (simp add: inj_on_def)
+      apply auto
+      unfolding prod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
+      apply (clarsimp simp add: natpermute_def nth_append)
+      done
+    finally show ?thesis .
+  qed
+qed
+
+text \<open>The special form for powers.\<close>
+lemma fps_power_nth_Suc:
+  fixes m :: nat
+    and a :: "'a::comm_ring_1 fps"
+  shows "(a ^ Suc m)$n = sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
+proof -
+  have th0: "a^Suc m = prod (\<lambda>i. a) {0..m}"
+    by (simp add: prod_constant)
+  show ?thesis unfolding th0 fps_prod_nth ..
+qed
+
+lemma fps_power_nth:
+  fixes m :: nat
+    and a :: "'a::comm_ring_1 fps"
+  shows "(a ^m)$n =
+    (if m=0 then 1$n else sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
+  by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc)
+
+lemma fps_nth_power_0:
+  fixes m :: nat
+    and a :: "'a::comm_ring_1 fps"
+  shows "(a ^m)$0 = (a$0) ^ m"
+proof (cases m)
+  case 0
+  then show ?thesis by simp
+next
+  case (Suc n)
+  then have c: "m = card {0..n}" by simp
+  have "(a ^m)$0 = prod (\<lambda>i. a$0) {0..n}"
+    by (simp add: Suc fps_power_nth del: replicate.simps power_Suc)
+  also have "\<dots> = (a$0) ^ m"
+   unfolding c by (rule prod_constant)
+ finally show ?thesis .
+qed
+
+lemma natpermute_max_card:
+  assumes n0: "n \<noteq> 0"
+  shows "card {xs \<in> natpermute n (k + 1). n \<in> set xs} = k + 1"
+  unfolding natpermute_contain_maximal
+proof -
+  let ?A = "\<lambda>i. {replicate (k + 1) 0[i := n]}"
+  let ?K = "{0 ..k}"
+  have fK: "finite ?K"
+    by simp
+  have fAK: "\<forall>i\<in>?K. finite (?A i)"
+    by auto
+  have d: "\<forall>i\<in> ?K. \<forall>j\<in> ?K. i \<noteq> j \<longrightarrow>
+    {replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
+  proof clarify
+    fix i j
+    assume i: "i \<in> ?K" and j: "j \<in> ?K" and ij: "i \<noteq> j"
+    have False if eq: "replicate (k+1) 0 [i:=n] = replicate (k+1) 0 [j:= n]"
+    proof -
+      have "(replicate (k+1) 0 [i:=n] ! i) = n"
+        using i by (simp del: replicate.simps)
+      moreover
+      have "(replicate (k+1) 0 [j:=n] ! i) = 0"
+        using i ij by (simp del: replicate.simps)
+      ultimately show ?thesis
+        using eq n0 by (simp del: replicate.simps)
+    qed
+    then show "{replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
+      by auto
+  qed
+  from card_UN_disjoint[OF fK fAK d]
+  show "card (\<Union>i\<in>{0..k}. {replicate (k + 1) 0[i := n]}) = k + 1"
+    by simp
+qed
+
+lemma fps_power_Suc_nth:
+  fixes f :: "'a :: comm_ring_1 fps"
+  assumes k: "k > 0"
+  shows "(f ^ Suc m) $ k = 
+           of_nat (Suc m) * (f $ k * (f $ 0) ^ m) +
+           (\<Sum>v\<in>{v\<in>natpermute k (m+1). k \<notin> set v}. \<Prod>j = 0..m. f $ v ! j)"
+proof -
+  define A B 
+    where "A = {v\<in>natpermute k (m+1). k \<in> set v}" 
+      and  "B = {v\<in>natpermute k (m+1). k \<notin> set v}"
+  have [simp]: "finite A" "finite B" "A \<inter> B = {}" by (auto simp: A_def B_def natpermute_finite)
+
+  from natpermute_max_card[of k m] k have card_A: "card A = m + 1" by (simp add: A_def)
+  {
+    fix v assume v: "v \<in> A"
+    from v have [simp]: "length v = Suc m" by (simp add: A_def natpermute_def)
+    from v have "\<exists>j. j \<le> m \<and> v ! j = k" 
+      by (auto simp: set_conv_nth A_def natpermute_def less_Suc_eq_le)
+    then guess j by (elim exE conjE) note j = this
+    
+    from v have "k = sum_list v" by (simp add: A_def natpermute_def)
+    also have "\<dots> = (\<Sum>i=0..m. v ! i)"
+      by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum_op_ivl_Suc)
+    also from j have "{0..m} = insert j ({0..m}-{j})" by auto
+    also from j have "(\<Sum>i\<in>\<dots>. v ! i) = k + (\<Sum>i\<in>{0..m}-{j}. v ! i)"
+      by (subst sum.insert) simp_all
+    finally have "(\<Sum>i\<in>{0..m}-{j}. v ! i) = 0" by simp
+    hence zero: "v ! i = 0" if "i \<in> {0..m}-{j}" for i using that
+      by (subst (asm) sum_eq_0_iff) auto
+      
+    from j have "{0..m} = insert j ({0..m} - {j})" by auto
+    also from j have "(\<Prod>i\<in>\<dots>. f $ (v ! i)) = f $ k * (\<Prod>i\<in>{0..m} - {j}. f $ (v ! i))"
+      by (subst prod.insert) auto
+    also have "(\<Prod>i\<in>{0..m} - {j}. f $ (v ! i)) = (\<Prod>i\<in>{0..m} - {j}. f $ 0)"
+      by (intro prod.cong) (simp_all add: zero)
+    also from j have "\<dots> = (f $ 0) ^ m" by (subst prod_constant) simp_all
+    finally have "(\<Prod>j = 0..m. f $ (v ! j)) = f $ k * (f $ 0) ^ m" .
+  } note A = this
+  
+  have "(f ^ Suc m) $ k = (\<Sum>v\<in>natpermute k (m + 1). \<Prod>j = 0..m. f $ v ! j)"
+    by (rule fps_power_nth_Suc)
+  also have "natpermute k (m+1) = A \<union> B" unfolding A_def B_def by blast
+  also have "(\<Sum>v\<in>\<dots>. \<Prod>j = 0..m. f $ (v ! j)) = 
+               (\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) + (\<Sum>v\<in>B. \<Prod>j = 0..m. f $ (v ! j))"
+    by (intro sum.union_disjoint) simp_all   
+  also have "(\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)"
+    by (simp add: A card_A)
+  finally show ?thesis by (simp add: B_def)
+qed 
+  
+lemma fps_power_Suc_eqD:
+  fixes f g :: "'a :: {idom,semiring_char_0} fps"
+  assumes "f ^ Suc m = g ^ Suc m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0"
+  shows   "f = g"
+proof (rule fps_ext)
+  fix k :: nat
+  show "f $ k = g $ k"
+  proof (induction k rule: less_induct)
+    case (less k)
+    show ?case
+    proof (cases "k = 0")
+      case False
+      let ?h = "\<lambda>f. (\<Sum>v | v \<in> natpermute k (m + 1) \<and> k \<notin> set v. \<Prod>j = 0..m. f $ v ! j)"
+      from False fps_power_Suc_nth[of k f m] fps_power_Suc_nth[of k g m]
+        have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h f =
+                g $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h g" using assms 
+        by (simp add: mult_ac del: power_Suc of_nat_Suc)
+      also have "v ! i < k" if "v \<in> {v\<in>natpermute k (m+1). k \<notin> set v}" "i \<le> m" for v i
+        using that elem_le_sum_list_nat[of i v] unfolding natpermute_def
+        by (auto simp: set_conv_nth dest!: spec[of _ i])
+      hence "?h f = ?h g"
+        by (intro sum.cong refl prod.cong less lessI) (auto simp: natpermute_def)
+      finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)"
+        by simp
+      with assms show "f $ k = g $ k" 
+        by (subst (asm) mult_right_cancel) (auto simp del: of_nat_Suc)
+    qed (simp_all add: assms)
+  qed
+qed
+
+lemma fps_power_Suc_eqD':
+  fixes f g :: "'a :: {idom,semiring_char_0} fps"
+  assumes "f ^ Suc m = g ^ Suc m" "f $ subdegree f = g $ subdegree g"
+  shows   "f = g"
+proof (cases "f = 0")
+  case False
+  have "Suc m * subdegree f = subdegree (f ^ Suc m)"
+    by (rule subdegree_power [symmetric])
+  also have "f ^ Suc m = g ^ Suc m" by fact
+  also have "subdegree \<dots> = Suc m * subdegree g" by (rule subdegree_power)
+  finally have [simp]: "subdegree f = subdegree g"
+    by (subst (asm) Suc_mult_cancel1)
+  have "fps_shift (subdegree f) f * X ^ subdegree f = f"
+    by (rule subdegree_decompose [symmetric])
+  also have "\<dots> ^ Suc m = g ^ Suc m" by fact
+  also have "g = fps_shift (subdegree g) g * X ^ subdegree g"
+    by (rule subdegree_decompose)
+  also have "subdegree f = subdegree g" by fact
+  finally have "fps_shift (subdegree g) f ^ Suc m = fps_shift (subdegree g) g ^ Suc m"
+    by (simp add: algebra_simps power_mult_distrib del: power_Suc)
+  hence "fps_shift (subdegree g) f = fps_shift (subdegree g) g"
+    by (rule fps_power_Suc_eqD) (insert assms False, auto)
+  with subdegree_decompose[of f] subdegree_decompose[of g] show ?thesis by simp
+qed (insert assms, simp_all)
+
+lemma fps_power_eqD':
+  fixes f g :: "'a :: {idom,semiring_char_0} fps"
+  assumes "f ^ m = g ^ m" "f $ subdegree f = g $ subdegree g" "m > 0"
+  shows   "f = g"
+  using fps_power_Suc_eqD'[of f "m-1" g] assms by simp
+
+lemma fps_power_eqD:
+  fixes f g :: "'a :: {idom,semiring_char_0} fps"
+  assumes "f ^ m = g ^ m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0" "m > 0"
+  shows   "f = g"
+  by (rule fps_power_eqD'[of f m g]) (insert assms, simp_all)
+
+lemma fps_compose_inj_right:
+  assumes a0: "a$0 = (0::'a::idom)"
+    and a1: "a$1 \<noteq> 0"
+  shows "(b oo a = c oo a) \<longleftrightarrow> b = c"
+  (is "?lhs \<longleftrightarrow>?rhs")
+proof
+  show ?lhs if ?rhs using that by simp
+  show ?rhs if ?lhs
+  proof -
+    have "b$n = c$n" for n
+    proof (induct n rule: nat_less_induct)
+      fix n
+      assume H: "\<forall>m<n. b$m = c$m"
+      show "b$n = c$n"
+      proof (cases n)
+        case 0
+        from \<open>?lhs\<close> have "(b oo a)$n = (c oo a)$n"
+          by simp
+        then show ?thesis
+          using 0 by (simp add: fps_compose_nth)
+      next
+        case (Suc n1)
+        have f: "finite {0 .. n1}" "finite {n}" by simp_all
+        have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using Suc by auto
+        have d: "{0 .. n1} \<inter> {n} = {}" using Suc by auto
+        have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
+          apply (rule sum.cong)
+          using H Suc
+          apply auto
+          done
+        have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
+          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq
+          using startsby_zero_power_nth_same[OF a0]
+          by simp
+        have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
+          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq]
+          using startsby_zero_power_nth_same[OF a0]
+          by simp
+        from \<open>?lhs\<close>[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
+        show ?thesis by auto
+      qed
+    qed
+    then show ?rhs by (simp add: fps_eq_iff)
+  qed
+qed
+
+
+subsection \<open>Radicals\<close>
+
+declare prod.cong [fundef_cong]
+
+function radical :: "(nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a::field fps \<Rightarrow> nat \<Rightarrow> 'a"
+where
+  "radical r 0 a 0 = 1"
+| "radical r 0 a (Suc n) = 0"
+| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
+| "radical r (Suc k) a (Suc n) =
+    (a$ Suc n - sum (\<lambda>xs. prod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
+      {xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
+    (of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
+  by pat_completeness auto
+
+termination radical
+proof
+  let ?R = "measure (\<lambda>(r, k, a, n). n)"
+  {
+    show "wf ?R" by auto
+  next
+    fix r k a n xs i
+    assume xs: "xs \<in> {xs \<in> natpermute (Suc n) (Suc k). Suc n \<notin> set xs}" and i: "i \<in> {0..k}"
+    have False if c: "Suc n \<le> xs ! i"
+    proof -
+      from xs i have "xs !i \<noteq> Suc n"
+        by (auto simp add: in_set_conv_nth natpermute_def)
+      with c have c': "Suc n < xs!i" by arith
+      have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
+        by simp_all
+      have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
+        by auto
+      have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
+        using i by auto
+      from xs have "Suc n = sum_list xs"
+        by (simp add: natpermute_def)
+      also have "\<dots> = sum (nth xs) {0..<Suc k}" using xs
+        by (simp add: natpermute_def sum_list_sum_nth)
+      also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
+        unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+        unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
+        by simp
+      finally show ?thesis using c' by simp
+    qed
+    then show "((r, Suc k, a, xs!i), r, Suc k, a, Suc n) \<in> ?R"
+      apply auto
+      apply (metis not_less)
+      done
+  next
+    fix r k a n
+    show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \<in> ?R" by simp
+  }
+qed
+
+definition "fps_radical r n a = Abs_fps (radical r n a)"
+
+lemma fps_radical0[simp]: "fps_radical r 0 a = 1"
+  apply (auto simp add: fps_eq_iff fps_radical_def)
+  apply (case_tac n)
+  apply auto
+  done
+
+lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n = 0 then 1 else r n (a$0))"
+  by (cases n) (simp_all add: fps_radical_def)
+
+lemma fps_radical_power_nth[simp]:
+  assumes r: "(r k (a$0)) ^ k = a$0"
+  shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)"
+proof (cases k)
+  case 0
+  then show ?thesis by simp
+next
+  case (Suc h)
+  have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
+    unfolding fps_power_nth Suc by simp
+  also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
+    apply (rule prod.cong)
+    apply simp
+    using Suc
+    apply (subgoal_tac "replicate k 0 ! x = 0")
+    apply (auto intro: nth_replicate simp del: replicate.simps)
+    done
+  also have "\<dots> = a$0"
+    using r Suc by (simp add: prod_constant)
+  finally show ?thesis
+    using Suc by simp
+qed
+
+lemma power_radical:
+  fixes a:: "'a::field_char_0 fps"
+  assumes a0: "a$0 \<noteq> 0"
+  shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \<longleftrightarrow> (fps_radical r (Suc k) a) ^ (Suc k) = a"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  let ?r = "fps_radical r (Suc k) a"
+  show ?rhs if r0: ?lhs
+  proof -
+    from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
+    have "?r ^ Suc k $ z = a$z" for z
+    proof (induct z rule: nat_less_induct)
+      fix n
+      assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
+      show "?r ^ Suc k $ n = a $n"
+      proof (cases n)
+        case 0
+        then show ?thesis
+          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp
+      next
+        case (Suc n1)
+        then have "n \<noteq> 0" by simp
+        let ?Pnk = "natpermute n (k + 1)"
+        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+        have f: "finite ?Pnkn" "finite ?Pnknn"
+          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+          by (metis natpermute_finite)+
+        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+        proof (rule sum.cong)
+          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
+          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
+            fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding natpermute_contain_maximal by auto
+          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
+              (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
+            apply (rule prod.cong, simp)
+            using i r0
+            apply (simp del: replicate.simps)
+            done
+          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
+            using i r0 by (simp add: prod_gen_delta)
+          finally show ?ths .
+        qed rule
+        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+          by (simp add: natpermute_max_card[OF \<open>n \<noteq> 0\<close>, simplified])
+        also have "\<dots> = a$n - sum ?f ?Pnknn"
+          unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
+        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
+        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
+          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
+        also have "\<dots> = a$n" unfolding fn by simp
+        finally show ?thesis .
+      qed
+    qed
+    then show ?thesis using r0 by (simp add: fps_eq_iff)
+  qed
+  show ?lhs if ?rhs
+  proof -
+    from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0"
+      by simp
+    then show ?thesis
+      unfolding fps_power_nth_Suc
+      by (simp add: prod_constant del: replicate.simps)
+  qed
+qed
+
+(*
+lemma power_radical:
+  fixes a:: "'a::field_char_0 fps"
+  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \<noteq> 0"
+  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
+proof-
+  let ?r = "fps_radical r (Suc k) a"
+  from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
+  {fix z have "?r ^ Suc k $ z = a$z"
+    proof(induct z rule: nat_less_induct)
+      fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
+      {assume "n = 0" then have "?r ^ Suc k $ n = a $n"
+          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
+      moreover
+      {fix n1 assume n1: "n = Suc n1"
+        have fK: "finite {0..k}" by simp
+        have nz: "n \<noteq> 0" using n1 by arith
+        let ?Pnk = "natpermute n (k + 1)"
+        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+        have f: "finite ?Pnkn" "finite ?Pnknn"
+          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+          by (metis natpermute_finite)+
+        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+        proof(rule sum.cong2)
+          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
+          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding natpermute_contain_maximal by auto
+          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
+            apply (rule prod.cong, simp)
+            using i r0 by (simp del: replicate.simps)
+          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
+            unfolding prod_gen_delta[OF fK] using i r0 by simp
+          finally show ?ths .
+        qed
+        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+          by (simp add: natpermute_max_card[OF nz, simplified])
+        also have "\<dots> = a$n - sum ?f ?Pnknn"
+          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
+        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
+        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
+          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
+        also have "\<dots> = a$n" unfolding fn by simp
+        finally have "?r ^ Suc k $ n = a $n" .}
+      ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
+  qed }
+  then show ?thesis by (simp add: fps_eq_iff)
+qed
+
+*)
+lemma eq_divide_imp':
+  fixes c :: "'a::field"
+  shows "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
+  by (simp add: field_simps)
+
+lemma radical_unique:
+  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
+    and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0"
+    and b0: "b$0 \<noteq> 0"
+  shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
+    (is "?lhs \<longleftrightarrow> ?rhs" is "_ \<longleftrightarrow> a = ?r")
+proof
+  show ?lhs if ?rhs
+    using that using power_radical[OF b0, of r k, unfolded r0] by simp
+  show ?rhs if ?lhs
+  proof -
+    have r00: "r (Suc k) (b$0) \<noteq> 0" using b0 r0 by auto
+    have ceq: "card {0..k} = Suc k" by simp
+    from a0 have a0r0: "a$0 = ?r$0" by simp
+    have "a $ n = ?r $ n" for n
+    proof (induct n rule: nat_less_induct)
+      fix n
+      assume h: "\<forall>m<n. a$m = ?r $m"
+      show "a$n = ?r $ n"
+      proof (cases n)
+        case 0
+        then show ?thesis using a0 by simp
+      next
+        case (Suc n1)
+        have fK: "finite {0..k}" by simp
+        have nz: "n \<noteq> 0" using Suc by simp
+        let ?Pnk = "natpermute n (Suc k)"
+        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
+        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
+        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
+        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
+        have f: "finite ?Pnkn" "finite ?Pnknn"
+          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
+          by (metis natpermute_finite)+
+        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
+        let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
+        have "sum ?g ?Pnkn = sum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
+        proof (rule sum.cong)
+          fix v
+          assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
+          let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
+          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
+            unfolding Suc_eq_plus1 natpermute_contain_maximal
+            by (auto simp del: replicate.simps)
+          have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
+            apply (rule prod.cong, simp)
+            using i a0
+            apply (simp del: replicate.simps)
+            done
+          also have "\<dots> = a $ n * (?r $ 0)^k"
+            using i by (simp add: prod_gen_delta)
+          finally show ?ths .
+        qed rule
+        then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
+          by (simp add: natpermute_max_card[OF nz, simplified])
+        have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn"
+        proof (rule sum.cong, rule refl, rule prod.cong, simp)
+          fix xs i
+          assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
+          have False if c: "n \<le> xs ! i"
+          proof -
+            from xs i have "xs ! i \<noteq> n"
+              by (auto simp add: in_set_conv_nth natpermute_def)
+            with c have c': "n < xs!i" by arith
+            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
+              by simp_all
+            have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
+              by auto
+            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
+              using i by auto
+            from xs have "n = sum_list xs"
+              by (simp add: natpermute_def)
+            also have "\<dots> = sum (nth xs) {0..<Suc k}"
+              using xs by (simp add: natpermute_def sum_list_sum_nth)
+            also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
+              unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+              unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
+              by simp
+            finally show ?thesis using c' by simp
+          qed
+          then have thn: "xs!i < n" by presburger
+          from h[rule_format, OF thn] show "a$(xs !i) = ?r$(xs!i)" .
+        qed
+        have th00: "\<And>x::'a. of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
+          by (simp add: field_simps del: of_nat_Suc)
+        from \<open>?lhs\<close> have "b$n = a^Suc k $ n"
+          by (simp add: fps_eq_iff)
+        also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn"
+          unfolding fps_power_nth_Suc
+          using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
+            unfolded eq, of ?g] by simp
+        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn"
+          unfolding th0 th1 ..
+        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn"
+          by simp
+        then have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
+          apply -
+          apply (rule eq_divide_imp')
+          using r00
+          apply (simp del: of_nat_Suc)
+          apply (simp add: ac_simps)
+          done
+        then show ?thesis
+          apply (simp del: of_nat_Suc)
+          unfolding fps_radical_def Suc
+          apply (simp add: field_simps Suc th00 del: of_nat_Suc)
+          done
+      qed
+    qed
+    then show ?rhs by (simp add: fps_eq_iff)
+  qed
+qed
+
+
+lemma radical_power:
+  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
+    and a0: "(a$0 :: 'a::field_char_0) \<noteq> 0"
+  shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
+proof -
+  let ?ak = "a^ Suc k"
+  have ak0: "?ak $ 0 = (a$0) ^ Suc k"
+    by (simp add: fps_nth_power_0 del: power_Suc)
+  from r0 have th0: "r (Suc k) (a ^ Suc k $ 0) ^ Suc k = a ^ Suc k $ 0"
+    using ak0 by auto
+  from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0"
+    by auto
+  from ak0 a0 have ak00: "?ak $ 0 \<noteq>0 "
+    by auto
+  from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis
+    by metis
+qed
+
+lemma fps_deriv_radical:
+  fixes a :: "'a::field_char_0 fps"
+  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
+    and a0: "a$0 \<noteq> 0"
+  shows "fps_deriv (fps_radical r (Suc k) a) =
+    fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)"
+proof -
+  let ?r = "fps_radical r (Suc k) a"
+  let ?w = "(fps_const (of_nat (Suc k)) * ?r ^ k)"
+  from a0 r0 have r0': "r (Suc k) (a$0) \<noteq> 0"
+    by auto
+  from r0' have w0: "?w $ 0 \<noteq> 0"
+    by (simp del: of_nat_Suc)
+  note th0 = inverse_mult_eq_1[OF w0]
+  let ?iw = "inverse ?w"
+  from iffD1[OF power_radical[of a r], OF a0 r0]
+  have "fps_deriv (?r ^ Suc k) = fps_deriv a"
+    by simp
+  then have "fps_deriv ?r * ?w = fps_deriv a"
+    by (simp add: fps_deriv_power ac_simps del: power_Suc)
+  then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a"
+    by simp
+  with a0 r0 have "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w"
+    by (subst fps_divide_unit) (auto simp del: of_nat_Suc)
+  then show ?thesis unfolding th0 by simp
+qed
+
+lemma radical_mult_distrib:
+  fixes a :: "'a::field_char_0 fps"
+  assumes k: "k > 0"
+    and ra0: "r k (a $ 0) ^ k = a $ 0"
+    and rb0: "r k (b $ 0) ^ k = b $ 0"
+    and a0: "a $ 0 \<noteq> 0"
+    and b0: "b $ 0 \<noteq> 0"
+  shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \<longleftrightarrow>
+    fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show ?rhs if r0': ?lhs
+  proof -
+    from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0"
+      by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
+    show ?thesis
+    proof (cases k)
+      case 0
+      then show ?thesis using r0' by simp
+    next
+      case (Suc h)
+      let ?ra = "fps_radical r (Suc h) a"
+      let ?rb = "fps_radical r (Suc h) b"
+      have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
+        using r0' Suc by (simp add: fps_mult_nth)
+      have ab0: "(a*b) $ 0 \<noteq> 0"
+        using a0 b0 by (simp add: fps_mult_nth)
+      from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded Suc] th0 ab0, symmetric]
+        iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded Suc]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded Suc]] Suc r0'
+      show ?thesis
+        by (auto simp add: power_mult_distrib simp del: power_Suc)
+    qed
+  qed
+  show ?lhs if ?rhs
+  proof -
+    from that have "(fps_radical r k (a * b)) $ 0 = (fps_radical r k a * fps_radical r k b) $ 0"
+      by simp
+    then show ?thesis
+      using k by (simp add: fps_mult_nth)
+  qed
+qed
+
+(*
+lemma radical_mult_distrib:
+  fixes a:: "'a::field_char_0 fps"
+  assumes
+  ra0: "r k (a $ 0) ^ k = a $ 0"
+  and rb0: "r k (b $ 0) ^ k = b $ 0"
+  and r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)"
+  and a0: "a$0 \<noteq> 0"
+  and b0: "b$0 \<noteq> 0"
+  shows "fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)"
+proof-
+  from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
+    by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
+  {assume "k=0" then have ?thesis by simp}
+  moreover
+  {fix h assume k: "k = Suc h"
+  let ?ra = "fps_radical r (Suc h) a"
+  let ?rb = "fps_radical r (Suc h) b"
+  have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
+    using r0' k by (simp add: fps_mult_nth)
+  have ab0: "(a*b) $ 0 \<noteq> 0" using a0 b0 by (simp add: fps_mult_nth)
+  from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric]
+    power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k
+  have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)}
+ultimately show ?thesis by (cases k, auto)
+qed
+*)
+
+lemma fps_divide_1 [simp]: "(a :: 'a::field fps) / 1 = a"
+  by (fact div_by_1)
+
+lemma radical_divide:
+  fixes a :: "'a::field_char_0 fps"
+  assumes kp: "k > 0"
+    and ra0: "(r k (a $ 0)) ^ k = a $ 0"
+    and rb0: "(r k (b $ 0)) ^ k = b $ 0"
+    and a0: "a$0 \<noteq> 0"
+    and b0: "b$0 \<noteq> 0"
+  shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \<longleftrightarrow>
+    fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
+  (is "?lhs = ?rhs")
+proof
+  let ?r = "fps_radical r k"
+  from kp obtain h where k: "k = Suc h"
+    by (cases k) auto
+  have ra0': "r k (a$0) \<noteq> 0" using a0 ra0 k by auto
+  have rb0': "r k (b$0) \<noteq> 0" using b0 rb0 k by auto
+
+  show ?lhs if ?rhs
+  proof -
+    from that have "?r (a/b) $ 0 = (?r a / ?r b)$0"
+      by simp
+    then show ?thesis
+      using k a0 b0 rb0' by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
+  qed
+  show ?rhs if ?lhs
+  proof -
+    from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0"
+      by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def)
+    have th0: "r k ((a/b)$0) ^ k = (a/b)$0"
+      by (simp add: \<open>?lhs\<close> power_divide ra0 rb0)
+    from a0 b0 ra0' rb0' kp \<open>?lhs\<close>
+    have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0"
+      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
+    from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \<noteq> 0"
+      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero)
+    note tha[simp] = iffD1[OF power_radical[where r=r and k=h], OF a0 ra0[unfolded k], unfolded k[symmetric]]
+    note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]]
+    from b0 rb0' have th2: "(?r a / ?r b)^k = a/b"
+      by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric])
+
+    from iffD1[OF radical_unique[where r=r and a="?r a / ?r b" and b="a/b" and k=h], symmetric, unfolded k[symmetric], OF th0 th1 ab0' th2]
+    show ?thesis .
+  qed
+qed
+
+lemma radical_inverse:
+  fixes a :: "'a::field_char_0 fps"
+  assumes k: "k > 0"
+    and ra0: "r k (a $ 0) ^ k = a $ 0"
+    and r1: "(r k 1)^k = 1"
+    and a0: "a$0 \<noteq> 0"
+  shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \<longleftrightarrow>
+    fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a"
+  using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0
+  by (simp add: divide_inverse fps_divide_def)
+
+
+subsection \<open>Derivative of composition\<close>
+
+lemma fps_compose_deriv:
+  fixes a :: "'a::idom fps"
+  assumes b0: "b$0 = 0"
+  shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b"
+proof -
+  have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n
+  proof -
+    have "(fps_deriv (a oo b))$n = sum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
+      by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc)
+    also have "\<dots> = sum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
+      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
+    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
+      unfolding fps_mult_left_const_nth  by (simp add: field_simps)
+    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
+      unfolding fps_mult_nth ..
+    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
+      apply (rule sum.mono_neutral_right)
+      apply (auto simp add: mult_delta_left sum.delta not_le)
+      done
+    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
+      unfolding fps_deriv_nth
+      by (rule sum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
+    finally have th0: "(fps_deriv (a oo b))$n =
+      sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
+
+    have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
+      unfolding fps_mult_nth by (simp add: ac_simps)
+    also have "\<dots> = sum (\<lambda>i. sum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
+      unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc
+      apply (rule sum.cong)
+      apply (rule refl)
+      apply (rule sum.mono_neutral_left)
+      apply (simp_all add: subset_eq)
+      apply clarify
+      apply (subgoal_tac "b^i$x = 0")
+      apply simp
+      apply (rule startsby_zero_power_prefix[OF b0, rule_format])
+      apply simp
+      done
+    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
+      unfolding sum_distrib_left
+      apply (subst sum.commute)
+      apply (rule sum.cong, rule refl)+
+      apply simp
+      done
+    finally show ?thesis
+      unfolding th0 by simp
+  qed
+  then show ?thesis by (simp add: fps_eq_iff)
+qed
+
+lemma fps_mult_X_plus_1_nth:
+  "((1+X)*a) $n = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
+proof (cases n)
+  case 0
+  then show ?thesis
+    by (simp add: fps_mult_nth)
+next
+  case (Suc m)
+  have "((1 + X)*a) $ n = sum (\<lambda>i. (1 + X) $ i * a $ (n - i)) {0..n}"
+    by (simp add: fps_mult_nth)
+  also have "\<dots> = sum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
+    unfolding Suc by (rule sum.mono_neutral_right) auto
+  also have "\<dots> = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
+    by (simp add: Suc)
+  finally show ?thesis .
+qed
+
+
+subsection \<open>Finite FPS (i.e. polynomials) and X\<close>
+
+lemma fps_poly_sum_X:
+  assumes "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
+  shows "a = sum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
+proof -
+  have "a$i = ?r$i" for i
+    unfolding fps_sum_nth fps_mult_left_const_nth X_power_nth
+    by (simp add: mult_delta_right sum.delta' assms)
+  then show ?thesis
+    unfolding fps_eq_iff by blast
+qed
+
+
+subsection \<open>Compositional inverses\<close>
+
+fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
+where
+  "compinv a 0 = X$0"
+| "compinv a (Suc n) =
+    (X$ Suc n - sum (\<lambda>i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
+
+definition "fps_inv a = Abs_fps (compinv a)"
+
+lemma fps_inv:
+  assumes a0: "a$0 = 0"
+    and a1: "a$1 \<noteq> 0"
+  shows "fps_inv a oo a = X"
+proof -
+  let ?i = "fps_inv a oo a"
+  have "?i $n = X$n" for n
+  proof (induct n rule: nat_less_induct)
+    fix n
+    assume h: "\<forall>m<n. ?i$m = X$m"
+    show "?i $ n = X$n"
+    proof (cases n)
+      case 0
+      then show ?thesis using a0
+        by (simp add: fps_compose_nth fps_inv_def)
+    next
+      case (Suc n1)
+      have "?i $ n = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
+        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
+      also have "\<dots> = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
+        (X$ Suc n1 - sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
+        using a0 a1 Suc by (simp add: fps_inv_def)
+      also have "\<dots> = X$n" using Suc by simp
+      finally show ?thesis .
+    qed
+  qed
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+
+fun gcompinv :: "'a fps \<Rightarrow> 'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
+where
+  "gcompinv b a 0 = b$0"
+| "gcompinv b a (Suc n) =
+    (b$ Suc n - sum (\<lambda>i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
+
+definition "fps_ginv b a = Abs_fps (gcompinv b a)"
+
+lemma fps_ginv:
+  assumes a0: "a$0 = 0"
+    and a1: "a$1 \<noteq> 0"
+  shows "fps_ginv b a oo a = b"
+proof -
+  let ?i = "fps_ginv b a oo a"
+  have "?i $n = b$n" for n
+  proof (induct n rule: nat_less_induct)
+    fix n
+    assume h: "\<forall>m<n. ?i$m = b$m"
+    show "?i $ n = b$n"
+    proof (cases n)
+      case 0
+      then show ?thesis using a0
+        by (simp add: fps_compose_nth fps_ginv_def)
+    next
+      case (Suc n1)
+      have "?i $ n = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
+        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
+      also have "\<dots> = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
+        (b$ Suc n1 - sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
+        using a0 a1 Suc by (simp add: fps_ginv_def)
+      also have "\<dots> = b$n" using Suc by simp
+      finally show ?thesis .
+    qed
+  qed
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+lemma fps_inv_ginv: "fps_inv = fps_ginv X"
+  apply (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def)
+  apply (induct_tac n rule: nat_less_induct)
+  apply auto
+  apply (case_tac na)
+  apply simp
+  apply simp
+  done
+
+lemma fps_compose_1[simp]: "1 oo a = 1"
+  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
+
+lemma fps_compose_0[simp]: "0 oo a = 0"
+  by (simp add: fps_eq_iff fps_compose_nth)
+
+lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)"
+  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral)
+
+lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
+  by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib)
+
+lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\<lambda>i. f i oo a) S"
+proof (cases "finite S")
+  case True
+  show ?thesis
+  proof (rule finite_induct[OF True])
+    show "sum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
+      by simp
+  next
+    fix x F
+    assume fF: "finite F"
+      and xF: "x \<notin> F"
+      and h: "sum f F oo a = sum (\<lambda>i. f i oo a) F"
+    show "sum f (insert x F) oo a  = sum (\<lambda>i. f i oo a) (insert x F)"
+      using fF xF h by (simp add: fps_compose_add_distrib)
+  qed
+next
+  case False
+  then show ?thesis by simp
+qed
+
+lemma convolution_eq:
+  "sum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
+    sum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
+  by (rule sum.reindex_bij_witness[where i=fst and j="\<lambda>i. (i, n - i)"]) auto
+
+lemma product_composition_lemma:
+  assumes c0: "c$0 = (0::'a::idom)"
+    and d0: "d$0 = 0"
+  shows "((a oo c) * (b oo d))$n =
+    sum (\<lambda>(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}"  (is "?l = ?r")
+proof -
+  let ?S = "{(k::nat, m::nat). k + m \<le> n}"
+  have s: "?S \<subseteq> {0..n} \<times> {0..n}" by (auto simp add: subset_eq)
+  have f: "finite {(k::nat, m::nat). k + m \<le> n}"
+    apply (rule finite_subset[OF s])
+    apply auto
+    done
+  have "?r =  sum (\<lambda>i. sum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
+    apply (simp add: fps_mult_nth sum_distrib_left)
+    apply (subst sum.commute)
+    apply (rule sum.cong)
+    apply (auto simp add: field_simps)
+    done
+  also have "\<dots> = ?l"
+    apply (simp add: fps_mult_nth fps_compose_nth sum_product)
+    apply (rule sum.cong)
+    apply (rule refl)
+    apply (simp add: sum.cartesian_product mult.assoc)
+    apply (rule sum.mono_neutral_right[OF f])
+    apply (simp add: subset_eq)
+    apply presburger
+    apply clarsimp
+    apply (rule ccontr)
+    apply (clarsimp simp add: not_le)
+    apply (case_tac "x < aa")
+    apply simp
+    apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0])
+    apply blast
+    apply simp
+    apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0])
+    apply blast
+    done
+  finally show ?thesis by simp
+qed
+
+lemma product_composition_lemma':
+  assumes c0: "c$0 = (0::'a::idom)"
+    and d0: "d$0 = 0"
+  shows "((a oo c) * (b oo d))$n =
+    sum (\<lambda>k. sum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}"  (is "?l = ?r")
+  unfolding product_composition_lemma[OF c0 d0]
+  unfolding sum.cartesian_product
+  apply (rule sum.mono_neutral_left)
+  apply simp
+  apply (clarsimp simp add: subset_eq)
+  apply clarsimp
+  apply (rule ccontr)
+  apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
+  apply simp
+  unfolding fps_mult_nth
+  apply (rule sum.neutral)
+  apply (clarsimp simp add: not_le)
+  apply (case_tac "x < aa")
+  apply (rule startsby_zero_power_prefix[OF c0, rule_format])
+  apply simp
+  apply (subgoal_tac "n - x < ba")
+  apply (frule_tac k = "ba" in startsby_zero_power_prefix[OF d0, rule_format])
+  apply simp
+  apply arith
+  done
+
+
+lemma sum_pair_less_iff:
+  "sum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
+    sum (\<lambda>s. sum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
+  (is "?l = ?r")
+proof -
+  let ?KM = "{(k,m). k + m \<le> n}"
+  let ?f = "\<lambda>s. UNION {(0::nat)..s} (\<lambda>i. {(i,s - i)})"
+  have th0: "?KM = UNION {0..n} ?f"
+    by auto
+  show "?l = ?r "
+    unfolding th0
+    apply (subst sum.UNION_disjoint)
+    apply auto
+    apply (subst sum.UNION_disjoint)
+    apply auto
+    done
+qed
+
+lemma fps_compose_mult_distrib_lemma:
+  assumes c0: "c$0 = (0::'a::idom)"
+  shows "((a oo c) * (b oo c))$n = sum (\<lambda>s. sum (\<lambda>i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
+  unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
+  unfolding sum_pair_less_iff[where a = "\<lambda>k. a$k" and b="\<lambda>m. b$m" and c="\<lambda>s. (c ^ s)$n" and n = n] ..
+
+lemma fps_compose_mult_distrib:
+  assumes c0: "c $ 0 = (0::'a::idom)"
+  shows "(a * b) oo c = (a oo c) * (b oo c)"
+  apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
+  apply (simp add: fps_compose_nth fps_mult_nth sum_distrib_right)
+  done
+
+lemma fps_compose_prod_distrib:
+  assumes c0: "c$0 = (0::'a::idom)"
+  shows "prod a S oo c = prod (\<lambda>k. a k oo c) S"
+  apply (cases "finite S")
+  apply simp_all
+  apply (induct S rule: finite_induct)
+  apply simp
+  apply (simp add: fps_compose_mult_distrib[OF c0])
+  done
+
+lemma fps_compose_divide:
+  assumes [simp]: "g dvd f" "h $ 0 = 0"
+  shows   "fps_compose f h = fps_compose (f / g :: 'a :: field fps) h * fps_compose g h"
+proof -
+  have "f = (f / g) * g" by simp
+  also have "fps_compose \<dots> h = fps_compose (f / g) h * fps_compose g h"
+    by (subst fps_compose_mult_distrib) simp_all
+  finally show ?thesis .
+qed
+
+lemma fps_compose_divide_distrib:
+  assumes "g dvd f" "h $ 0 = 0" "fps_compose g h \<noteq> 0"
+  shows   "fps_compose (f / g :: 'a :: field fps) h = fps_compose f h / fps_compose g h"
+  using fps_compose_divide[OF assms(1,2)] assms(3) by simp
+
+lemma fps_compose_power:
+  assumes c0: "c$0 = (0::'a::idom)"
+  shows "(a oo c)^n = a^n oo c"
+proof (cases n)
+  case 0
+  then show ?thesis by simp
+next
+  case (Suc m)
+  have th0: "a^n = prod (\<lambda>k. a) {0..m}" "(a oo c) ^ n = prod (\<lambda>k. a oo c) {0..m}"
+    by (simp_all add: prod_constant Suc)
+  then show ?thesis
+    by (simp add: fps_compose_prod_distrib[OF c0])
+qed
+
+lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
+  by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric])
+    
+lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
+  using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
+
+lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
+  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
+
+lemma fps_inverse_compose:
+  assumes b0: "(b$0 :: 'a::field) = 0"
+    and a0: "a$0 \<noteq> 0"
+  shows "inverse a oo b = inverse (a oo b)"
+proof -
+  let ?ia = "inverse a"
+  let ?ab = "a oo b"
+  let ?iab = "inverse ?ab"
+
+  from a0 have ia0: "?ia $ 0 \<noteq> 0" by simp
+  from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
+  have "(?ia oo b) *  (a oo b) = 1"
+    unfolding fps_compose_mult_distrib[OF b0, symmetric]
+    unfolding inverse_mult_eq_1[OF a0]
+    fps_compose_1 ..
+
+  then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
+  then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
+  then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp
+qed
+
+lemma fps_divide_compose:
+  assumes c0: "(c$0 :: 'a::field) = 0"
+    and b0: "b$0 \<noteq> 0"
+  shows "(a/b) oo c = (a oo c) / (b oo c)"
+    using b0 c0 by (simp add: fps_divide_unit fps_inverse_compose fps_compose_mult_distrib)
+
+lemma gp:
+  assumes a0: "a$0 = (0::'a::field)"
+  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)"
+    (is "?one oo a = _")
+proof -
+  have o0: "?one $ 0 \<noteq> 0" by simp
+  have th0: "(1 - X) $ 0 \<noteq> (0::'a)" by simp
+  from fps_inverse_gp[where ?'a = 'a]
+  have "inverse ?one = 1 - X" by (simp add: fps_eq_iff)
+  then have "inverse (inverse ?one) = inverse (1 - X)" by simp
+  then have th: "?one = 1/(1 - X)" unfolding fps_inverse_idempotent[OF o0]
+    by (simp add: fps_divide_def)
+  show ?thesis
+    unfolding th
+    unfolding fps_divide_compose[OF a0 th0]
+    fps_compose_1 fps_compose_sub_distrib X_fps_compose_startby0[OF a0] ..
+qed
+
+lemma fps_const_power [simp]: "fps_const (c::'a::ring_1) ^ n = fps_const (c^n)"
+  by (induct n) auto
+
+lemma fps_compose_radical:
+  assumes b0: "b$0 = (0::'a::field_char_0)"
+    and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
+    and a0: "a$0 \<noteq> 0"
+  shows "fps_radical r (Suc k)  a oo b = fps_radical r (Suc k) (a oo b)"
+proof -
+  let ?r = "fps_radical r (Suc k)"
+  let ?ab = "a oo b"
+  have ab0: "?ab $ 0 = a$0"
+    by (simp add: fps_compose_def)
+  from ab0 a0 ra0 have rab0: "?ab $ 0 \<noteq> 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0"
+    by simp_all
+  have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0"
+    by (simp add: ab0 fps_compose_def)
+  have th0: "(?r a oo b) ^ (Suc k) = a  oo b"
+    unfolding fps_compose_power[OF b0]
+    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  ..
+  from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0]
+  show ?thesis  .
+qed
+
+lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
+  by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc)
+
+lemma fps_const_mult_apply_right:
+  "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
+  by (auto simp add: fps_const_mult_apply_left mult.commute)
+
+lemma fps_compose_assoc:
+  assumes c0: "c$0 = (0::'a::idom)"
+    and b0: "b$0 = 0"
+  shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
+proof -
+  have "?l$n = ?r$n" for n
+  proof -
+    have "?l$n = (sum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
+      by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
+        sum_distrib_left mult.assoc fps_sum_nth)
+    also have "\<dots> = ((sum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
+      by (simp add: fps_compose_sum_distrib)
+    also have "\<dots> = ?r$n"
+      apply (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc)
+      apply (rule sum.cong)
+      apply (rule refl)
+      apply (rule sum.mono_neutral_right)
+      apply (auto simp add: not_le)
+      apply (erule startsby_zero_power_prefix[OF b0, rule_format])
+      done
+    finally show ?thesis .
+  qed
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+
+lemma fps_X_power_compose:
+  assumes a0: "a$0=0"
+  shows "X^k oo a = (a::'a::idom fps)^k"
+  (is "?l = ?r")
+proof (cases k)
+  case 0
+  then show ?thesis by simp
+next
+  case (Suc h)
+  have "?l $ n = ?r $n" for n
+  proof -
+    consider "k > n" | "k \<le> n" by arith
+    then show ?thesis
+    proof cases
+      case 1
+      then show ?thesis
+        using a0 startsby_zero_power_prefix[OF a0] Suc
+        by (simp add: fps_compose_nth del: power_Suc)
+    next
+      case 2
+      then show ?thesis
+        by (simp add: fps_compose_nth mult_delta_left sum.delta)
+    qed
+  qed
+  then show ?thesis
+    unfolding fps_eq_iff by blast
+qed
+
+lemma fps_inv_right:
+  assumes a0: "a$0 = 0"
+    and a1: "a$1 \<noteq> 0"
+  shows "a oo fps_inv a = X"
+proof -
+  let ?ia = "fps_inv a"
+  let ?iaa = "a oo fps_inv a"
+  have th0: "?ia $ 0 = 0"
+    by (simp add: fps_inv_def)
+  have th1: "?iaa $ 0 = 0"
+    using a0 a1 by (simp add: fps_inv_def fps_compose_nth)
+  have th2: "X$0 = 0"
+    by simp
+  from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo X"
+    by simp
+  then have "(a oo fps_inv a) oo a = X oo a"
+    by (simp add: fps_compose_assoc[OF a0 th0] X_fps_compose_startby0[OF a0])
+  with fps_compose_inj_right[OF a0 a1] show ?thesis
+    by simp
+qed
+
+lemma fps_inv_deriv:
+  assumes a0: "a$0 = (0::'a::field)"
+    and a1: "a$1 \<noteq> 0"
+  shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)"
+proof -
+  let ?ia = "fps_inv a"
+  let ?d = "fps_deriv a oo ?ia"
+  let ?dia = "fps_deriv ?ia"
+  have ia0: "?ia$0 = 0"
+    by (simp add: fps_inv_def)
+  have th0: "?d$0 \<noteq> 0"
+    using a1 by (simp add: fps_compose_nth)
+  from fps_inv_right[OF a0 a1] have "?d * ?dia = 1"
+    by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] )
+  then have "inverse ?d * ?d * ?dia = inverse ?d * 1"
+    by simp
+  with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d"
+    by simp
+qed
+
+lemma fps_inv_idempotent:
+  assumes a0: "a$0 = 0"
+    and a1: "a$1 \<noteq> 0"
+  shows "fps_inv (fps_inv a) = a"
+proof -
+  let ?r = "fps_inv"
+  have ra0: "?r a $ 0 = 0"
+    by (simp add: fps_inv_def)
+  from a1 have ra1: "?r a $ 1 \<noteq> 0"
+    by (simp add: fps_inv_def field_simps)
+  have X0: "X$0 = 0"
+    by simp
+  from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
+  then have "?r (?r a) oo ?r a oo a = X oo a"
+    by simp
+  then have "?r (?r a) oo (?r a oo a) = a"
+    unfolding X_fps_compose_startby0[OF a0]
+    unfolding fps_compose_assoc[OF a0 ra0, symmetric] .
+  then show ?thesis
+    unfolding fps_inv[OF a0 a1] by simp
+qed
+
+lemma fps_ginv_ginv:
+  assumes a0: "a$0 = 0"
+    and a1: "a$1 \<noteq> 0"
+    and c0: "c$0 = 0"
+    and  c1: "c$1 \<noteq> 0"
+  shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c"
+proof -
+  let ?r = "fps_ginv"
+  from c0 have rca0: "?r c a $0 = 0"
+    by (simp add: fps_ginv_def)
+  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0"
+    by (simp add: fps_ginv_def field_simps)
+  from fps_ginv[OF rca0 rca1]
+  have "?r b (?r c a) oo ?r c a = b" .
+  then have "?r b (?r c a) oo ?r c a oo a = b oo a"
+    by simp
+  then have "?r b (?r c a) oo (?r c a oo a) = b oo a"
+    apply (subst fps_compose_assoc)
+    using a0 c0
+    apply (auto simp add: fps_ginv_def)
+    done
+  then have "?r b (?r c a) oo c = b oo a"
+    unfolding fps_ginv[OF a0 a1] .
+  then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c"
+    by simp
+  then have "?r b (?r c a) oo (c oo fps_inv c) = b oo a oo fps_inv c"
+    apply (subst fps_compose_assoc)
+    using a0 c0
+    apply (auto simp add: fps_inv_def)
+    done
+  then show ?thesis
+    unfolding fps_inv_right[OF c0 c1] by simp
+qed
+
+lemma fps_ginv_deriv:
+  assumes a0:"a$0 = (0::'a::field)"
+    and a1: "a$1 \<noteq> 0"
+  shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv X a"
+proof -
+  let ?ia = "fps_ginv b a"
+  let ?iXa = "fps_ginv X a"
+  let ?d = "fps_deriv"
+  let ?dia = "?d ?ia"
+  have iXa0: "?iXa $ 0 = 0"
+    by (simp add: fps_ginv_def)
+  have da0: "?d a $ 0 \<noteq> 0"
+    using a1 by simp
+  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b"
+    by simp
+  then have "(?d ?ia oo a) * ?d a = ?d b"
+    unfolding fps_compose_deriv[OF a0] .
+  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)"
+    by simp
+  with a1 have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a"
+    by (simp add: fps_divide_unit)
+  then have "(?d ?ia oo a) oo ?iXa =  (?d b / ?d a) oo ?iXa"
+    unfolding inverse_mult_eq_1[OF da0] by simp
+  then have "?d ?ia oo (a oo ?iXa) =  (?d b / ?d a) oo ?iXa"
+    unfolding fps_compose_assoc[OF iXa0 a0] .
+  then show ?thesis unfolding fps_inv_ginv[symmetric]
+    unfolding fps_inv_right[OF a0 a1] by simp
+qed
+
+lemma fps_compose_linear:
+  "fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * X) = Abs_fps (\<lambda>n. c^n * f $ n)"
+  by (simp add: fps_eq_iff fps_compose_def power_mult_distrib
+                if_distrib sum.delta' cong: if_cong)
+              
+lemma fps_compose_uminus': 
+  "fps_compose f (-X :: 'a :: comm_ring_1 fps) = Abs_fps (\<lambda>n. (-1)^n * f $ n)"
+  using fps_compose_linear[of f "-1"] 
+  by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp
+
+subsection \<open>Elementary series\<close>
+
+subsubsection \<open>Exponential series\<close>
+
+definition "fps_exp x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
+
+lemma fps_exp_deriv[simp]: "fps_deriv (fps_exp a) = fps_const (a::'a::field_char_0) * fps_exp a" 
+  (is "?l = ?r")
+proof -
+  have "?l$n = ?r $ n" for n
+    apply (auto simp add: fps_exp_def field_simps power_Suc[symmetric]
+      simp del: fact_Suc of_nat_Suc power_Suc)
+    apply (simp add: field_simps)
+    done
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+lemma fps_exp_unique_ODE:
+  "fps_deriv a = fps_const c * a \<longleftrightarrow> a = fps_const (a$0) * fps_exp (c::'a::field_char_0)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show ?rhs if ?lhs
+  proof -
+    from that have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
+      by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
+    have th': "a$n = a$0 * c ^ n/ (fact n)" for n
+    proof (induct n)
+      case 0
+      then show ?case by simp
+    next
+      case Suc
+      then show ?case
+        unfolding th
+        using fact_gt_zero
+        apply (simp add: field_simps del: of_nat_Suc fact_Suc)
+        apply simp
+        done
+    qed
+    show ?thesis
+      by (auto simp add: fps_eq_iff fps_const_mult_left fps_exp_def intro: th')
+  qed
+  show ?lhs if ?rhs
+    using that by (metis fps_exp_deriv fps_deriv_mult_const_left mult.left_commute)
+qed
+
+lemma fps_exp_add_mult: "fps_exp (a + b) = fps_exp (a::'a::field_char_0) * fps_exp b" (is "?l = ?r")
+proof -
+  have "fps_deriv ?r = fps_const (a + b) * ?r"
+    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
+  then have "?r = ?l"
+    by (simp only: fps_exp_unique_ODE) (simp add: fps_mult_nth fps_exp_def)
+  then show ?thesis ..
+qed
+
+lemma fps_exp_nth[simp]: "fps_exp a $ n = a^n / of_nat (fact n)"
+  by (simp add: fps_exp_def)
+
+lemma fps_exp_0[simp]: "fps_exp (0::'a::field) = 1"
+  by (simp add: fps_eq_iff power_0_left)
+
+lemma fps_exp_neg: "fps_exp (- a) = inverse (fps_exp (a::'a::field_char_0))"
+proof -
+  from fps_exp_add_mult[of a "- a"] have th0: "fps_exp a * fps_exp (- a) = 1" by simp
+  from fps_inverse_unique[OF th0] show ?thesis by simp
+qed
+
+lemma fps_exp_nth_deriv[simp]: 
+  "fps_nth_deriv n (fps_exp (a::'a::field_char_0)) = (fps_const a)^n * (fps_exp a)"
+  by (induct n) auto
+
+lemma X_compose_fps_exp[simp]: "X oo fps_exp (a::'a::field) = fps_exp a - 1"
+  by (simp add: fps_eq_iff X_fps_compose)
+
+lemma fps_inv_fps_exp_compose:
+  assumes a: "a \<noteq> 0"
+  shows "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = X"
+    and "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = X"
+proof -
+  let ?b = "fps_exp a - 1"
+  have b0: "?b $ 0 = 0"
+    by simp
+  have b1: "?b $ 1 \<noteq> 0"
+    by (simp add: a)
+  from fps_inv[OF b0 b1] show "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = X" .
+  from fps_inv_right[OF b0 b1] show "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = X" .
+qed
+
+lemma fps_exp_power_mult: "(fps_exp (c::'a::field_char_0))^n = fps_exp (of_nat n * c)"
+  by (induct n) (auto simp add: field_simps fps_exp_add_mult)
+
+lemma radical_fps_exp:
+  assumes r: "r (Suc k) 1 = 1"
+  shows "fps_radical r (Suc k) (fps_exp (c::'a::field_char_0)) = fps_exp (c / of_nat (Suc k))"
+proof -
+  let ?ck = "(c / of_nat (Suc k))"
+  let ?r = "fps_radical r (Suc k)"
+  have eq0[simp]: "?ck * of_nat (Suc k) = c" "of_nat (Suc k) * ?ck = c"
+    by (simp_all del: of_nat_Suc)
+  have th0: "fps_exp ?ck ^ (Suc k) = fps_exp c" unfolding fps_exp_power_mult eq0 ..
+  have th: "r (Suc k) (fps_exp c $0) ^ Suc k = fps_exp c $ 0"
+    "r (Suc k) (fps_exp c $ 0) = fps_exp ?ck $ 0" "fps_exp c $ 0 \<noteq> 0" using r by simp_all
+  from th0 radical_unique[where r=r and k=k, OF th] show ?thesis
+    by auto
+qed
+
+lemma fps_exp_compose_linear [simp]: 
+  "fps_exp (d::'a::field_char_0) oo (fps_const c * X) = fps_exp (c * d)"
+  by (simp add: fps_compose_linear fps_exp_def fps_eq_iff power_mult_distrib)
+  
+lemma fps_fps_exp_compose_minus [simp]: 
+  "fps_compose (fps_exp c) (-X) = fps_exp (-c :: 'a :: field_char_0)"
+  using fps_exp_compose_linear[of c "-1 :: 'a"] 
+  unfolding fps_const_neg [symmetric] fps_const_1_eq_1 by simp
+
+lemma fps_exp_eq_iff [simp]: "fps_exp c = fps_exp d \<longleftrightarrow> c = (d :: 'a :: field_char_0)"
+proof
+  assume "fps_exp c = fps_exp d"
+  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] show "c = d" by simp
+qed simp_all
+
+lemma fps_exp_eq_fps_const_iff [simp]: 
+  "fps_exp (c :: 'a :: field_char_0) = fps_const c' \<longleftrightarrow> c = 0 \<and> c' = 1"
+proof
+  assume "c = 0 \<and> c' = 1"
+  thus "fps_exp c = fps_const c'" by (auto simp: fps_eq_iff)
+next
+  assume "fps_exp c = fps_const c'"
+  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] 
+    show "c = 0 \<and> c' = 1" by simp_all
+qed
+
+lemma fps_exp_neq_0 [simp]: "\<not>fps_exp (c :: 'a :: field_char_0) = 0"
+  unfolding fps_const_0_eq_0 [symmetric] fps_exp_eq_fps_const_iff by simp  
+
+lemma fps_exp_eq_1_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = 1 \<longleftrightarrow> c = 0"
+  unfolding fps_const_1_eq_1 [symmetric] fps_exp_eq_fps_const_iff by simp
+    
+lemma fps_exp_neq_numeral_iff [simp]: 
+  "fps_exp (c :: 'a :: field_char_0) = numeral n \<longleftrightarrow> c = 0 \<and> n = Num.One"
+  unfolding numeral_fps_const fps_exp_eq_fps_const_iff by simp
+
+
+subsubsection \<open>Logarithmic series\<close>
+
+lemma Abs_fps_if_0:
+  "Abs_fps (\<lambda>n. if n = 0 then (v::'a::ring_1) else f n) =
+    fps_const v + X * Abs_fps (\<lambda>n. f (Suc n))"
+  by (auto simp add: fps_eq_iff)
+
+definition fps_ln :: "'a::field_char_0 \<Rightarrow> 'a fps"
+  where "fps_ln c = fps_const (1/c) * Abs_fps (\<lambda>n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)"
+
+lemma fps_ln_deriv: "fps_deriv (fps_ln c) = fps_const (1/c) * inverse (1 + X)"
+  unfolding fps_inverse_X_plus1
+  by (simp add: fps_ln_def fps_eq_iff del: of_nat_Suc)
+
+lemma fps_ln_nth: "fps_ln c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
+  by (simp add: fps_ln_def field_simps)
+
+lemma fps_ln_0 [simp]: "fps_ln c $ 0 = 0" by (simp add: fps_ln_def)
+
+lemma fps_ln_fps_exp_inv:
+  fixes a :: "'a::field_char_0"
+  assumes a: "a \<noteq> 0"
+  shows "fps_ln a = fps_inv (fps_exp a - 1)"  (is "?l = ?r")
+proof -
+  let ?b = "fps_exp a - 1"
+  have b0: "?b $ 0 = 0" by simp
+  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
+  have "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) =
+    (fps_const a * (fps_exp a - 1) + fps_const a) oo fps_inv (fps_exp a - 1)"
+    by (simp add: field_simps)
+  also have "\<dots> = fps_const a * (X + 1)"
+    apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
+    apply (simp add: field_simps)
+    done
+  finally have eq: "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_const a * (X + 1)" .
+  from fps_inv_deriv[OF b0 b1, unfolded eq]
+  have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
+    using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
+  then have "fps_deriv ?l = fps_deriv ?r"
+    by (simp add: fps_ln_deriv add.commute fps_divide_def divide_inverse)
+  then show ?thesis unfolding fps_deriv_eq_iff
+    by (simp add: fps_ln_nth fps_inv_def)
+qed
+
+lemma fps_ln_mult_add:
+  assumes c0: "c\<noteq>0"
+    and d0: "d\<noteq>0"
+  shows "fps_ln c + fps_ln d = fps_const (c+d) * fps_ln (c*d)"
+  (is "?r = ?l")
+proof-
+  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
+  have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + X)"
+    by (simp add: fps_ln_deriv fps_const_add[symmetric] algebra_simps del: fps_const_add)
+  also have "\<dots> = fps_deriv ?l"
+    apply (simp add: fps_ln_deriv)
+    apply (simp add: fps_eq_iff eq)
+    done
+  finally show ?thesis
+    unfolding fps_deriv_eq_iff by simp
+qed
+
+lemma X_dvd_fps_ln [simp]: "X dvd fps_ln c"
+proof -
+  have "fps_ln c = X * Abs_fps (\<lambda>n. (-1) ^ n / (of_nat (Suc n) * c))"
+    by (intro fps_ext) (auto simp: fps_ln_def of_nat_diff)
+  thus ?thesis by simp
+qed
+
+
+subsubsection \<open>Binomial series\<close>
+
+definition "fps_binomial a = Abs_fps (\<lambda>n. a gchoose n)"
+
+lemma fps_binomial_nth[simp]: "fps_binomial a $ n = a gchoose n"
+  by (simp add: fps_binomial_def)
+
+lemma fps_binomial_ODE_unique:
+  fixes c :: "'a::field_char_0"
+  shows "fps_deriv a = (fps_const c * a) / (1 + X) \<longleftrightarrow> a = fps_const (a$0) * fps_binomial c"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  let ?da = "fps_deriv a"
+  let ?x1 = "(1 + X):: 'a fps"
+  let ?l = "?x1 * ?da"
+  let ?r = "fps_const c * a"
+
+  have eq: "?l = ?r \<longleftrightarrow> ?lhs"
+  proof -
+    have x10: "?x1 $ 0 \<noteq> 0" by simp
+    have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
+    also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
+      apply (simp only: fps_divide_def  mult.assoc[symmetric] inverse_mult_eq_1[OF x10])
+      apply (simp add: field_simps)
+      done
+    finally show ?thesis .
+  qed
+
+  show ?rhs if ?lhs
+  proof -
+    from eq that have h: "?l = ?r" ..
+    have th0: "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" for n
+    proof -
+      from h have "?l $ n = ?r $ n" by simp
+      then show ?thesis
+        apply (simp add: field_simps del: of_nat_Suc)
+        apply (cases n)
+        apply (simp_all add: field_simps del: of_nat_Suc)
+        done
+    qed
+    have th1: "a $ n = (c gchoose n) * a $ 0" for n
+    proof (induct n)
+      case 0
+      then show ?case by simp
+    next
+      case (Suc m)
+      then show ?case
+        unfolding th0
+        apply (simp add: field_simps del: of_nat_Suc)
+        unfolding mult.assoc[symmetric] gbinomial_mult_1
+        apply (simp add: field_simps)
+        done
+    qed
+    show ?thesis
+      apply (simp add: fps_eq_iff)
+      apply (subst th1)
+      apply (simp add: field_simps)
+      done
+  qed
+
+  show ?lhs if ?rhs
+  proof -
+    have th00: "x * (a $ 0 * y) = a $ 0 * (x * y)" for x y
+      by (simp add: mult.commute)
+    have "?l = ?r"
+      apply (subst \<open>?rhs\<close>)
+      apply (subst (2) \<open>?rhs\<close>)
+      apply (clarsimp simp add: fps_eq_iff field_simps)
+      unfolding mult.assoc[symmetric] th00 gbinomial_mult_1
+      apply (simp add: field_simps gbinomial_mult_1)
+      done
+    with eq show ?thesis ..
+  qed
+qed
+
+lemma fps_binomial_ODE_unique':
+  "(fps_deriv a = fps_const c * a / (1 + X) \<and> a $ 0 = 1) \<longleftrightarrow> (a = fps_binomial c)"
+  by (subst fps_binomial_ODE_unique) auto
+
+lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + X)"
+proof -
+  let ?a = "fps_binomial c"
+  have th0: "?a = fps_const (?a$0) * ?a" by (simp)
+  from iffD2[OF fps_binomial_ODE_unique, OF th0] show ?thesis .
+qed
+
+lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r")
+proof -
+  let ?P = "?r - ?l"
+  let ?b = "fps_binomial"
+  let ?db = "\<lambda>x. fps_deriv (?b x)"
+  have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
+  also have "\<dots> = inverse (1 + X) *
+      (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
+    unfolding fps_binomial_deriv
+    by (simp add: fps_divide_def field_simps)
+  also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
+    by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add)
+  finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
+    by (simp add: fps_divide_def)
+  have "?P = fps_const (?P$0) * ?b (c + d)"
+    unfolding fps_binomial_ODE_unique[symmetric]
+    using th0 by simp
+  then have "?P = 0" by (simp add: fps_mult_nth)
+  then show ?thesis by simp
+qed
+
+lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + X)"
+  (is "?l = inverse ?r")
+proof-
+  have th: "?r$0 \<noteq> 0" by simp
+  have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)"
+    by (simp add: fps_inverse_deriv[OF th] fps_divide_def
+      power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg)
+  have eq: "inverse ?r $ 0 = 1"
+    by (simp add: fps_inverse_def)
+  from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq
+  show ?thesis by (simp add: fps_inverse_def)
+qed
+
+lemma fps_binomial_of_nat: "fps_binomial (of_nat n) = (1 + X :: 'a :: field_char_0 fps) ^ n"
+proof (cases "n = 0")
+  case [simp]: True
+  have "fps_deriv ((1 + X) ^ n :: 'a fps) = 0" by simp
+  also have "\<dots> = fps_const (of_nat n) * (1 + X) ^ n / (1 + X)" by (simp add: fps_binomial_def)
+  finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) simp_all
+next
+  case False
+  have "fps_deriv ((1 + X) ^ n :: 'a fps) = fps_const (of_nat n) * (1 + X) ^ (n - 1)"
+    by (simp add: fps_deriv_power)
+  also have "(1 + X :: 'a fps) $ 0 \<noteq> 0" by simp
+  hence "(1 + X :: 'a fps) \<noteq> 0" by (intro notI) (simp only: , simp)
+  with False have "(1 + X :: 'a fps) ^ (n - 1) = (1 + X) ^ n / (1 + X)"
+    by (cases n) (simp_all )
+  also have "fps_const (of_nat n :: 'a) * ((1 + X) ^ n / (1 + X)) =
+               fps_const (of_nat n) * (1 + X) ^ n / (1 + X)"
+    by (simp add: unit_div_mult_swap)
+  finally show ?thesis
+    by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) (simp_all add: fps_power_nth)
+qed
+
+lemma fps_binomial_0 [simp]: "fps_binomial 0 = 1"
+  using fps_binomial_of_nat[of 0] by simp
+  
+lemma fps_binomial_power: "fps_binomial a ^ n = fps_binomial (of_nat n * a)"
+  by (induction n) (simp_all add: fps_binomial_add_mult ring_distribs)
+
+lemma fps_binomial_1: "fps_binomial 1 = 1 + X"
+  using fps_binomial_of_nat[of 1] by simp
+
+lemma fps_binomial_minus_of_nat:
+  "fps_binomial (- of_nat n) = inverse ((1 + X :: 'a :: field_char_0 fps) ^ n)"
+  by (rule sym, rule fps_inverse_unique)
+     (simp add: fps_binomial_of_nat [symmetric] fps_binomial_add_mult [symmetric])
+
+lemma one_minus_const_X_power:
+  "c \<noteq> 0 \<Longrightarrow> (1 - fps_const c * X) ^ n =
+     fps_compose (fps_binomial (of_nat n)) (-fps_const c * X)"
+  by (subst fps_binomial_of_nat)
+     (simp add: fps_compose_power [symmetric] fps_compose_add_distrib fps_const_neg [symmetric] 
+           del: fps_const_neg)
+
+lemma one_minus_X_const_neg_power:
+  "inverse ((1 - fps_const c * X) ^ n) = 
+       fps_compose (fps_binomial (-of_nat n)) (-fps_const c * X)"
+proof (cases "c = 0")
+  case False
+  thus ?thesis
+  by (subst fps_binomial_minus_of_nat)
+     (simp add: fps_compose_power [symmetric] fps_inverse_compose fps_compose_add_distrib
+                fps_const_neg [symmetric] del: fps_const_neg)
+qed simp
+
+lemma X_plus_const_power:
+  "c \<noteq> 0 \<Longrightarrow> (X + fps_const c) ^ n =
+     fps_const (c^n) * fps_compose (fps_binomial (of_nat n)) (fps_const (inverse c) * X)"
+  by (subst fps_binomial_of_nat)
+     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
+                fps_const_power [symmetric] power_mult_distrib [symmetric] 
+                algebra_simps inverse_mult_eq_1' del: fps_const_power)
+
+lemma X_plus_const_neg_power:
+  "c \<noteq> 0 \<Longrightarrow> inverse ((X + fps_const c) ^ n) =
+     fps_const (inverse c^n) * fps_compose (fps_binomial (-of_nat n)) (fps_const (inverse c) * X)"
+  by (subst fps_binomial_minus_of_nat)
+     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
+                fps_const_power [symmetric] power_mult_distrib [symmetric] fps_inverse_compose 
+                algebra_simps fps_const_inverse [symmetric] fps_inverse_mult [symmetric]
+                fps_inverse_power [symmetric] inverse_mult_eq_1'
+           del: fps_const_power)
+
+
+lemma one_minus_const_X_neg_power':
+  "n > 0 \<Longrightarrow> inverse ((1 - fps_const (c :: 'a :: field_char_0) * X) ^ n) =
+       Abs_fps (\<lambda>k. of_nat ((n + k - 1) choose k) * c^k)"
+  apply (rule fps_ext)
+  apply (subst one_minus_X_const_neg_power, subst fps_const_neg, subst fps_compose_linear)
+  apply (simp add: power_mult_distrib [symmetric] mult.assoc [symmetric] 
+                   gbinomial_minus binomial_gbinomial of_nat_diff)
+  done
+
+text \<open>Vandermonde's Identity as a consequence.\<close>
+lemma gbinomial_Vandermonde:
+  "sum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
+proof -
+  let ?ba = "fps_binomial a"
+  let ?bb = "fps_binomial b"
+  let ?bab = "fps_binomial (a + b)"
+  from fps_binomial_add_mult[of a b] have "?bab $ n = (?ba * ?bb)$n" by simp
+  then show ?thesis by (simp add: fps_mult_nth)
+qed
+
+lemma binomial_Vandermonde:
+  "sum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
+  using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n]
+  by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric]
+                 of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
+
+lemma binomial_Vandermonde_same: "sum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
+  using binomial_Vandermonde[of n n n, symmetric]
+  unfolding mult_2
+  apply (simp add: power2_eq_square)
+  apply (rule sum.cong)
+  apply (auto intro:  binomial_symmetric)
+  done
+
+lemma Vandermonde_pochhammer_lemma:
+  fixes a :: "'a::field_char_0"
+  assumes b: "\<forall>j\<in>{0 ..<n}. b \<noteq> of_nat j"
+  shows "sum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
+      (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} =
+    pochhammer (- (a + b)) n / pochhammer (- b) n"
+  (is "?l = ?r")
+proof -
+  let ?m1 = "\<lambda>m. (- 1 :: 'a) ^ m"
+  let ?f = "\<lambda>m. of_nat (fact m)"
+  let ?p = "\<lambda>(x::'a). pochhammer (- x)"
+  from b have bn0: "?p b n \<noteq> 0"
+    unfolding pochhammer_eq_0_iff by simp
+  have th00:
+    "b gchoose (n - k) =
+        (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
+      (is ?gchoose)
+    "pochhammer (1 + b - of_nat n) k \<noteq> 0"
+      (is ?pochhammer)
+    if kn: "k \<in> {0..n}" for k
+  proof -
+    from kn have "k \<le> n" by simp
+    have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0"
+    proof
+      assume "pochhammer (1 + b - of_nat n) n = 0"
+      then have c: "pochhammer (b - of_nat n + 1) n = 0"
+        by (simp add: algebra_simps)
+      then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j"
+        unfolding pochhammer_eq_0_iff by blast
+      from j have "b = of_nat n - of_nat j - of_nat 1"
+        by (simp add: algebra_simps)
+      then have "b = of_nat (n - j - 1)"
+        using j kn by (simp add: of_nat_diff)
+      with b show False using j by auto
+    qed
+
+    from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \<noteq> 0"
+      by (rule pochhammer_neq_0_mono)
+
+    consider "k = 0 \<or> n = 0" | "k \<noteq> 0" "n \<noteq> 0"
+      by blast
+    then have "b gchoose (n - k) =
+      (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
+    proof cases
+      case 1
+      then show ?thesis
+        using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer)
+    next
+      case neq: 2
+      then obtain m where m: "n = Suc m"
+        by (cases n) auto
+      from neq(1) obtain h where h: "k = Suc h"
+        by (cases k) auto
+      show ?thesis
+      proof (cases "k = n")
+        case True
+        then show ?thesis
+          using pochhammer_minus'[where k=k and b=b]
+          apply (simp add: pochhammer_same)
+          using bn0
+          apply (simp add: field_simps power_add[symmetric])
+          done
+      next
+        case False
+        with kn have kn': "k < n"
+          by simp
+        have m1nk: "?m1 n = prod (\<lambda>i. - 1) {..m}" "?m1 k = prod (\<lambda>i. - 1) {0..h}"
+          by (simp_all add: prod_constant m h)
+        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
+          using bn0 kn
+          unfolding pochhammer_eq_0_iff
+          apply auto
+          apply (erule_tac x= "n - ka - 1" in allE)
+          apply (auto simp add: algebra_simps of_nat_diff)
+          done
+        have eq1: "prod (\<lambda>k. (1::'a) + of_nat m - of_nat k) {..h} =
+          prod of_nat {Suc (m - h) .. Suc m}"
+          using kn' h m
+          by (intro prod.reindex_bij_witness[where i="\<lambda>k. Suc m - k" and j="\<lambda>k. Suc m - k"])
+             (auto simp: of_nat_diff)
+        have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
+          apply (simp add: pochhammer_minus field_simps)
+          using \<open>k \<le> n\<close> apply (simp add: fact_split [of k n])
+          apply (simp add: pochhammer_prod)
+          using prod.atLeast_lessThan_shift_bounds [where ?'a = 'a, of "\<lambda>i. 1 + of_nat i" 0 "n - k" k]
+          apply (auto simp add: of_nat_diff field_simps)
+          done
+        have th20: "?m1 n * ?p b n = prod (\<lambda>i. b - of_nat i) {0..m}"
+          apply (simp add: pochhammer_minus field_simps m)
+          apply (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift)
+          done
+        have th21:"pochhammer (b - of_nat n + 1) k = prod (\<lambda>i. b - of_nat i) {n - k .. n - 1}"
+          using kn apply (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift)
+          using prod.atLeast_atMost_shift_0 [of "m - h" m, where ?'a = 'a]
+          apply (auto simp add: of_nat_diff field_simps)
+          done
+        have "?m1 n * ?p b n =
+          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1} * pochhammer (b - of_nat n + 1) k"
+          using kn' m h unfolding th20 th21 apply simp
+          apply (subst prod.union_disjoint [symmetric])
+          apply auto
+          apply (rule prod.cong)
+          apply auto
+          done
+        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k =
+          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1}"
+          using nz' by (simp add: field_simps)
+        have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) =
+          ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
+          using bnz0
+          by (simp add: field_simps)
+        also have "\<dots> = b gchoose (n - k)"
+          unfolding th1 th2
+          using kn' m h
+          apply (simp add: field_simps gbinomial_mult_fact)
+          apply (rule prod.cong)
+          apply auto
+          done
+        finally show ?thesis by simp
+      qed
+    qed
+    then show ?gchoose and ?pochhammer
+      apply (cases "n = 0")
+      using nz'
+      apply auto
+      done
+  qed
+  have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))"
+    unfolding gbinomial_pochhammer
+    using bn0 by (auto simp add: field_simps)
+  also have "\<dots> = ?l"
+    unfolding gbinomial_Vandermonde[symmetric]
+    apply (simp add: th00)
+    unfolding gbinomial_pochhammer
+    using bn0
+    apply (simp add: sum_distrib_right sum_distrib_left field_simps)
+    done
+  finally show ?thesis by simp
+qed
+
+lemma Vandermonde_pochhammer:
+  fixes a :: "'a::field_char_0"
+  assumes c: "\<forall>i \<in> {0..< n}. c \<noteq> - of_nat i"
+  shows "sum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
+    (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
+proof -
+  let ?a = "- a"
+  let ?b = "c + of_nat n - 1"
+  have h: "\<forall> j \<in>{0..< n}. ?b \<noteq> of_nat j"
+    using c
+    apply (auto simp add: algebra_simps of_nat_diff)
+    apply (erule_tac x = "n - j - 1" in ballE)
+    apply (auto simp add: of_nat_diff algebra_simps)
+    done
+  have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n"
+    unfolding pochhammer_minus
+    by (simp add: algebra_simps)
+  have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n"
+    unfolding pochhammer_minus
+    by simp
+  have nz: "pochhammer c n \<noteq> 0" using c
+    by (simp add: pochhammer_eq_0_iff)
+  from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
+  show ?thesis
+    using nz by (simp add: field_simps sum_distrib_left)
+qed
+
+
+subsubsection \<open>Formal trigonometric functions\<close>
+
+definition "fps_sin (c::'a::field_char_0) =
+  Abs_fps (\<lambda>n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))"
+
+definition "fps_cos (c::'a::field_char_0) =
+  Abs_fps (\<lambda>n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)"
+
+lemma fps_sin_deriv:
+  "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
+  (is "?lhs = ?rhs")
+proof (rule fps_ext)
+  fix n :: nat
+  show "?lhs $ n = ?rhs $ n"
+  proof (cases "even n")
+    case True
+    have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
+    also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
+      using True by (simp add: fps_sin_def)
+    also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
+      unfolding fact_Suc of_nat_mult
+      by (simp add: field_simps del: of_nat_add of_nat_Suc)
+    also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
+      by (simp add: field_simps del: of_nat_add of_nat_Suc)
+    finally show ?thesis
+      using True by (simp add: fps_cos_def field_simps)
+  next
+    case False
+    then show ?thesis
+      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
+  qed
+qed
+
+lemma fps_cos_deriv: "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)"
+  (is "?lhs = ?rhs")
+proof (rule fps_ext)
+  have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n
+    by simp
+  show "?lhs $ n = ?rhs $ n" for n
+  proof (cases "even n")
+    case False
+    then have n0: "n \<noteq> 0" by presburger
+    from False have th1: "Suc ((n - 1) div 2) = Suc n div 2"
+      by (cases n) simp_all
+    have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
+    also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
+      using False by (simp add: fps_cos_def)
+    also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
+      unfolding fact_Suc of_nat_mult
+      by (simp add: field_simps del: of_nat_add of_nat_Suc)
+    also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
+      by (simp add: field_simps del: of_nat_add of_nat_Suc)
+    also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
+      unfolding th0 unfolding th1 by simp
+    finally show ?thesis
+      using False by (simp add: fps_sin_def field_simps)
+  next
+    case True
+    then show ?thesis
+      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
+  qed
+qed
+
+lemma fps_sin_cos_sum_of_squares: "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1"
+  (is "?lhs = _")
+proof -
+  have "fps_deriv ?lhs = 0"
+    apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv)
+    apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
+    done
+  then have "?lhs = fps_const (?lhs $ 0)"
+    unfolding fps_deriv_eq_0_iff .
+  also have "\<dots> = 1"
+    by (auto simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def)
+  finally show ?thesis .
+qed
+
+lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0"
+  unfolding fps_sin_def by simp
+
+lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c"
+  unfolding fps_sin_def by simp
+
+lemma fps_sin_nth_add_2:
+    "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
+  unfolding fps_sin_def
+  apply (cases n)
+  apply simp
+  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
+  apply simp
+  done
+
+lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1"
+  unfolding fps_cos_def by simp
+
+lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0"
+  unfolding fps_cos_def by simp
+
+lemma fps_cos_nth_add_2:
+  "fps_cos c $ (n + 2) = - (c * c * fps_cos c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
+  unfolding fps_cos_def
+  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
+  apply simp
+  done
+
+lemma nat_induct2: "P 0 \<Longrightarrow> P 1 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (n + 2)) \<Longrightarrow> P (n::nat)"
+  unfolding One_nat_def numeral_2_eq_2
+  apply (induct n rule: nat_less_induct)
+  apply (case_tac n)
+  apply simp
+  apply (rename_tac m)
+  apply (case_tac m)
+  apply simp
+  apply (rename_tac k)
+  apply (case_tac k)
+  apply simp_all
+  done
+
+lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2"
+  by simp
+
+lemma eq_fps_sin:
+  assumes 0: "a $ 0 = 0"
+    and 1: "a $ 1 = c"
+    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
+  shows "a = fps_sin c"
+  apply (rule fps_ext)
+  apply (induct_tac n rule: nat_induct2)
+  apply (simp add: 0)
+  apply (simp add: 1 del: One_nat_def)
+  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
+  apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2
+              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
+  apply (subst minus_divide_left)
+  apply (subst nonzero_eq_divide_eq)
+  apply (simp del: of_nat_add of_nat_Suc)
+  apply (simp only: ac_simps)
+  done
+
+lemma eq_fps_cos:
+  assumes 0: "a $ 0 = 1"
+    and 1: "a $ 1 = 0"
+    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
+  shows "a = fps_cos c"
+  apply (rule fps_ext)
+  apply (induct_tac n rule: nat_induct2)
+  apply (simp add: 0)
+  apply (simp add: 1 del: One_nat_def)
+  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
+  apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2
+              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
+  apply (subst minus_divide_left)
+  apply (subst nonzero_eq_divide_eq)
+  apply (simp del: of_nat_add of_nat_Suc)
+  apply (simp only: ac_simps)
+  done
+
+lemma mult_nth_0 [simp]: "(a * b) $ 0 = a $ 0 * b $ 0"
+  by (simp add: fps_mult_nth)
+
+lemma mult_nth_1 [simp]: "(a * b) $ 1 = a $ 0 * b $ 1 + a $ 1 * b $ 0"
+  by (simp add: fps_mult_nth)
+
+lemma fps_sin_add: "fps_sin (a + b) = fps_sin a * fps_cos b + fps_cos a * fps_sin b"
+  apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def)
+  apply (simp del: fps_const_neg fps_const_add fps_const_mult
+              add: fps_const_add [symmetric] fps_const_neg [symmetric]
+                   fps_sin_deriv fps_cos_deriv algebra_simps)
+  done
+
+lemma fps_cos_add: "fps_cos (a + b) = fps_cos a * fps_cos b - fps_sin a * fps_sin b"
+  apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def)
+  apply (simp del: fps_const_neg fps_const_add fps_const_mult
+              add: fps_const_add [symmetric] fps_const_neg [symmetric]
+                   fps_sin_deriv fps_cos_deriv algebra_simps)
+  done
+
+lemma fps_sin_even: "fps_sin (- c) = - fps_sin c"
+  by (auto simp add: fps_eq_iff fps_sin_def)
+
+lemma fps_cos_odd: "fps_cos (- c) = fps_cos c"
+  by (auto simp add: fps_eq_iff fps_cos_def)
+
+definition "fps_tan c = fps_sin c / fps_cos c"
+
+lemma fps_tan_deriv: "fps_deriv (fps_tan c) = fps_const c / (fps_cos c)\<^sup>2"
+proof -
+  have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
+  from this have "fps_cos c \<noteq> 0" by (intro notI) simp
+  hence "fps_deriv (fps_tan c) =
+           fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)"
+    by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps
+                  fps_sin_deriv fps_cos_deriv fps_const_neg[symmetric] div_mult_swap
+             del: fps_const_neg)
+  also note fps_sin_cos_sum_of_squares
+  finally show ?thesis by simp
+qed
+
+text \<open>Connection to @{const "fps_exp"} over the complex numbers --- Euler and de Moivre.\<close>
+
+lemma fps_exp_ii_sin_cos: "fps_exp (\<i> * c) = fps_cos c + fps_const \<i> * fps_sin c"
+  (is "?l = ?r")
+proof -
+  have "?l $ n = ?r $ n" for n
+  proof (cases "even n")
+    case True
+    then obtain m where m: "n = 2 * m" ..
+    show ?thesis
+      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"])
+  next
+    case False
+    then obtain m where m: "n = 2 * m + 1" ..
+    show ?thesis
+      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
+        power_mult power_minus [of "c ^ 2"])
+  qed
+  then show ?thesis
+    by (simp add: fps_eq_iff)
+qed
+
+lemma fps_exp_minus_ii_sin_cos: "fps_exp (- (\<i> * c)) = fps_cos c - fps_const \<i> * fps_sin c"
+  unfolding minus_mult_right fps_exp_ii_sin_cos by (simp add: fps_sin_even fps_cos_odd)
+
+lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
+  by (fact fps_const_sub)
+
+lemma fps_of_int: "fps_const (of_int c) = of_int c"
+  by (induction c) (simp_all add: fps_const_minus [symmetric] fps_of_nat fps_const_neg [symmetric] 
+                             del: fps_const_minus fps_const_neg)
+
+lemma fps_deriv_of_int [simp]: "fps_deriv (of_int n) = 0"
+  by (simp add: fps_of_int [symmetric])
+
+lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a::comm_ring_1)"
+  by (fact numeral_fps_const) (* FIXME: duplicate *)
+
+lemma fps_cos_fps_exp_ii: "fps_cos c = (fps_exp (\<i> * c) + fps_exp (- \<i> * c)) / fps_const 2"
+proof -
+  have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2"
+    by (simp add: numeral_fps_const)
+  show ?thesis
+    unfolding fps_exp_ii_sin_cos minus_mult_commute
+    by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_unit fps_const_inverse th)
+qed
+
+lemma fps_sin_fps_exp_ii: "fps_sin c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / fps_const (2*\<i>)"
+proof -
+  have th: "fps_const \<i> * fps_sin c + fps_const \<i> * fps_sin c = fps_sin c * fps_const (2 * \<i>)"
+    by (simp add: fps_eq_iff numeral_fps_const)
+  show ?thesis
+    unfolding fps_exp_ii_sin_cos minus_mult_commute
+    by (simp add: fps_sin_even fps_cos_odd fps_divide_unit fps_const_inverse th)
+qed
+
+lemma fps_tan_fps_exp_ii:
+  "fps_tan c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / 
+      (fps_const \<i> * (fps_exp (\<i> * c) + fps_exp (- \<i> * c)))"
+  unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii mult_minus_left fps_exp_neg
+  apply (simp add: fps_divide_unit fps_inverse_mult fps_const_mult[symmetric] fps_const_inverse del: fps_const_mult)
+  apply simp
+  done
+
+lemma fps_demoivre:
+  "(fps_cos a + fps_const \<i> * fps_sin a)^n =
+    fps_cos (of_nat n * a) + fps_const \<i> * fps_sin (of_nat n * a)"
+  unfolding fps_exp_ii_sin_cos[symmetric] fps_exp_power_mult
+  by (simp add: ac_simps)
+
+
+subsection \<open>Hypergeometric series\<close>
+
+definition "fps_hypergeo as bs (c::'a::{field_char_0,field}) =
+  Abs_fps (\<lambda>n. (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
+    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
+
+lemma fps_hypergeo_nth[simp]: "fps_hypergeo as bs c $ n =
+  (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
+    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
+  by (simp add: fps_hypergeo_def)
+
+lemma foldl_mult_start:
+  fixes v :: "'a::comm_ring_1"
+  shows "foldl (\<lambda>r x. r * f x) v as * x = foldl (\<lambda>r x. r * f x) (v * x) as "
+  by (induct as arbitrary: x v) (auto simp add: algebra_simps)
+
+lemma foldr_mult_foldl:
+  fixes v :: "'a::comm_ring_1"
+  shows "foldr (\<lambda>x r. r * f x) as v = foldl (\<lambda>r x. r * f x) v as"
+  by (induct as arbitrary: v) (auto simp add: foldl_mult_start)
+
+lemma fps_hypergeo_nth_alt:
+  "fps_hypergeo as bs c $ n = foldr (\<lambda>a r. r * pochhammer a n) as (c ^ n) /
+    foldr (\<lambda>b r. r * pochhammer b n) bs (of_nat (fact n))"
+  by (simp add: foldl_mult_start foldr_mult_foldl)
+
+lemma fps_hypergeo_fps_exp[simp]: "fps_hypergeo [] [] c = fps_exp c"
+  by (simp add: fps_eq_iff)
+
+lemma fps_hypergeo_1_0[simp]: "fps_hypergeo [1] [] c = 1/(1 - fps_const c * X)"
+proof -
+  let ?a = "(Abs_fps (\<lambda>n. 1)) oo (fps_const c * X)"
+  have th0: "(fps_const c * X) $ 0 = 0" by simp
+  show ?thesis unfolding gp[OF th0, symmetric]
+    by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
+      fps_compose_nth power_mult_distrib cond_value_iff sum.delta' cong del: if_weak_cong)
+qed
+
+lemma fps_hypergeo_B[simp]: "fps_hypergeo [-a] [] (- 1) = fps_binomial a"
+  by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps)
+
+lemma fps_hypergeo_0[simp]: "fps_hypergeo as bs c $ 0 = 1"
+  apply simp
+  apply (subgoal_tac "\<forall>as. foldl (\<lambda>(r::'a) (a::'a). r) 1 as = 1")
+  apply auto
+  apply (induct_tac as)
+  apply auto
+  done
+
+lemma foldl_prod_prod:
+  "foldl (\<lambda>(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (\<lambda>r x. r * g x) w as =
+    foldl (\<lambda>r x. r * f x * g x) (v * w) as"
+  by (induct as arbitrary: v w) (auto simp add: algebra_simps)
+
+
+lemma fps_hypergeo_rec:
+  "fps_hypergeo as bs c $ Suc n = ((foldl (\<lambda>r a. r* (a + of_nat n)) c as) /
+    (foldl (\<lambda>r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * fps_hypergeo as bs c $ n"
+  apply (simp del: of_nat_Suc of_nat_add fact_Suc)
+  apply (simp add: foldl_mult_start del: fact_Suc of_nat_Suc)
+  unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc
+  apply (simp add: algebra_simps)
+  done
+
+lemma XD_nth[simp]: "XD a $ n = (if n = 0 then 0 else of_nat n * a$n)"
+  by (simp add: XD_def)
+
+lemma XD_0th[simp]: "XD a $ 0 = 0"
+  by simp
+lemma XD_Suc[simp]:" XD a $ Suc n = of_nat (Suc n) * a $ Suc n"
+  by simp
+
+definition "XDp c a = XD a + fps_const c * a"
+
+lemma XDp_nth[simp]: "XDp c a $ n = (c + of_nat n) * a$n"
+  by (simp add: XDp_def algebra_simps)
+
+lemma XDp_commute: "XDp b \<circ> XDp (c::'a::comm_ring_1) = XDp c \<circ> XDp b"
+  by (auto simp add: XDp_def fun_eq_iff fps_eq_iff algebra_simps)
+
+lemma XDp0 [simp]: "XDp 0 = XD"
+  by (simp add: fun_eq_iff fps_eq_iff)
+
+lemma XDp_fps_integral [simp]: "XDp 0 (fps_integral a c) = X * a"
+  by (simp add: fps_eq_iff fps_integral_def)
+
+lemma fps_hypergeo_minus_nat:
+  "fps_hypergeo [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0,field}) $ k =
+    (if k \<le> n then
+      pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k))
+     else 0)"
+  "fps_hypergeo [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0,field}) $ k =
+    (if k \<le> m then
+      pochhammer (- of_nat m) k * c ^ k / (pochhammer (- of_nat (m + n)) k * of_nat (fact k))
+     else 0)"
+  by (auto simp add: pochhammer_eq_0_iff)
+
+lemma sum_eq_if: "sum f {(n::nat) .. m} = (if m < n then 0 else f n + sum f {n+1 .. m})"
+  apply simp
+  apply (subst sum.insert[symmetric])
+  apply (auto simp add: not_less sum_head_Suc)
+  done
+
+lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))"
+  by (cases n) (simp_all add: pochhammer_rec)
+
+lemma XDp_foldr_nth [simp]: "foldr (\<lambda>c r. XDp c \<circ> r) cs (\<lambda>c. XDp c a) c0 $ n =
+    foldr (\<lambda>c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n"
+  by (induct cs arbitrary: c0) (auto simp add: algebra_simps)
+
+lemma genric_XDp_foldr_nth:
+  assumes f: "\<forall>n c a. f c a $ n = (of_nat n + k c) * a$n"
+  shows "foldr (\<lambda>c r. f c \<circ> r) cs (\<lambda>c. g c a) c0 $ n =
+    foldr (\<lambda>c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
+  by (induct cs arbitrary: c0) (auto simp add: algebra_simps f)
+
+lemma dist_less_imp_nth_equal:
+  assumes "dist f g < inverse (2 ^ i)"
+    and"j \<le> i"
+  shows "f $ j = g $ j"
+proof (rule ccontr)
+  assume "f $ j \<noteq> g $ j"
+  hence "f \<noteq> g" by auto
+  with assms have "i < subdegree (f - g)"
+    by (simp add: if_split_asm dist_fps_def)
+  also have "\<dots> \<le> j"
+    using \<open>f $ j \<noteq> g $ j\<close> by (intro subdegree_leI) simp_all
+  finally show False using \<open>j \<le> i\<close> by simp
+qed
+
+lemma nth_equal_imp_dist_less:
+  assumes "\<And>j. j \<le> i \<Longrightarrow> f $ j = g $ j"
+  shows "dist f g < inverse (2 ^ i)"
+proof (cases "f = g")
+  case True
+  then show ?thesis by simp
+next
+  case False
+  with assms have "dist f g = inverse (2 ^ subdegree (f - g))"
+    by (simp add: if_split_asm dist_fps_def)
+  moreover
+  from assms and False have "i < subdegree (f - g)"
+    by (intro subdegree_greaterI) simp_all
+  ultimately show ?thesis by simp
+qed
+
+lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \<longleftrightarrow> (\<forall>j \<le> i. f $ j = g $ j)"
+  using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast
+
+instance fps :: (comm_ring_1) complete_space
+proof
+  fix X :: "nat \<Rightarrow> 'a fps"
+  assume "Cauchy X"
+  obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j"
+  proof -
+    have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. X M $ j = X m $ j" for i
+    proof -
+      have "0 < inverse ((2::real)^i)" by simp
+      from metric_CauchyD[OF \<open>Cauchy X\<close> this] dist_less_imp_nth_equal
+      show ?thesis by blast
+    qed
+    then show ?thesis using that by metis
+  qed
+
+  show "convergent X"
+  proof (rule convergentI)
+    show "X \<longlonglongrightarrow> Abs_fps (\<lambda>i. X (M i) $ i)"
+      unfolding tendsto_iff
+    proof safe
+      fix e::real assume e: "0 < e"
+      have "(\<lambda>n. inverse (2 ^ n) :: real) \<longlonglongrightarrow> 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all
+      from this and e have "eventually (\<lambda>i. inverse (2 ^ i) < e) sequentially"
+        by (rule order_tendstoD)
+      then obtain i where "inverse (2 ^ i) < e"
+        by (auto simp: eventually_sequentially)
+      have "eventually (\<lambda>x. M i \<le> x) sequentially"
+        by (auto simp: eventually_sequentially)
+      then show "eventually (\<lambda>x. dist (X x) (Abs_fps (\<lambda>i. X (M i) $ i)) < e) sequentially"
+      proof eventually_elim
+        fix x
+        assume x: "M i \<le> x"
+        have "X (M i) $ j = X (M j) $ j" if "j \<le> i" for j
+          using M that by (metis nat_le_linear)
+        with x have "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < inverse (2 ^ i)"
+          using M by (force simp: dist_less_eq_nth_equal)
+        also note \<open>inverse (2 ^ i) < e\<close>
+        finally show "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < e" .
+      qed
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Fraction_Field.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,476 @@
+(*  Title:      HOL/Library/Fraction_Field.thy
+    Author:     Amine Chaieb, University of Cambridge
+*)
+
+section\<open>A formalization of the fraction field of any integral domain;
+         generalization of theory Rat from int to any integral domain\<close>
+
+theory Fraction_Field
+imports Main
+begin
+
+subsection \<open>General fractions construction\<close>
+
+subsubsection \<open>Construction of the type of fractions\<close>
+
+context idom begin
+
+definition fractrel :: "'a \<times> 'a \<Rightarrow> 'a * 'a \<Rightarrow> bool" where
+  "fractrel = (\<lambda>x y. snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"
+
+lemma fractrel_iff [simp]:
+  "fractrel x y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
+  by (simp add: fractrel_def)
+
+lemma symp_fractrel: "symp fractrel"
+  by (simp add: symp_def)
+
+lemma transp_fractrel: "transp fractrel"
+proof (rule transpI, unfold split_paired_all)
+  fix a b a' b' a'' b'' :: 'a
+  assume A: "fractrel (a, b) (a', b')"
+  assume B: "fractrel (a', b') (a'', b'')"
+  have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps)
+  also from A have "a * b' = a' * b" by auto
+  also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps)
+  also from B have "a' * b'' = a'' * b'" by auto
+  also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: ac_simps)
+  finally have "b' * (a * b'') = b' * (a'' * b)" .
+  moreover from B have "b' \<noteq> 0" by auto
+  ultimately have "a * b'' = a'' * b" by simp
+  with A B show "fractrel (a, b) (a'', b'')" by auto
+qed
+
+lemma part_equivp_fractrel: "part_equivp fractrel"
+using _ symp_fractrel transp_fractrel
+by(rule part_equivpI)(rule exI[where x="(0, 1)"]; simp)
+
+end
+
+quotient_type (overloaded) 'a fract = "'a :: idom \<times> 'a" / partial: "fractrel"
+by(rule part_equivp_fractrel)
+
+subsubsection \<open>Representation and basic operations\<close>
+
+lift_definition Fract :: "'a :: idom \<Rightarrow> 'a \<Rightarrow> 'a fract"
+  is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)"
+  by simp
+
+lemma Fract_cases [cases type: fract]:
+  obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0"
+by transfer simp
+
+lemma Fract_induct [case_names Fract, induct type: fract]:
+  "(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q"
+  by (cases q) simp
+
+lemma eq_fract:
+  shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
+    and "\<And>a. Fract a 0 = Fract 0 1"
+    and "\<And>a c. Fract 0 a = Fract 0 c"
+by(transfer; simp)+
+
+instantiation fract :: (idom) comm_ring_1
+begin
+
+lift_definition zero_fract :: "'a fract" is "(0, 1)" by simp
+
+lemma Zero_fract_def: "0 = Fract 0 1"
+by transfer simp
+
+lift_definition one_fract :: "'a fract" is "(1, 1)" by simp
+
+lemma One_fract_def: "1 = Fract 1 1"
+by transfer simp
+
+lift_definition plus_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>q r. (fst q * snd r + fst r * snd q, snd q * snd r)"
+by(auto simp add: algebra_simps)
+
+lemma add_fract [simp]:
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
+by transfer simp
+
+lift_definition uminus_fract :: "'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>x. (- fst x, snd x)"
+by simp
+
+lemma minus_fract [simp]:
+  fixes a b :: "'a::idom"
+  shows "- Fract a b = Fract (- a) b"
+by transfer simp
+
+lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
+  by (cases "b = 0") (simp_all add: eq_fract)
+
+definition diff_fract_def: "q - r = q + - (r::'a fract)"
+
+lemma diff_fract [simp]:
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
+  by (simp add: diff_fract_def)
+
+lift_definition times_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>q r. (fst q * fst r, snd q * snd r)"
+by(simp add: algebra_simps)
+
+lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"
+by transfer simp
+
+lemma mult_fract_cancel:
+  "c \<noteq> 0 \<Longrightarrow> Fract (c * a) (c * b) = Fract a b"
+by transfer simp
+
+instance
+proof
+  fix q r s :: "'a fract"
+  show "(q * r) * s = q * (r * s)"
+    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+  show "q * r = r * q"
+    by (cases q, cases r) (simp add: eq_fract algebra_simps)
+  show "1 * q = q"
+    by (cases q) (simp add: One_fract_def eq_fract)
+  show "(q + r) + s = q + (r + s)"
+    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+  show "q + r = r + q"
+    by (cases q, cases r) (simp add: eq_fract algebra_simps)
+  show "0 + q = q"
+    by (cases q) (simp add: Zero_fract_def eq_fract)
+  show "- q + q = 0"
+    by (cases q) (simp add: Zero_fract_def eq_fract)
+  show "q - r = q + - r"
+    by (cases q, cases r) (simp add: eq_fract)
+  show "(q + r) * s = q * s + r * s"
+    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
+  show "(0::'a fract) \<noteq> 1"
+    by (simp add: Zero_fract_def One_fract_def eq_fract)
+qed
+
+end
+
+lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1"
+  by (induct k) (simp_all add: Zero_fract_def One_fract_def)
+
+lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
+  by (rule of_nat_fract [symmetric])
+
+lemma fract_collapse:
+  "Fract 0 k = 0"
+  "Fract 1 1 = 1"
+  "Fract k 0 = 0"
+by(transfer; simp)+
+
+lemma fract_expand:
+  "0 = Fract 0 1"
+  "1 = Fract 1 1"
+  by (simp_all add: fract_collapse)
+
+lemma Fract_cases_nonzero:
+  obtains (Fract) a b where "q = Fract a b" and "b \<noteq> 0" and "a \<noteq> 0"
+    | (0) "q = 0"
+proof (cases "q = 0")
+  case True
+  then show thesis using 0 by auto
+next
+  case False
+  then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
+  with False have "0 \<noteq> Fract a b" by simp
+  with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
+  with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto
+qed
+
+
+subsubsection \<open>The field of rational numbers\<close>
+
+context idom
+begin
+
+subclass ring_no_zero_divisors ..
+
+end
+
+instantiation fract :: (idom) field
+begin
+
+lift_definition inverse_fract :: "'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>x. if fst x = 0 then (0, 1) else (snd x, fst x)"
+by(auto simp add: algebra_simps)
+
+lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"
+by transfer simp
+
+definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)"
+
+lemma divide_fract [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)"
+  by (simp add: divide_fract_def)
+
+instance
+proof
+  fix q :: "'a fract"
+  assume "q \<noteq> 0"
+  then show "inverse q * q = 1"
+    by (cases q rule: Fract_cases_nonzero)
+      (simp_all add: fract_expand eq_fract mult.commute)
+next
+  fix q r :: "'a fract"
+  show "q div r = q * inverse r" by (simp add: divide_fract_def)
+next
+  show "inverse 0 = (0:: 'a fract)"
+    by (simp add: fract_expand) (simp add: fract_collapse)
+qed
+
+end
+
+
+subsubsection \<open>The ordered field of fractions over an ordered idom\<close>
+
+instantiation fract :: (linordered_idom) linorder
+begin
+
+lemma less_eq_fract_respect:
+  fixes a b a' b' c d c' d' :: 'a
+  assumes neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assumes eq1: "a * b' = a' * b"
+  assumes eq2: "c * d' = c' * d"
+  shows "((a * d) * (b * d) \<le> (c * b) * (b * d)) \<longleftrightarrow> ((a' * d') * (b' * d') \<le> (c' * b') * (b' * d'))"
+proof -
+  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+  {
+    fix a b c d x :: 'a
+    assume x: "x \<noteq> 0"
+    have "?le a b c d = ?le (a * x) (b * x) c d"
+    proof -
+      from x have "0 < x * x"
+        by (auto simp add: zero_less_mult_iff)
+      then have "?le a b c d =
+          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
+        by (simp add: mult_le_cancel_right)
+      also have "... = ?le (a * x) (b * x) c d"
+        by (simp add: ac_simps)
+      finally show ?thesis .
+    qed
+  } note le_factor = this
+
+  let ?D = "b * d" and ?D' = "b' * d'"
+  from neq have D: "?D \<noteq> 0" by simp
+  from neq have "?D' \<noteq> 0" by simp
+  then have "?le a b c d = ?le (a * ?D') (b * ?D') c d"
+    by (rule le_factor)
+  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
+    by (simp add: ac_simps)
+  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
+    by (simp only: eq1 eq2)
+  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
+    by (simp add: ac_simps)
+  also from D have "... = ?le a' b' c' d'"
+    by (rule le_factor [symmetric])
+  finally show "?le a b c d = ?le a' b' c' d'" .
+qed
+
+lift_definition less_eq_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> bool"
+  is "\<lambda>q r. (fst q * snd r) * (snd q * snd r) \<le> (fst r * snd q) * (snd q * snd r)"
+by (clarsimp simp add: less_eq_fract_respect)
+
+definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
+
+lemma le_fract [simp]:
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
+  by transfer simp
+
+lemma less_fract [simp]:
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
+  by (simp add: less_fract_def less_le_not_le ac_simps)
+
+instance
+proof
+  fix q r s :: "'a fract"
+  assume "q \<le> r" and "r \<le> s"
+  then show "q \<le> s"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
+    assume 1: "Fract a b \<le> Fract c d"
+    assume 2: "Fract c d \<le> Fract e f"
+    show "Fract a b \<le> Fract e f"
+    proof -
+      from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
+        by (auto simp add: zero_less_mult_iff linorder_neq_iff)
+      have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
+      proof -
+        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+          by simp
+        with ff show ?thesis by (simp add: mult_le_cancel_right)
+      qed
+      also have "... = (c * f) * (d * f) * (b * b)"
+        by (simp only: ac_simps)
+      also have "... \<le> (e * d) * (d * f) * (b * b)"
+      proof -
+        from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
+          by simp
+        with bb show ?thesis by (simp add: mult_le_cancel_right)
+      qed
+      finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
+        by (simp only: ac_simps)
+      with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by simp
+    qed
+  qed
+next
+  fix q r :: "'a fract"
+  assume "q \<le> r" and "r \<le> q"
+  then show "q = r"
+  proof (induct q, induct r)
+    fix a b c d :: 'a
+    assume neq: "b \<noteq> 0" "d \<noteq> 0"
+    assume 1: "Fract a b \<le> Fract c d"
+    assume 2: "Fract c d \<le> Fract a b"
+    show "Fract a b = Fract c d"
+    proof -
+      from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by simp
+      also have "... \<le> (a * d) * (b * d)"
+      proof -
+        from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
+          by simp
+        then show ?thesis by (simp only: ac_simps)
+      qed
+      finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
+      moreover from neq have "b * d \<noteq> 0" by simp
+      ultimately have "a * d = c * b" by simp
+      with neq show ?thesis by (simp add: eq_fract)
+    qed
+  qed
+next
+  fix q r :: "'a fract"
+  show "q \<le> q"
+    by (induct q) simp
+  show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
+    by (simp only: less_fract_def)
+  show "q \<le> r \<or> r \<le> q"
+    by (induct q, induct r)
+       (simp add: mult.commute, rule linorder_linear)
+qed
+
+end
+
+instantiation fract :: (linordered_idom) linordered_field
+begin
+
+definition abs_fract_def2:
+  "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
+
+definition sgn_fract_def:
+  "sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
+
+theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
+  unfolding abs_fract_def2 not_le [symmetric]
+  by transfer (auto simp add: zero_less_mult_iff le_less)
+
+instance proof
+  fix q r s :: "'a fract"
+  assume "q \<le> r"
+  then show "s + q \<le> s + r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
+    assume le: "Fract a b \<le> Fract c d"
+    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
+    proof -
+      let ?F = "f * f" from neq have F: "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by simp
+      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by (simp add: field_simps)
+    qed
+  qed
+next
+  fix q r s :: "'a fract"
+  assume "q < r" and "0 < s"
+  then show "s * q < s * r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
+    assume le: "Fract a b < Fract c d"
+    assume gt: "0 < Fract e f"
+    show "Fract e f * Fract a b < Fract e f * Fract c d"
+    proof -
+      let ?E = "e * f" and ?F = "f * f"
+      from neq gt have "0 < ?E"
+        by (auto simp add: Zero_fract_def order_less_le eq_fract)
+      moreover from neq have "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
+        by simp
+      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
+        by (simp add: mult_less_cancel_right)
+      with neq show ?thesis
+        by (simp add: ac_simps)
+    qed
+  qed
+qed (fact sgn_fract_def abs_fract_def2)+
+
+end
+
+instantiation fract :: (linordered_idom) distrib_lattice
+begin
+
+definition inf_fract_def:
+  "(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
+
+definition sup_fract_def:
+  "(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
+
+instance
+  by standard (simp_all add: inf_fract_def sup_fract_def max_min_distrib2)
+  
+end
+
+lemma fract_induct_pos [case_names Fract]:
+  fixes P :: "'a::linordered_idom fract \<Rightarrow> bool"
+  assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
+  shows "P q"
+proof (cases q)
+  case (Fract a b)
+  {
+    fix a b :: 'a
+    assume b: "b < 0"
+    have "P (Fract a b)"
+    proof -
+      from b have "0 < - b" by simp
+      then have "P (Fract (- a) (- b))"
+        by (rule step)
+      then show "P (Fract a b)"
+        by (simp add: order_less_imp_not_eq [OF b])
+    qed
+  }
+  with Fract show "P q"
+    by (auto simp add: linorder_neq_iff step)
+qed
+
+lemma zero_less_Fract_iff: "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
+  by (auto simp add: Zero_fract_def zero_less_mult_iff)
+
+lemma Fract_less_zero_iff: "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
+  by (auto simp add: Zero_fract_def mult_less_0_iff)
+
+lemma zero_le_Fract_iff: "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
+  by (auto simp add: Zero_fract_def zero_le_mult_iff)
+
+lemma Fract_le_zero_iff: "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
+  by (auto simp add: Zero_fract_def mult_le_0_iff)
+
+lemma one_less_Fract_iff: "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
+  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
+
+lemma Fract_less_one_iff: "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
+  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
+
+lemma one_le_Fract_iff: "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
+  by (auto simp add: One_fract_def mult_le_cancel_right)
+
+lemma Fract_le_one_iff: "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
+  by (auto simp add: One_fract_def mult_le_cancel_right)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Fundamental_Theorem_Algebra.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,1207 @@
+(* Author: Amine Chaieb, TU Muenchen *)
+
+section \<open>Fundamental Theorem of Algebra\<close>
+
+theory Fundamental_Theorem_Algebra
+imports Polynomial Complex_Main
+begin
+
+subsection \<open>More lemmas about module of complex numbers\<close>
+
+text \<open>The triangle inequality for cmod\<close>
+
+lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
+  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
+
+
+subsection \<open>Basic lemmas about polynomials\<close>
+
+lemma poly_bound_exists:
+  fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
+  shows "\<exists>m. m > 0 \<and> (\<forall>z. norm z \<le> r \<longrightarrow> norm (poly p z) \<le> m)"
+proof (induct p)
+  case 0
+  then show ?case by (rule exI[where x=1]) simp
+next
+  case (pCons c cs)
+  from pCons.hyps obtain m where m: "\<forall>z. norm z \<le> r \<longrightarrow> norm (poly cs z) \<le> m"
+    by blast
+  let ?k = " 1 + norm c + \<bar>r * m\<bar>"
+  have kp: "?k > 0"
+    using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
+  have "norm (poly (pCons c cs) z) \<le> ?k" if H: "norm z \<le> r" for z
+  proof -
+    from m H have th: "norm (poly cs z) \<le> m"
+      by blast
+    from H have rp: "r \<ge> 0"
+      using norm_ge_zero[of z] by arith
+    have "norm (poly (pCons c cs) z) \<le> norm c + norm (z * poly cs z)"
+      using norm_triangle_ineq[of c "z* poly cs z"] by simp
+    also have "\<dots> \<le> norm c + r * m"
+      using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]]
+      by (simp add: norm_mult)
+    also have "\<dots> \<le> ?k"
+      by simp
+    finally show ?thesis .
+  qed
+  with kp show ?case by blast
+qed
+
+
+text \<open>Offsetting the variable in a polynomial gives another of same degree\<close>
+
+definition offset_poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
+  where "offset_poly p h = fold_coeffs (\<lambda>a q. smult h q + pCons a q) p 0"
+
+lemma offset_poly_0: "offset_poly 0 h = 0"
+  by (simp add: offset_poly_def)
+
+lemma offset_poly_pCons:
+  "offset_poly (pCons a p) h =
+    smult h (offset_poly p h) + pCons a (offset_poly p h)"
+  by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def)
+
+lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
+  by (simp add: offset_poly_pCons offset_poly_0)
+
+lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
+  apply (induct p)
+  apply (simp add: offset_poly_0)
+  apply (simp add: offset_poly_pCons algebra_simps)
+  done
+
+lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
+  by (induct p arbitrary: a) (simp, force)
+
+lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
+  apply (safe intro!: offset_poly_0)
+  apply (induct p)
+  apply simp
+  apply (simp add: offset_poly_pCons)
+  apply (frule offset_poly_eq_0_lemma, simp)
+  done
+
+lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
+  apply (induct p)
+  apply (simp add: offset_poly_0)
+  apply (case_tac "p = 0")
+  apply (simp add: offset_poly_0 offset_poly_pCons)
+  apply (simp add: offset_poly_pCons)
+  apply (subst degree_add_eq_right)
+  apply (rule le_less_trans [OF degree_smult_le])
+  apply (simp add: offset_poly_eq_0_iff)
+  apply (simp add: offset_poly_eq_0_iff)
+  done
+
+definition "psize p = (if p = 0 then 0 else Suc (degree p))"
+
+lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
+  unfolding psize_def by simp
+
+lemma poly_offset:
+  fixes p :: "'a::comm_ring_1 poly"
+  shows "\<exists>q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))"
+proof (intro exI conjI)
+  show "psize (offset_poly p a) = psize p"
+    unfolding psize_def
+    by (simp add: offset_poly_eq_0_iff degree_offset_poly)
+  show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
+    by (simp add: poly_offset_poly)
+qed
+
+text \<open>An alternative useful formulation of completeness of the reals\<close>
+lemma real_sup_exists:
+  assumes ex: "\<exists>x. P x"
+    and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
+  shows "\<exists>s::real. \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
+proof
+  from bz have "bdd_above (Collect P)"
+    by (force intro: less_imp_le)
+  then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
+    using ex bz by (subst less_cSup_iff) auto
+qed
+
+
+subsection \<open>Fundamental theorem of algebra\<close>
+
+lemma unimodular_reduce_norm:
+  assumes md: "cmod z = 1"
+  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + \<i>) < 1 \<or> cmod (z - \<i>) < 1"
+proof -
+  obtain x y where z: "z = Complex x y "
+    by (cases z) auto
+  from md z have xy: "x\<^sup>2 + y\<^sup>2 = 1"
+    by (simp add: cmod_def)
+  have False if "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + \<i>) \<ge> 1" "cmod (z - \<i>) \<ge> 1"
+  proof -
+    from that z xy have "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1"
+      by (simp_all add: cmod_def power2_eq_square algebra_simps)
+    then have "\<bar>2 * x\<bar> \<le> 1" "\<bar>2 * y\<bar> \<le> 1"
+      by simp_all
+    then have "\<bar>2 * x\<bar>\<^sup>2 \<le> 1\<^sup>2" "\<bar>2 * y\<bar>\<^sup>2 \<le> 1\<^sup>2"
+      by - (rule power_mono, simp, simp)+
+    then have th0: "4 * x\<^sup>2 \<le> 1" "4 * y\<^sup>2 \<le> 1"
+      by (simp_all add: power_mult_distrib)
+    from add_mono[OF th0] xy show ?thesis
+      by simp
+  qed
+  then show ?thesis
+    unfolding linorder_not_le[symmetric] by blast
+qed
+
+text \<open>Hence we can always reduce modulus of \<open>1 + b z^n\<close> if nonzero\<close>
+lemma reduce_poly_simple:
+  assumes b: "b \<noteq> 0"
+    and n: "n \<noteq> 0"
+  shows "\<exists>z. cmod (1 + b * z^n) < 1"
+  using n
+proof (induct n rule: nat_less_induct)
+  fix n
+  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)"
+  assume n: "n \<noteq> 0"
+  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
+  show "\<exists>z. ?P z n"
+  proof cases
+    assume "even n"
+    then have "\<exists>m. n = 2 * m"
+      by presburger
+    then obtain m where m: "n = 2 * m"
+      by blast
+    from n m have "m \<noteq> 0" "m < n"
+      by presburger+
+    with IH[rule_format, of m] obtain z where z: "?P z m"
+      by blast
+    from z have "?P (csqrt z) n"
+      by (simp add: m power_mult)
+    then show ?thesis ..
+  next
+    assume "odd n"
+    then have "\<exists>m. n = Suc (2 * m)"
+      by presburger+
+    then obtain m where m: "n = Suc (2 * m)"
+      by blast
+    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
+      using b by (simp add: norm_divide)
+    from unimodular_reduce_norm[OF th0] \<open>odd n\<close>
+    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
+      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1")
+      apply (rule_tac x="1" in exI)
+      apply simp
+      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1")
+      apply (rule_tac x="-1" in exI)
+      apply simp
+      apply (cases "cmod (complex_of_real (cmod b) / b + \<i>) < 1")
+      apply (cases "even m")
+      apply (rule_tac x="\<i>" in exI)
+      apply (simp add: m power_mult)
+      apply (rule_tac x="- \<i>" in exI)
+      apply (simp add: m power_mult)
+      apply (cases "even m")
+      apply (rule_tac x="- \<i>" in exI)
+      apply (simp add: m power_mult)
+      apply (auto simp add: m power_mult)
+      apply (rule_tac x="\<i>" in exI)
+      apply (auto simp add: m power_mult)
+      done
+    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1"
+      by blast
+    let ?w = "v / complex_of_real (root n (cmod b))"
+    from odd_real_root_pow[OF \<open>odd n\<close>, of "cmod b"]
+    have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
+      by (simp add: power_divide of_real_power[symmetric])
+    have th2:"cmod (complex_of_real (cmod b) / b) = 1"
+      using b by (simp add: norm_divide)
+    then have th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0"
+      by simp
+    have th4: "cmod (complex_of_real (cmod b) / b) *
+        cmod (1 + b * (v ^ n / complex_of_real (cmod b))) <
+        cmod (complex_of_real (cmod b) / b) * 1"
+      apply (simp only: norm_mult[symmetric] distrib_left)
+      using b v
+      apply (simp add: th2)
+      done
+    from mult_left_less_imp_less[OF th4 th3]
+    have "?P ?w n" unfolding th1 .
+    then show ?thesis ..
+  qed
+qed
+
+text \<open>Bolzano-Weierstrass type property for closed disc in complex plane.\<close>
+
+lemma metric_bound_lemma: "cmod (x - y) \<le> \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
+  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y"]
+  unfolding cmod_def by simp
+
+lemma bolzano_weierstrass_complex_disc:
+  assumes r: "\<forall>n. cmod (s n) \<le> r"
+  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
+proof -
+  from seq_monosub[of "Re \<circ> s"]
+  obtain f where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))"
+    unfolding o_def by blast
+  from seq_monosub[of "Im \<circ> s \<circ> f"]
+  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s (f (g n))))"
+    unfolding o_def by blast
+  let ?h = "f \<circ> g"
+  from r[rule_format, of 0] have rp: "r \<ge> 0"
+    using norm_ge_zero[of "s 0"] by arith
+  have th: "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>"
+  proof
+    fix n
+    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]
+    show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
+  qed
+  have conv1: "convergent (\<lambda>n. Re (s (f n)))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (metis gt_ex le_less_linear less_trans order.trans th)
+    apply (rule f(2))
+    done
+  have th: "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>"
+  proof
+    fix n
+    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]
+    show "\<bar>Im (s n)\<bar> \<le> r + 1"
+      by arith
+  qed
+
+  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (metis gt_ex le_less_linear less_trans order.trans th)
+    apply (rule g(2))
+    done
+
+  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
+    by blast
+  then have x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Re (s (f n)) - x\<bar> < r"
+    unfolding LIMSEQ_iff real_norm_def .
+
+  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
+    by blast
+  then have y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Im (s (f (g n))) - y\<bar> < r"
+    unfolding LIMSEQ_iff real_norm_def .
+  let ?w = "Complex x y"
+  from f(1) g(1) have hs: "subseq ?h"
+    unfolding subseq_def by auto
+  have "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" if "e > 0" for e
+  proof -
+    from that have e2: "e/2 > 0"
+      by simp
+    from x[rule_format, OF e2] y[rule_format, OF e2]
+    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2"
+      and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2"
+      by blast
+    have "cmod (s (?h n) - ?w) < e" if "n \<ge> N1 + N2" for n
+    proof -
+      from that have nN1: "g n \<ge> N1" and nN2: "n \<ge> N2"
+        using seq_suble[OF g(1), of n] by arith+
+      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
+      show ?thesis
+        using metric_bound_lemma[of "s (f (g n))" ?w] by simp
+    qed
+    then show ?thesis by blast
+  qed
+  with hs show ?thesis by blast
+qed
+
+text \<open>Polynomial is continuous.\<close>
+
+lemma poly_cont:
+  fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
+  assumes ep: "e > 0"
+  shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e"
+proof -
+  obtain q where q: "degree q = degree p" "poly q x = poly p (z + x)" for x
+  proof
+    show "degree (offset_poly p z) = degree p"
+      by (rule degree_offset_poly)
+    show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
+      by (rule poly_offset_poly)
+  qed
+  have th: "\<And>w. poly q (w - z) = poly p w"
+    using q(2)[of "w - z" for w] by simp
+  show ?thesis unfolding th[symmetric]
+  proof (induct q)
+    case 0
+    then show ?case
+      using ep by auto
+  next
+    case (pCons c cs)
+    from poly_bound_exists[of 1 "cs"]
+    obtain m where m: "m > 0" "norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m" for z
+      by blast
+    from ep m(1) have em0: "e/m > 0"
+      by (simp add: field_simps)
+    have one0: "1 > (0::real)"
+      by arith
+    from real_lbound_gt_zero[OF one0 em0]
+    obtain d where d: "d > 0" "d < 1" "d < e / m"
+      by blast
+    from d(1,3) m(1) have dm: "d * m > 0" "d * m < e"
+      by (simp_all add: field_simps)
+    show ?case
+    proof (rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
+      fix d w
+      assume H: "d > 0" "d < 1" "d < e/m" "w \<noteq> z" "norm (w - z) < d"
+      then have d1: "norm (w-z) \<le> 1" "d \<ge> 0"
+        by simp_all
+      from H(3) m(1) have dme: "d*m < e"
+        by (simp add: field_simps)
+      from H have th: "norm (w - z) \<le> d"
+        by simp
+      from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
+      show "norm (w - z) * norm (poly cs (w - z)) < e"
+        by simp
+    qed
+  qed
+qed
+
+text \<open>Hence a polynomial attains minimum on a closed disc
+  in the complex plane.\<close>
+lemma poly_minimum_modulus_disc: "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
+proof -
+  show ?thesis
+  proof (cases "r \<ge> 0")
+    case False
+    then show ?thesis
+      by (metis norm_ge_zero order.trans)
+  next
+    case True
+    then have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))"
+      by simp
+    then have mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"
+      by blast
+    have False if "cmod z \<le> r" "cmod (poly p z) = - x" "\<not> x < 1" for x z
+    proof -
+      from that have "- x < 0 "
+        by arith
+      with that(2) norm_ge_zero[of "poly p z"] show ?thesis
+        by simp
+    qed
+    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z"
+      by blast
+    from real_sup_exists[OF mth1 mth2] obtain s where
+      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow> y < s"
+      by blast
+    let ?m = "- s"
+    have s1[unfolded minus_minus]:
+      "(\<exists>z x. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" for y
+      using s[rule_format, of "-y"]
+      unfolding minus_less_iff[of y] equation_minus_iff by blast
+    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
+      by auto
+    have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" for n
+      using s1[rule_format, of "?m + 1/real (Suc n)"] by simp
+    then have th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
+    from choice[OF th] obtain g where
+        g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)"
+      by blast
+    from bolzano_weierstrass_complex_disc[OF g(1)]
+    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
+      by blast
+    {
+      fix w
+      assume wr: "cmod w \<le> r"
+      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
+      {
+        assume e: "?e > 0"
+        then have e2: "?e/2 > 0"
+          by simp
+        from poly_cont[OF e2, of z p] obtain d where
+            d: "d > 0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2"
+          by blast
+        have th1: "cmod(poly p w - poly p z) < ?e / 2" if w: "cmod (w - z) < d" for w
+          using d(2)[rule_format, of w] w e by (cases "w = z") simp_all
+        from fz(2) d(1) obtain N1 where N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d"
+          by blast
+        from reals_Archimedean2[of "2/?e"] obtain N2 :: nat where N2: "2/?e < real N2"
+          by blast
+        have th2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2"
+          using N1[rule_format, of "N1 + N2"] th1 by simp
+        have th0: "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False"
+          for a b e2 m :: real
+          by arith
+        have ath: "m \<le> x \<Longrightarrow> x < m + e \<Longrightarrow> \<bar>x - m\<bar> < e" for m x e :: real
+          by arith
+        from s1m[OF g(1)[rule_format]] have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
+        from seq_suble[OF fz(1), of "N1 + N2"]
+        have th00: "real (Suc (N1 + N2)) \<le> real (Suc (f (N1 + N2)))"
+          by simp
+        have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1 + N2)) > 0"
+          using N2 by auto
+        from frac_le[OF th000 th00]
+        have th00: "?m + 1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))"
+          by simp
+        from g(2)[rule_format, of "f (N1 + N2)"]
+        have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
+        from order_less_le_trans[OF th01 th00]
+        have th32: "cmod (poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
+        from N2 have "2/?e < real (Suc (N1 + N2))"
+          by arith
+        with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
+        have "?e/2 > 1/ real (Suc (N1 + N2))"
+          by (simp add: inverse_eq_divide)
+        with ath[OF th31 th32] have thc1: "\<bar>cmod (poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2"
+          by arith
+        have ath2: "\<bar>a - b\<bar> \<le> c \<Longrightarrow> \<bar>b - m\<bar> \<le> \<bar>a - m\<bar> + c" for a b c m :: real
+          by arith
+        have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar> \<le>
+            cmod (poly p (g (f (N1 + N2))) - poly p z)"
+          by (simp add: norm_triangle_ineq3)
+        from ath2[OF th22, of ?m]
+        have thc2: "2 * (?e/2) \<le>
+            \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)"
+          by simp
+        from th0[OF th2 thc1 thc2] have False .
+      }
+      then have "?e = 0"
+        by auto
+      then have "cmod (poly p z) = ?m"
+        by simp
+      with s1m[OF wr] have "cmod (poly p z) \<le> cmod (poly p w)"
+        by simp
+    }
+    then show ?thesis by blast
+  qed
+qed
+
+text \<open>Nonzero polynomial in z goes to infinity as z does.\<close>
+
+lemma poly_infinity:
+  fixes p:: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
+  assumes ex: "p \<noteq> 0"
+  shows "\<exists>r. \<forall>z. r \<le> norm z \<longrightarrow> d \<le> norm (poly (pCons a p) z)"
+  using ex
+proof (induct p arbitrary: a d)
+  case 0
+  then show ?case by simp
+next
+  case (pCons c cs a d)
+  show ?case
+  proof (cases "cs = 0")
+    case False
+    with pCons.hyps obtain r where r: "\<forall>z. r \<le> norm z \<longrightarrow> d + norm a \<le> norm (poly (pCons c cs) z)"
+      by blast
+    let ?r = "1 + \<bar>r\<bar>"
+    have "d \<le> norm (poly (pCons a (pCons c cs)) z)" if "1 + \<bar>r\<bar> \<le> norm z" for z
+    proof -
+      have r0: "r \<le> norm z"
+        using that by arith
+      from r[rule_format, OF r0] have th0: "d + norm a \<le> 1 * norm(poly (pCons c cs) z)"
+        by arith
+      from that have z1: "norm z \<ge> 1"
+        by arith
+      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
+      have th1: "d \<le> norm(z * poly (pCons c cs) z) - norm a"
+        unfolding norm_mult by (simp add: algebra_simps)
+      from norm_diff_ineq[of "z * poly (pCons c cs) z" a]
+      have th2: "norm (z * poly (pCons c cs) z) - norm a \<le> norm (poly (pCons a (pCons c cs)) z)"
+        by (simp add: algebra_simps)
+      from th1 th2 show ?thesis
+        by arith
+    qed
+    then show ?thesis by blast
+  next
+    case True
+    with pCons.prems have c0: "c \<noteq> 0"
+      by simp
+    have "d \<le> norm (poly (pCons a (pCons c cs)) z)"
+      if h: "(\<bar>d\<bar> + norm a) / norm c \<le> norm z" for z :: 'a
+    proof -
+      from c0 have "norm c > 0"
+        by simp
+      from h c0 have th0: "\<bar>d\<bar> + norm a \<le> norm (z * c)"
+        by (simp add: field_simps norm_mult)
+      have ath: "\<And>mzh mazh ma. mzh \<le> mazh + ma \<Longrightarrow> \<bar>d\<bar> + ma \<le> mzh \<Longrightarrow> d \<le> mazh"
+        by arith
+      from norm_diff_ineq[of "z * c" a] have th1: "norm (z * c) \<le> norm (a + z * c) + norm a"
+        by (simp add: algebra_simps)
+      from ath[OF th1 th0] show ?thesis
+        using True by simp
+    qed
+    then show ?thesis by blast
+  qed
+qed
+
+text \<open>Hence polynomial's modulus attains its minimum somewhere.\<close>
+lemma poly_minimum_modulus: "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
+proof (induct p)
+  case 0
+  then show ?case by simp
+next
+  case (pCons c cs)
+  show ?case
+  proof (cases "cs = 0")
+    case False
+    from poly_infinity[OF False, of "cmod (poly (pCons c cs) 0)" c]
+    obtain r where r: "cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)"
+      if "r \<le> cmod z" for z
+      by blast
+    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>"
+      by arith
+    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
+    obtain v where v: "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)"
+      if "cmod w \<le> \<bar>r\<bar>" for w
+      by blast
+    have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)" if z: "r \<le> cmod z" for z
+      using v[of 0] r[OF z] by simp
+    with v ath[of r] show ?thesis
+      by blast
+  next
+    case True
+    with pCons.hyps show ?thesis
+      by simp
+  qed
+qed
+
+text \<open>Constant function (non-syntactic characterization).\<close>
+definition "constant f \<longleftrightarrow> (\<forall>x y. f x = f y)"
+
+lemma nonconstant_length: "\<not> constant (poly p) \<Longrightarrow> psize p \<ge> 2"
+  by (induct p) (auto simp: constant_def psize_def)
+
+lemma poly_replicate_append: "poly (monom 1 n * p) (x::'a::comm_ring_1) = x^n * poly p x"
+  by (simp add: poly_monom)
+
+text \<open>Decomposition of polynomial, skipping zero coefficients after the first.\<close>
+
+lemma poly_decompose_lemma:
+  assumes nz: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly p z = (0::'a::idom))"
+  shows "\<exists>k a q. a \<noteq> 0 \<and> Suc (psize q + k) = psize p \<and> (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
+  unfolding psize_def
+  using nz
+proof (induct p)
+  case 0
+  then show ?case by simp
+next
+  case (pCons c cs)
+  show ?case
+  proof (cases "c = 0")
+    case True
+    from pCons.hyps pCons.prems True show ?thesis
+      apply auto
+      apply (rule_tac x="k+1" in exI)
+      apply (rule_tac x="a" in exI)
+      apply clarsimp
+      apply (rule_tac x="q" in exI)
+      apply auto
+      done
+  next
+    case False
+    show ?thesis
+      apply (rule exI[where x=0])
+      apply (rule exI[where x=c])
+      apply (auto simp: False)
+      done
+  qed
+qed
+
+lemma poly_decompose:
+  assumes nc: "\<not> constant (poly p)"
+  shows "\<exists>k a q. a \<noteq> (0::'a::idom) \<and> k \<noteq> 0 \<and>
+               psize q + k + 1 = psize p \<and>
+              (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
+  using nc
+proof (induct p)
+  case 0
+  then show ?case
+    by (simp add: constant_def)
+next
+  case (pCons c cs)
+  have "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)"
+  proof
+    assume "\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
+    then have "poly (pCons c cs) x = poly (pCons c cs) y" for x y
+      by (cases "x = 0") auto
+    with pCons.prems show False
+      by (auto simp add: constant_def)
+  qed
+  from poly_decompose_lemma[OF this]
+  show ?case
+    apply clarsimp
+    apply (rule_tac x="k+1" in exI)
+    apply (rule_tac x="a" in exI)
+    apply simp
+    apply (rule_tac x="q" in exI)
+    apply (auto simp add: psize_def split: if_splits)
+    done
+qed
+
+text \<open>Fundamental theorem of algebra\<close>
+
+lemma fundamental_theorem_of_algebra:
+  assumes nc: "\<not> constant (poly p)"
+  shows "\<exists>z::complex. poly p z = 0"
+  using nc
+proof (induct "psize p" arbitrary: p rule: less_induct)
+  case less
+  let ?p = "poly p"
+  let ?ths = "\<exists>z. ?p z = 0"
+
+  from nonconstant_length[OF less(2)] have n2: "psize p \<ge> 2" .
+  from poly_minimum_modulus obtain c where c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)"
+    by blast
+
+  show ?ths
+  proof (cases "?p c = 0")
+    case True
+    then show ?thesis by blast
+  next
+    case False
+    from poly_offset[of p c] obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)"
+      by blast
+    have False if h: "constant (poly q)"
+    proof -
+      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x"
+        by auto
+      have "?p x = ?p y" for x y
+      proof -
+        from th have "?p x = poly q (x - c)"
+          by auto
+        also have "\<dots> = poly q (y - c)"
+          using h unfolding constant_def by blast
+        also have "\<dots> = ?p y"
+          using th by auto
+        finally show ?thesis .
+      qed
+      with less(2) show ?thesis
+        unfolding constant_def by blast
+    qed
+    then have qnc: "\<not> constant (poly q)"
+      by blast
+    from q(2) have pqc0: "?p c = poly q 0"
+      by simp
+    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)"
+      by simp
+    let ?a0 = "poly q 0"
+    from False pqc0 have a00: "?a0 \<noteq> 0"
+      by simp
+    from a00 have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
+      by simp
+    let ?r = "smult (inverse ?a0) q"
+    have lgqr: "psize q = psize ?r"
+      using a00
+      unfolding psize_def degree_def
+      by (simp add: poly_eq_iff)
+    have False if h: "\<And>x y. poly ?r x = poly ?r y"
+    proof -
+      have "poly q x = poly q y" for x y
+      proof -
+        from qr[rule_format, of x] have "poly q x = poly ?r x * ?a0"
+          by auto
+        also have "\<dots> = poly ?r y * ?a0"
+          using h by simp
+        also have "\<dots> = poly q y"
+          using qr[rule_format, of y] by simp
+        finally show ?thesis .
+      qed
+      with qnc show ?thesis
+        unfolding constant_def by blast
+    qed
+    then have rnc: "\<not> constant (poly ?r)"
+      unfolding constant_def by blast
+    from qr[rule_format, of 0] a00 have r01: "poly ?r 0 = 1"
+      by auto
+    have mrmq_eq: "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" for w
+    proof -
+      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
+        using qr[rule_format, of w] a00 by (simp add: divide_inverse ac_simps)
+      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
+        using a00 unfolding norm_divide by (simp add: field_simps)
+      finally show ?thesis .
+    qed
+    from poly_decompose[OF rnc] obtain k a s where
+      kas: "a \<noteq> 0" "k \<noteq> 0" "psize s + k + 1 = psize ?r"
+        "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
+    have "\<exists>w. cmod (poly ?r w) < 1"
+    proof (cases "psize p = k + 1")
+      case True
+      with kas(3) lgqr[symmetric] q(1) have s0: "s = 0"
+        by auto
+      have hth[symmetric]: "cmod (poly ?r w) = cmod (1 + a * w ^ k)" for w
+        using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)
+      from reduce_poly_simple[OF kas(1,2)] show ?thesis
+        unfolding hth by blast
+    next
+      case False note kn = this
+      from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p"
+        by simp
+      have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
+        unfolding constant_def poly_pCons poly_monom
+        using kas(1)
+        apply simp
+        apply (rule exI[where x=0])
+        apply (rule exI[where x=1])
+        apply simp
+        done
+      from kas(1) kas(2) have th02: "k + 1 = psize (pCons 1 (monom a (k - 1)))"
+        by (simp add: psize_def degree_monom_eq)
+      from less(1) [OF k1n [simplified th02] th01]
+      obtain w where w: "1 + w^k * a = 0"
+        unfolding poly_pCons poly_monom
+        using kas(2) by (cases k) (auto simp add: algebra_simps)
+      from poly_bound_exists[of "cmod w" s] obtain m where
+        m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
+      have w0: "w \<noteq> 0"
+        using kas(2) w by (auto simp add: power_0_left)
+      from w have "(1 + w ^ k * a) - 1 = 0 - 1"
+        by simp
+      then have wm1: "w^k * a = - 1"
+        by simp
+      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
+        using norm_ge_zero[of w] w0 m(1)
+        by (simp add: inverse_eq_divide zero_less_mult_iff)
+      with real_lbound_gt_zero[OF zero_less_one] obtain t where
+        t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
+      let ?ct = "complex_of_real t"
+      let ?w = "?ct * w"
+      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w"
+        using kas(1) by (simp add: algebra_simps power_mult_distrib)
+      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
+        unfolding wm1 by simp
+      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) =
+        cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
+        by metis
+      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
+      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)"
+        unfolding norm_of_real by simp
+      have ath: "\<And>x t::real. 0 \<le> x \<Longrightarrow> x < t \<Longrightarrow> t \<le> 1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1"
+        by arith
+      have "t * cmod w \<le> 1 * cmod w"
+        apply (rule mult_mono)
+        using t(1,2)
+        apply auto
+        done
+      then have tw: "cmod ?w \<le> cmod w"
+        using t(1) by (simp add: norm_mult)
+      from t inv0 have "t * (cmod w ^ (k + 1) * m) < 1"
+        by (simp add: field_simps)
+      with zero_less_power[OF t(1), of k] have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
+        by simp
+      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k + 1) * cmod (poly s ?w)))"
+        using w0 t(1)
+        by (simp add: algebra_simps power_mult_distrib norm_power norm_mult)
+      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
+        using t(1,2) m(2)[rule_format, OF tw] w0
+        by auto
+      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k"
+        by simp
+      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
+        by auto
+      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
+      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
+      from th11 th12 have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"
+        by arith
+      then have "cmod (poly ?r ?w) < 1"
+        unfolding kas(4)[rule_format, of ?w] r01 by simp
+      then show ?thesis
+        by blast
+    qed
+    with cq0 q(2) show ?thesis
+      unfolding mrmq_eq not_less[symmetric] by auto
+  qed
+qed
+
+text \<open>Alternative version with a syntactic notion of constant polynomial.\<close>
+
+lemma fundamental_theorem_of_algebra_alt:
+  assumes nc: "\<not> (\<exists>a l. a \<noteq> 0 \<and> l = 0 \<and> p = pCons a l)"
+  shows "\<exists>z. poly p z = (0::complex)"
+  using nc
+proof (induct p)
+  case 0
+  then show ?case by simp
+next
+  case (pCons c cs)
+  show ?case
+  proof (cases "c = 0")
+    case True
+    then show ?thesis by auto
+  next
+    case False
+    have "\<not> constant (poly (pCons c cs))"
+    proof
+      assume nc: "constant (poly (pCons c cs))"
+      from nc[unfolded constant_def, rule_format, of 0]
+      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
+      then have "cs = 0"
+      proof (induct cs)
+        case 0
+        then show ?case by simp
+      next
+        case (pCons d ds)
+        show ?case
+        proof (cases "d = 0")
+          case True
+          then show ?thesis
+            using pCons.prems pCons.hyps by simp
+        next
+          case False
+          from poly_bound_exists[of 1 ds] obtain m where
+            m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
+          have dm: "cmod d / m > 0"
+            using False m(1) by (simp add: field_simps)
+          from real_lbound_gt_zero[OF dm zero_less_one]
+          obtain x where x: "x > 0" "x < cmod d / m" "x < 1"
+            by blast
+          let ?x = "complex_of_real x"
+          from x have cx: "?x \<noteq> 0" "cmod ?x \<le> 1"
+            by simp_all
+          from pCons.prems[rule_format, OF cx(1)]
+          have cth: "cmod (?x*poly ds ?x) = cmod d"
+            by (simp add: eq_diff_eq[symmetric])
+          from m(2)[rule_format, OF cx(2)] x(1)
+          have th0: "cmod (?x*poly ds ?x) \<le> x*m"
+            by (simp add: norm_mult)
+          from x(2) m(1) have "x * m < cmod d"
+            by (simp add: field_simps)
+          with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d"
+            by auto
+          with cth show ?thesis
+            by blast
+        qed
+      qed
+      then show False
+        using pCons.prems False by blast
+    qed
+    then show ?thesis
+      by (rule fundamental_theorem_of_algebra)
+  qed
+qed
+
+
+subsection \<open>Nullstellensatz, degrees and divisibility of polynomials\<close>
+
+lemma nullstellensatz_lemma:
+  fixes p :: "complex poly"
+  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
+    and "degree p = n"
+    and "n \<noteq> 0"
+  shows "p dvd (q ^ n)"
+  using assms
+proof (induct n arbitrary: p q rule: nat_less_induct)
+  fix n :: nat
+  fix p q :: "complex poly"
+  assume IH: "\<forall>m<n. \<forall>p q.
+                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
+                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
+    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
+    and dpn: "degree p = n"
+    and n0: "n \<noteq> 0"
+  from dpn n0 have pne: "p \<noteq> 0" by auto
+  show "p dvd (q ^ n)"
+  proof (cases "\<exists>a. poly p a = 0")
+    case True
+    then obtain a where a: "poly p a = 0" ..
+    have ?thesis if oa: "order a p \<noteq> 0"
+    proof -
+      let ?op = "order a p"
+      from pne have ap: "([:- a, 1:] ^ ?op) dvd p" "\<not> [:- a, 1:] ^ (Suc ?op) dvd p"
+        using order by blast+
+      note oop = order_degree[OF pne, unfolded dpn]
+      show ?thesis
+      proof (cases "q = 0")
+        case True
+        with n0 show ?thesis by (simp add: power_0_left)
+      next
+        case False
+        from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
+        obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
+        from ap(1) obtain s where s: "p = [:- a, 1:] ^ ?op * s"
+          by (rule dvdE)
+        have sne: "s \<noteq> 0"
+          using s pne by auto
+        show ?thesis
+        proof (cases "degree s = 0")
+          case True
+          then obtain k where kpn: "s = [:k:]"
+            by (cases s) (auto split: if_splits)
+          from sne kpn have k: "k \<noteq> 0" by simp
+          let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
+          have "q ^ n = p * ?w"
+            apply (subst r)
+            apply (subst s)
+            apply (subst kpn)
+            using k oop [of a]
+            apply (subst power_mult_distrib)
+            apply simp
+            apply (subst power_add [symmetric])
+            apply simp
+            done
+          then show ?thesis
+            unfolding dvd_def by blast
+        next
+          case False
+          with sne dpn s oa have dsn: "degree s < n"
+            apply auto
+            apply (erule ssubst)
+            apply (simp add: degree_mult_eq degree_linear_power)
+            done
+          have "poly r x = 0" if h: "poly s x = 0" for x
+          proof -
+            have xa: "x \<noteq> a"
+            proof
+              assume "x = a"
+              from h[unfolded this poly_eq_0_iff_dvd] obtain u where u: "s = [:- a, 1:] * u"
+                by (rule dvdE)
+              have "p = [:- a, 1:] ^ (Suc ?op) * u"
+                apply (subst s)
+                apply (subst u)
+                apply (simp only: power_Suc ac_simps)
+                done
+              with ap(2)[unfolded dvd_def] show False
+                by blast
+            qed
+            from h have "poly p x = 0"
+              by (subst s) simp
+            with pq0 have "poly q x = 0"
+              by blast
+            with r xa show ?thesis
+              by auto
+          qed
+          with IH[rule_format, OF dsn, of s r] False have "s dvd (r ^ (degree s))"
+            by blast
+          then obtain u where u: "r ^ (degree s) = s * u" ..
+          then have u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
+            by (simp only: poly_mult[symmetric] poly_power[symmetric])
+          let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
+          from oop[of a] dsn have "q ^ n = p * ?w"
+            apply -
+            apply (subst s)
+            apply (subst r)
+            apply (simp only: power_mult_distrib)
+            apply (subst mult.assoc [where b=s])
+            apply (subst mult.assoc [where a=u])
+            apply (subst mult.assoc [where b=u, symmetric])
+            apply (subst u [symmetric])
+            apply (simp add: ac_simps power_add [symmetric])
+            done
+          then show ?thesis
+            unfolding dvd_def by blast
+        qed
+      qed
+    qed
+    then show ?thesis
+      using a order_root pne by blast
+  next
+    case False
+    with fundamental_theorem_of_algebra_alt[of p]
+    obtain c where ccs: "c \<noteq> 0" "p = pCons c 0"
+      by blast
+    then have pp: "poly p x = c" for x
+      by simp
+    let ?w = "[:1/c:] * (q ^ n)"
+    from ccs have "(q ^ n) = (p * ?w)"
+      by simp
+    then show ?thesis
+      unfolding dvd_def by blast
+  qed
+qed
+
+lemma nullstellensatz_univariate:
+  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow>
+    p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
+proof -
+  consider "p = 0" | "p \<noteq> 0" "degree p = 0" | n where "p \<noteq> 0" "degree p = Suc n"
+    by (cases "degree p") auto
+  then show ?thesis
+  proof cases
+    case p: 1
+    then have eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
+      by (auto simp add: poly_all_0_iff_0)
+    {
+      assume "p dvd (q ^ (degree p))"
+      then obtain r where r: "q ^ (degree p) = p * r" ..
+      from r p have False by simp
+    }
+    with eq p show ?thesis by blast
+  next
+    case dp: 2
+    then obtain k where k: "p = [:k:]" "k \<noteq> 0"
+      by (cases p) (simp split: if_splits)
+    then have th1: "\<forall>x. poly p x \<noteq> 0"
+      by simp
+    from k dp(2) have "q ^ (degree p) = p * [:1/k:]"
+      by (simp add: one_poly_def)
+    then have th2: "p dvd (q ^ (degree p))" ..
+    from dp(1) th1 th2 show ?thesis
+      by blast
+  next
+    case dp: 3
+    have False if dvd: "p dvd (q ^ (Suc n))" and h: "poly p x = 0" "poly q x \<noteq> 0" for x
+    proof -
+      from dvd obtain u where u: "q ^ (Suc n) = p * u" ..
+      from h have "poly (q ^ (Suc n)) x \<noteq> 0"
+        by simp
+      with u h(1) show ?thesis
+        by (simp only: poly_mult) simp
+    qed
+    with dp nullstellensatz_lemma[of p q "degree p"] show ?thesis
+      by auto
+  qed
+qed
+
+text \<open>Useful lemma\<close>
+lemma constant_degree:
+  fixes p :: "'a::{idom,ring_char_0} poly"
+  shows "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
+proof
+  show ?rhs if ?lhs
+  proof -
+    from that[unfolded constant_def, rule_format, of _ "0"]
+    have th: "poly p = poly [:poly p 0:]"
+      by auto
+    then have "p = [:poly p 0:]"
+      by (simp add: poly_eq_poly_eq_iff)
+    then have "degree p = degree [:poly p 0:]"
+      by simp
+    then show ?thesis
+      by simp
+  qed
+  show ?lhs if ?rhs
+  proof -
+    from that obtain k where "p = [:k:]"
+      by (cases p) (simp split: if_splits)
+    then show ?thesis
+      unfolding constant_def by auto
+  qed
+qed
+
+text \<open>Arithmetic operations on multivariate polynomials.\<close>
+
+lemma mpoly_base_conv:
+  fixes x :: "'a::comm_ring_1"
+  shows "0 = poly 0 x" "c = poly [:c:] x" "x = poly [:0,1:] x"
+  by simp_all
+
+lemma mpoly_norm_conv:
+  fixes x :: "'a::comm_ring_1"
+  shows "poly [:0:] x = poly 0 x" "poly [:poly 0 y:] x = poly 0 x"
+  by simp_all
+
+lemma mpoly_sub_conv:
+  fixes x :: "'a::comm_ring_1"
+  shows "poly p x - poly q x = poly p x + -1 * poly q x"
+  by simp
+
+lemma poly_pad_rule: "poly p x = 0 \<Longrightarrow> poly (pCons 0 p) x = 0"
+  by simp
+
+lemma poly_cancel_eq_conv:
+  fixes x :: "'a::field"
+  shows "x = 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> y = 0 \<longleftrightarrow> a * y - b * x = 0"
+  by auto
+
+lemma poly_divides_pad_rule:
+  fixes p:: "('a::comm_ring_1) poly"
+  assumes pq: "p dvd q"
+  shows "p dvd (pCons 0 q)"
+proof -
+  have "pCons 0 q = q * [:0,1:]" by simp
+  then have "q dvd (pCons 0 q)" ..
+  with pq show ?thesis by (rule dvd_trans)
+qed
+
+lemma poly_divides_conv0:
+  fixes p:: "'a::field poly"
+  assumes lgpq: "degree q < degree p"
+    and lq: "p \<noteq> 0"
+  shows "p dvd q \<longleftrightarrow> q = 0" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs
+  then have "q = p * 0" by simp
+  then show ?lhs ..
+next
+  assume l: ?lhs
+  show ?rhs
+  proof (cases "q = 0")
+    case True
+    then show ?thesis by simp
+  next
+    assume q0: "q \<noteq> 0"
+    from l q0 have "degree p \<le> degree q"
+      by (rule dvd_imp_degree_le)
+    with lgpq show ?thesis by simp
+  qed
+qed
+
+lemma poly_divides_conv1:
+  fixes p :: "'a::field poly"
+  assumes a0: "a \<noteq> 0"
+    and pp': "p dvd p'"
+    and qrp': "smult a q - p' = r"
+  shows "p dvd q \<longleftrightarrow> p dvd r" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  from pp' obtain t where t: "p' = p * t" ..
+  show ?rhs if ?lhs
+  proof -
+    from that obtain u where u: "q = p * u" ..
+    have "r = p * (smult a u - t)"
+      using u qrp' [symmetric] t by (simp add: algebra_simps)
+    then show ?thesis ..
+  qed
+  show ?lhs if ?rhs
+  proof -
+    from that obtain u where u: "r = p * u" ..
+    from u [symmetric] t qrp' [symmetric] a0
+    have "q = p * smult (1/a) (u + t)"
+      by (simp add: algebra_simps)
+    then show ?thesis ..
+  qed
+qed
+
+lemma basic_cqe_conv1:
+  "(\<exists>x. poly p x = 0 \<and> poly 0 x \<noteq> 0) \<longleftrightarrow> False"
+  "(\<exists>x. poly 0 x \<noteq> 0) \<longleftrightarrow> False"
+  "(\<exists>x. poly [:c:] x \<noteq> 0) \<longleftrightarrow> c \<noteq> 0"
+  "(\<exists>x. poly 0 x = 0) \<longleftrightarrow> True"
+  "(\<exists>x. poly [:c:] x = 0) \<longleftrightarrow> c = 0"
+  by simp_all
+
+lemma basic_cqe_conv2:
+  assumes l: "p \<noteq> 0"
+  shows "\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)"
+proof -
+  have False if "h \<noteq> 0" "t = 0" and "pCons a (pCons b p) = pCons h t" for h t
+    using l that by simp
+  then have th: "\<not> (\<exists> h t. h \<noteq> 0 \<and> t = 0 \<and> pCons a (pCons b p) = pCons h t)"
+    by blast
+  from fundamental_theorem_of_algebra_alt[OF th] show ?thesis
+    by auto
+qed
+
+lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<longleftrightarrow> p \<noteq> 0"
+  by (metis poly_all_0_iff_0)
+
+lemma basic_cqe_conv3:
+  fixes p q :: "complex poly"
+  assumes l: "p \<noteq> 0"
+  shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<longleftrightarrow> \<not> (pCons a p) dvd (q ^ psize p)"
+proof -
+  from l have dp: "degree (pCons a p) = psize p"
+    by (simp add: psize_def)
+  from nullstellensatz_univariate[of "pCons a p" q] l
+  show ?thesis
+    by (metis dp pCons_eq_0_iff)
+qed
+
+lemma basic_cqe_conv4:
+  fixes p q :: "complex poly"
+  assumes h: "\<And>x. poly (q ^ n) x = poly r x"
+  shows "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
+proof -
+  from h have "poly (q ^ n) = poly r"
+    by auto
+  then have "(q ^ n) = r"
+    by (simp add: poly_eq_poly_eq_iff)
+  then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
+    by simp
+qed
+
+lemma poly_const_conv:
+  fixes x :: "'a::comm_ring_1"
+  shows "poly [:c:] x = y \<longleftrightarrow> c = y"
+  by simp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Normalized_Fraction.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,338 @@
+(*  Title:      HOL/Library/Normalized_Fraction.thy
+    Author:     Manuel Eberl
+*)
+
+theory Normalized_Fraction
+imports 
+  Main 
+  Euclidean_Algorithm
+  Fraction_Field
+begin
+
+definition quot_to_fract :: "'a :: {idom} \<times> 'a \<Rightarrow> 'a fract" where
+  "quot_to_fract = (\<lambda>(a,b). Fraction_Field.Fract a b)"
+
+definition normalize_quot :: "'a :: {ring_gcd,idom_divide} \<times> 'a \<Rightarrow> 'a \<times> 'a" where
+  "normalize_quot = 
+     (\<lambda>(a,b). if b = 0 then (0,1) else let d = gcd a b * unit_factor b in (a div d, b div d))" 
+
+definition normalized_fracts :: "('a :: {ring_gcd,idom_divide} \<times> 'a) set" where
+  "normalized_fracts = {(a,b). coprime a b \<and> unit_factor b = 1}"
+  
+lemma not_normalized_fracts_0_denom [simp]: "(a, 0) \<notin> normalized_fracts"
+  by (auto simp: normalized_fracts_def)
+
+lemma unit_factor_snd_normalize_quot [simp]:
+  "unit_factor (snd (normalize_quot x)) = 1"
+  by (simp add: normalize_quot_def case_prod_unfold Let_def dvd_unit_factor_div
+                mult_unit_dvd_iff unit_factor_mult unit_factor_gcd)
+  
+lemma snd_normalize_quot_nonzero [simp]: "snd (normalize_quot x) \<noteq> 0"
+  using unit_factor_snd_normalize_quot[of x] 
+  by (auto simp del: unit_factor_snd_normalize_quot)
+  
+lemma normalize_quot_aux:
+  fixes a b
+  assumes "b \<noteq> 0"
+  defines "d \<equiv> gcd a b * unit_factor b"
+  shows   "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d"
+          "d dvd a" "d dvd b" "d \<noteq> 0"
+proof -
+  from assms show "d dvd a" "d dvd b"
+    by (simp_all add: d_def mult_unit_dvd_iff)
+  thus "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d \<noteq> 0"
+    by (auto simp: normalize_quot_def Let_def d_def \<open>b \<noteq> 0\<close>)
+qed
+
+lemma normalize_quotE:
+  assumes "b \<noteq> 0"
+  obtains d where "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d"
+                  "d dvd a" "d dvd b" "d \<noteq> 0"
+  using that[OF normalize_quot_aux[OF assms]] .
+  
+lemma normalize_quotE':
+  assumes "snd x \<noteq> 0"
+  obtains d where "fst x = fst (normalize_quot x) * d" "snd x = snd (normalize_quot x) * d"
+                  "d dvd fst x" "d dvd snd x" "d \<noteq> 0"
+proof -
+  from normalize_quotE[OF assms, of "fst x"] guess d .
+  from this show ?thesis unfolding prod.collapse by (intro that[of d])
+qed
+  
+lemma coprime_normalize_quot:
+  "coprime (fst (normalize_quot x)) (snd (normalize_quot x))"
+  by (simp add: normalize_quot_def case_prod_unfold Let_def
+        div_mult_unit2 gcd_div_unit1 gcd_div_unit2 div_gcd_coprime)
+
+lemma normalize_quot_in_normalized_fracts [simp]: "normalize_quot x \<in> normalized_fracts"
+  by (simp add: normalized_fracts_def coprime_normalize_quot case_prod_unfold)
+
+lemma normalize_quot_eq_iff:
+  assumes "b \<noteq> 0" "d \<noteq> 0"
+  shows   "normalize_quot (a,b) = normalize_quot (c,d) \<longleftrightarrow> a * d = b * c"
+proof -
+  define x y where "x = normalize_quot (a,b)" and "y = normalize_quot (c,d)" 
+  from normalize_quotE[OF assms(1), of a] normalize_quotE[OF assms(2), of c]
+    obtain d1 d2 
+      where "a = fst x * d1" "b = snd x * d1" "c = fst y * d2" "d = snd y * d2" "d1 \<noteq> 0" "d2 \<noteq> 0"
+    unfolding x_def y_def by metis
+  hence "a * d = b * c \<longleftrightarrow> fst x * snd y = snd x * fst y" by simp
+  also have "\<dots> \<longleftrightarrow> fst x = fst y \<and> snd x = snd y"
+    by (intro coprime_crossproduct') (simp_all add: x_def y_def coprime_normalize_quot)
+  also have "\<dots> \<longleftrightarrow> x = y" using prod_eqI by blast
+  finally show "x = y \<longleftrightarrow> a * d = b * c" ..
+qed
+
+lemma normalize_quot_eq_iff':
+  assumes "snd x \<noteq> 0" "snd y \<noteq> 0"
+  shows   "normalize_quot x = normalize_quot y \<longleftrightarrow> fst x * snd y = snd x * fst y"
+  using assms by (cases x, cases y, hypsubst) (subst normalize_quot_eq_iff, simp_all)
+
+lemma normalize_quot_id: "x \<in> normalized_fracts \<Longrightarrow> normalize_quot x = x"
+  by (auto simp: normalized_fracts_def normalize_quot_def case_prod_unfold)
+
+lemma normalize_quot_idem [simp]: "normalize_quot (normalize_quot x) = normalize_quot x"
+  by (rule normalize_quot_id) simp_all
+
+lemma fractrel_iff_normalize_quot_eq:
+  "fractrel x y \<longleftrightarrow> normalize_quot x = normalize_quot y \<and> snd x \<noteq> 0 \<and> snd y \<noteq> 0"
+  by (cases x, cases y) (auto simp: fractrel_def normalize_quot_eq_iff)
+  
+lemma fractrel_normalize_quot_left:
+  assumes "snd x \<noteq> 0"
+  shows   "fractrel (normalize_quot x) y \<longleftrightarrow> fractrel x y"
+  using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto
+
+lemma fractrel_normalize_quot_right:
+  assumes "snd x \<noteq> 0"
+  shows   "fractrel y (normalize_quot x) \<longleftrightarrow> fractrel y x"
+  using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto
+
+  
+lift_definition quot_of_fract :: "'a :: {ring_gcd,idom_divide} fract \<Rightarrow> 'a \<times> 'a" 
+    is normalize_quot
+  by (subst (asm) fractrel_iff_normalize_quot_eq) simp_all
+  
+lemma quot_to_fract_quot_of_fract [simp]: "quot_to_fract (quot_of_fract x) = x"
+  unfolding quot_to_fract_def
+proof transfer
+  fix x :: "'a \<times> 'a" assume rel: "fractrel x x"
+  define x' where "x' = normalize_quot x"
+  obtain a b where [simp]: "x = (a, b)" by (cases x)
+  from rel have "b \<noteq> 0" by simp
+  from normalize_quotE[OF this, of a] guess d .
+  hence "a = fst x' * d" "b = snd x' * d" "d \<noteq> 0" "snd x' \<noteq> 0" by (simp_all add: x'_def)
+  thus "fractrel (case x' of (a, b) \<Rightarrow> if b = 0 then (0, 1) else (a, b)) x"
+    by (auto simp add: case_prod_unfold)
+qed
+
+lemma quot_of_fract_quot_to_fract: "quot_of_fract (quot_to_fract x) = normalize_quot x"
+proof (cases "snd x = 0")
+  case True
+  thus ?thesis unfolding quot_to_fract_def
+    by transfer (simp add: case_prod_unfold normalize_quot_def)
+next
+  case False
+  thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold)
+qed
+
+lemma quot_of_fract_quot_to_fract': 
+  "x \<in> normalized_fracts \<Longrightarrow> quot_of_fract (quot_to_fract x) = x"
+  unfolding quot_to_fract_def by transfer (auto simp: normalize_quot_id)
+
+lemma quot_of_fract_in_normalized_fracts [simp]: "quot_of_fract x \<in> normalized_fracts"
+  by transfer simp
+
+lemma normalize_quotI:
+  assumes "a * d = b * c" "b \<noteq> 0" "(c, d) \<in> normalized_fracts"
+  shows   "normalize_quot (a, b) = (c, d)"
+proof -
+  from assms have "normalize_quot (a, b) = normalize_quot (c, d)"
+    by (subst normalize_quot_eq_iff) auto
+  also have "\<dots> = (c, d)" by (intro normalize_quot_id) fact
+  finally show ?thesis .
+qed
+
+lemma td_normalized_fract:
+  "type_definition quot_of_fract quot_to_fract normalized_fracts"
+  by standard (simp_all add: quot_of_fract_quot_to_fract')
+
+lemma quot_of_fract_add_aux:
+  assumes "snd x \<noteq> 0" "snd y \<noteq> 0" 
+  shows   "(fst x * snd y + fst y * snd x) * (snd (normalize_quot x) * snd (normalize_quot y)) =
+             snd x * snd y * (fst (normalize_quot x) * snd (normalize_quot y) +
+             snd (normalize_quot x) * fst (normalize_quot y))"
+proof -
+  from normalize_quotE'[OF assms(1)] guess d . note d = this
+  from normalize_quotE'[OF assms(2)] guess e . note e = this
+  show ?thesis by (simp_all add: d e algebra_simps)
+qed
+
+
+locale fract_as_normalized_quot
+begin
+setup_lifting td_normalized_fract
+end
+
+
+lemma quot_of_fract_add:
+  "quot_of_fract (x + y) = 
+     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y
+      in  normalize_quot (a * d + b * c, b * d))"
+  by transfer (insert quot_of_fract_add_aux, 
+               simp_all add: Let_def case_prod_unfold normalize_quot_eq_iff)
+
+lemma quot_of_fract_uminus:
+  "quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))"
+  by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff)
+
+lemma quot_of_fract_diff:
+  "quot_of_fract (x - y) = 
+     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y
+      in  normalize_quot (a * d - b * c, b * d))" (is "_ = ?rhs")
+proof -
+  have "x - y = x + -y" by simp
+  also have "quot_of_fract \<dots> = ?rhs"
+    by (simp only: quot_of_fract_add quot_of_fract_uminus Let_def case_prod_unfold) simp_all
+  finally show ?thesis .
+qed
+
+lemma normalize_quot_mult_coprime:
+  assumes "coprime a b" "coprime c d" "unit_factor b = 1" "unit_factor d = 1"
+  defines "e \<equiv> fst (normalize_quot (a, d))" and "f \<equiv> snd (normalize_quot (a, d))"
+     and  "g \<equiv> fst (normalize_quot (c, b))" and "h \<equiv> snd (normalize_quot (c, b))"
+  shows   "normalize_quot (a * c, b * d) = (e * g, f * h)"
+proof (rule normalize_quotI)
+  from assms have "b \<noteq> 0" "d \<noteq> 0" by auto
+  from normalize_quotE[OF \<open>b \<noteq> 0\<close>, of c] guess k . note k = this [folded assms]
+  from normalize_quotE[OF \<open>d \<noteq> 0\<close>, of a] guess l . note l = this [folded assms]
+  from k l show "a * c * (f * h) = b * d * (e * g)" by (simp_all)
+  from assms have [simp]: "unit_factor f = 1" "unit_factor h = 1"
+    by simp_all
+  from assms have "coprime e f" "coprime g h" by (simp_all add: coprime_normalize_quot)
+  with k l assms(1,2) show "(e * g, f * h) \<in> normalized_fracts"
+    by (simp add: normalized_fracts_def unit_factor_mult coprime_mul_eq coprime_mul_eq')
+qed (insert assms(3,4), auto)
+
+lemma normalize_quot_mult:
+  assumes "snd x \<noteq> 0" "snd y \<noteq> 0"
+  shows   "normalize_quot (fst x * fst y, snd x * snd y) = normalize_quot 
+             (fst (normalize_quot x) * fst (normalize_quot y),
+              snd (normalize_quot x) * snd (normalize_quot y))"
+proof -
+  from normalize_quotE'[OF assms(1)] guess d . note d = this
+  from normalize_quotE'[OF assms(2)] guess e . note e = this
+  show ?thesis by (simp_all add: d e algebra_simps normalize_quot_eq_iff)
+qed
+
+lemma quot_of_fract_mult:
+  "quot_of_fract (x * y) = 
+     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y;
+          (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b)
+      in  (e*g, f*h))"
+  by transfer (simp_all add: Let_def case_prod_unfold normalize_quot_mult_coprime [symmetric]
+                 coprime_normalize_quot normalize_quot_mult [symmetric])
+  
+lemma normalize_quot_0 [simp]: 
+    "normalize_quot (0, x) = (0, 1)" "normalize_quot (x, 0) = (0, 1)"
+  by (simp_all add: normalize_quot_def)
+  
+lemma normalize_quot_eq_0_iff [simp]: "fst (normalize_quot x) = 0 \<longleftrightarrow> fst x = 0 \<or> snd x = 0"
+  by (auto simp: normalize_quot_def case_prod_unfold Let_def div_mult_unit2 dvd_div_eq_0_iff)
+  find_theorems "_ div _ = 0"
+  
+lemma fst_quot_of_fract_0_imp: "fst (quot_of_fract x) = 0 \<Longrightarrow> snd (quot_of_fract x) = 1"
+  by transfer auto
+
+lemma normalize_quot_swap:
+  assumes "a \<noteq> 0" "b \<noteq> 0"
+  defines "a' \<equiv> fst (normalize_quot (a, b))" and "b' \<equiv> snd (normalize_quot (a, b))"
+  shows   "normalize_quot (b, a) = (b' div unit_factor a', a' div unit_factor a')"
+proof (rule normalize_quotI)
+  from normalize_quotE[OF assms(2), of a] guess d . note d = this [folded assms(3,4)]
+  show "b * (a' div unit_factor a') = a * (b' div unit_factor a')"
+    using assms(1,2) d 
+    by (simp add: div_unit_factor [symmetric] unit_div_mult_swap mult_ac del: div_unit_factor)
+  have "coprime a' b'" by (simp add: a'_def b'_def coprime_normalize_quot)
+  thus "(b' div unit_factor a', a' div unit_factor a') \<in> normalized_fracts"
+    using assms(1,2) d by (auto simp: normalized_fracts_def gcd_div_unit1 gcd_div_unit2 gcd.commute)
+qed fact+
+  
+lemma quot_of_fract_inverse:
+  "quot_of_fract (inverse x) = 
+     (let (a,b) = quot_of_fract x; d = unit_factor a 
+      in  if d = 0 then (0, 1) else (b div d, a div d))"
+proof (transfer, goal_cases)
+  case (1 x)
+  from normalize_quot_swap[of "fst x" "snd x"] show ?case
+    by (auto simp: Let_def case_prod_unfold)
+qed
+
+lemma normalize_quot_div_unit_left:
+  fixes x y u
+  assumes "is_unit u"
+  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
+  shows "normalize_quot (x div u, y) = (x' div u, y')"
+proof (cases "y = 0")
+  case False
+  from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)]
+  from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot)
+  with False d \<open>is_unit u\<close> show ?thesis
+    by (intro normalize_quotI)
+       (auto simp: normalized_fracts_def unit_div_mult_swap unit_div_commute unit_div_cancel
+          gcd_div_unit1)
+qed (simp_all add: assms)
+
+lemma normalize_quot_div_unit_right:
+  fixes x y u
+  assumes "is_unit u"
+  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
+  shows "normalize_quot (x, y div u) = (x' * u, y')"
+proof (cases "y = 0")
+  case False
+  from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)]
+  from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot)
+  with False d \<open>is_unit u\<close> show ?thesis
+    by (intro normalize_quotI)
+       (auto simp: normalized_fracts_def unit_div_mult_swap unit_div_commute unit_div_cancel
+          gcd_mult_unit1 unit_div_eq_0_iff mult.assoc [symmetric])
+qed (simp_all add: assms)
+
+lemma normalize_quot_normalize_left:
+  fixes x y u
+  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
+  shows "normalize_quot (normalize x, y) = (x' div unit_factor x, y')"
+  using normalize_quot_div_unit_left[of "unit_factor x" x y]
+  by (cases "x = 0") (simp_all add: assms)
+  
+lemma normalize_quot_normalize_right:
+  fixes x y u
+  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
+  shows "normalize_quot (x, normalize y) = (x' * unit_factor y, y')"
+  using normalize_quot_div_unit_right[of "unit_factor y" x y]
+  by (cases "y = 0") (simp_all add: assms)
+  
+lemma quot_of_fract_0 [simp]: "quot_of_fract 0 = (0, 1)"
+  by transfer auto
+
+lemma quot_of_fract_1 [simp]: "quot_of_fract 1 = (1, 1)"
+  by transfer (rule normalize_quotI, simp_all add: normalized_fracts_def)
+
+lemma quot_of_fract_divide:
+  "quot_of_fract (x / y) = (if y = 0 then (0, 1) else
+     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y;
+          (e,f) = normalize_quot (a,c); (g,h) = normalize_quot (d,b)
+      in  (e * g, f * h)))" (is "_ = ?rhs")
+proof (cases "y = 0")
+  case False
+  hence A: "fst (quot_of_fract y) \<noteq> 0" by transfer auto
+  have "x / y = x * inverse y" by (simp add: divide_inverse)
+  also from False A have "quot_of_fract \<dots> = ?rhs"
+    by (simp only: quot_of_fract_mult quot_of_fract_inverse)
+       (simp_all add: Let_def case_prod_unfold fst_quot_of_fract_0_imp
+          normalize_quot_div_unit_left normalize_quot_div_unit_right 
+          normalize_quot_normalize_right normalize_quot_normalize_left)
+  finally show ?thesis .
+qed simp_all
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Polynomial.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,4453 @@
+(*  Title:      HOL/Library/Polynomial.thy
+    Author:     Brian Huffman
+    Author:     Clemens Ballarin
+    Author:     Amine Chaieb
+    Author:     Florian Haftmann
+*)
+
+section \<open>Polynomials as type over a ring structure\<close>
+
+theory Polynomial
+imports
+  "~~/src/HOL/Deriv"
+  "~~/src/HOL/Library/More_List"
+  "~~/src/HOL/Library/Infinite_Set"
+begin
+
+subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
+
+definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixr "##" 65)
+  where "x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)"
+
+lemma cCons_0_Nil_eq [simp]: "0 ## [] = []"
+  by (simp add: cCons_def)
+
+lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys"
+  by (simp add: cCons_def)
+
+lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys"
+  by (simp add: cCons_def)
+
+lemma cCons_not_0_eq [simp]: "x \<noteq> 0 \<Longrightarrow> x ## xs = x # xs"
+  by (simp add: cCons_def)
+
+lemma strip_while_not_0_Cons_eq [simp]:
+  "strip_while (\<lambda>x. x = 0) (x # xs) = x ## strip_while (\<lambda>x. x = 0) xs"
+proof (cases "x = 0")
+  case False
+  then show ?thesis by simp
+next
+  case True
+  show ?thesis
+  proof (induct xs rule: rev_induct)
+    case Nil
+    with True show ?case by simp
+  next
+    case (snoc y ys)
+    then show ?case
+      by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons)
+  qed
+qed
+
+lemma tl_cCons [simp]: "tl (x ## xs) = xs"
+  by (simp add: cCons_def)
+
+
+subsection \<open>Definition of type \<open>poly\<close>\<close>
+
+typedef (overloaded) 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
+  morphisms coeff Abs_poly
+  by (auto intro!: ALL_MOST)
+
+setup_lifting type_definition_poly
+
+lemma poly_eq_iff: "p = q \<longleftrightarrow> (\<forall>n. coeff p n = coeff q n)"
+  by (simp add: coeff_inject [symmetric] fun_eq_iff)
+
+lemma poly_eqI: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q"
+  by (simp add: poly_eq_iff)
+
+lemma MOST_coeff_eq_0: "\<forall>\<^sub>\<infinity> n. coeff p n = 0"
+  using coeff [of p] by simp
+
+
+subsection \<open>Degree of a polynomial\<close>
+
+definition degree :: "'a::zero poly \<Rightarrow> nat"
+  where "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)"
+
+lemma coeff_eq_0:
+  assumes "degree p < n"
+  shows "coeff p n = 0"
+proof -
+  have "\<exists>n. \<forall>i>n. coeff p i = 0"
+    using MOST_coeff_eq_0 by (simp add: MOST_nat)
+  then have "\<forall>i>degree p. coeff p i = 0"
+    unfolding degree_def by (rule LeastI_ex)
+  with assms show ?thesis by simp
+qed
+
+lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p"
+  by (erule contrapos_np, rule coeff_eq_0, simp)
+
+lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n"
+  unfolding degree_def by (erule Least_le)
+
+lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0"
+  unfolding degree_def by (drule not_less_Least, simp)
+
+
+subsection \<open>The zero polynomial\<close>
+
+instantiation poly :: (zero) zero
+begin
+
+lift_definition zero_poly :: "'a poly"
+  is "\<lambda>_. 0"
+  by (rule MOST_I) simp
+
+instance ..
+
+end
+
+lemma coeff_0 [simp]: "coeff 0 n = 0"
+  by transfer rule
+
+lemma degree_0 [simp]: "degree 0 = 0"
+  by (rule order_antisym [OF degree_le le0]) simp
+
+lemma leading_coeff_neq_0:
+  assumes "p \<noteq> 0"
+  shows "coeff p (degree p) \<noteq> 0"
+proof (cases "degree p")
+  case 0
+  from \<open>p \<noteq> 0\<close> obtain n where "coeff p n \<noteq> 0"
+    by (auto simp add: poly_eq_iff)
+  then have "n \<le> degree p"
+    by (rule le_degree)
+  with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close> show "coeff p (degree p) \<noteq> 0"
+    by simp
+next
+  case (Suc n)
+  from \<open>degree p = Suc n\<close> have "n < degree p"
+    by simp
+  then have "\<exists>i>n. coeff p i \<noteq> 0"
+    by (rule less_degree_imp)
+  then obtain i where "n < i" and "coeff p i \<noteq> 0"
+    by blast
+  from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i"
+    by simp
+  also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p"
+    by (rule le_degree)
+  finally have "degree p = i" .
+  with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp
+qed
+
+lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0"
+  by (cases "p = 0") (simp_all 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)
+  from \<open>degree p \<le> n\<close> have "\<forall>i>n. coeff p i = 0"
+    by (simp add: coeff_eq_0)
+  with \<open>coeff p n = 0\<close> have "\<forall>i\<ge>n. coeff p i = 0"
+    by (simp add: le_less)
+  with \<open>n = Suc m\<close> have "\<forall>i>m. coeff p i = 0"
+    by (simp add: less_eq_Suc_le)
+  then have "degree p \<le> m"
+    by (rule degree_le)
+  with \<open>n = Suc m\<close> have "degree p < n"
+    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>
+
+lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  is "\<lambda>a p. case_nat a (coeff p)"
+  by (rule MOST_SucD) (simp add: MOST_coeff_eq_0)
+
+lemmas coeff_pCons = pCons.rep_eq
+
+lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a"
+  by transfer simp
+
+lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n"
+  by (simp add: coeff_pCons)
+
+lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)"
+  by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split)
+
+lemma degree_pCons_eq: "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)"
+  apply (rule order_antisym [OF degree_pCons_le])
+  apply (rule le_degree, simp)
+  done
+
+lemma degree_pCons_0: "degree (pCons a 0) = 0"
+  apply (rule order_antisym [OF _ le0])
+  apply (rule degree_le, simp add: coeff_pCons split: nat.split)
+  done
+
+lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))"
+  apply (cases "p = 0", simp_all)
+  apply (rule order_antisym [OF _ le0])
+  apply (rule degree_le, simp add: coeff_pCons split: nat.split)
+  apply (rule order_antisym [OF degree_pCons_le])
+  apply (rule le_degree, simp)
+  done
+
+lemma pCons_0_0 [simp]: "pCons 0 0 = 0"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q"
+proof safe
+  assume "pCons a p = pCons b q"
+  then have "coeff (pCons a p) 0 = coeff (pCons b q) 0"
+    by simp
+  then show "a = b"
+    by simp
+next
+  assume "pCons a p = pCons b q"
+  then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n
+    by simp
+  then show "p = q"
+    by (simp add: poly_eq_iff)
+qed
+
+lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0"
+  using pCons_eq_iff [of a p 0 0] by simp
+
+lemma pCons_cases [cases type: poly]:
+  obtains (pCons) a q where "p = pCons a q"
+proof
+  show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))"
+    by transfer
+      (simp_all add: MOST_inj[where f=Suc and P="\<lambda>n. p n = 0" for p] fun_eq_iff Abs_poly_inverse
+        split: nat.split)
+qed
+
+lemma pCons_induct [case_names 0 pCons, induct type: poly]:
+  assumes zero: "P 0"
+  assumes pCons: "\<And>a p. a \<noteq> 0 \<or> p \<noteq> 0 \<Longrightarrow> P p \<Longrightarrow> P (pCons a p)"
+  shows "P p"
+proof (induct p rule: measure_induct_rule [where f=degree])
+  case (less p)
+  obtain a q where "p = pCons a q" by (rule pCons_cases)
+  have "P q"
+  proof (cases "q = 0")
+    case True
+    then show "P q" by (simp add: zero)
+  next
+    case False
+    then have "degree (pCons a q) = Suc (degree q)"
+      by (rule degree_pCons_eq)
+    with \<open>p = pCons a q\<close> have "degree q < degree p"
+      by simp
+    then show "P q"
+      by (rule less.hyps)
+  qed
+  have "P (pCons a q)"
+  proof (cases "a \<noteq> 0 \<or> q \<noteq> 0")
+    case True
+    with \<open>P q\<close> show ?thesis by (auto intro: pCons)
+  next
+    case False
+    with zero show ?thesis by simp
+  qed
+  with \<open>p = pCons a q\<close> show ?case
+    by simp
+qed
+
+lemma degree_eq_zeroE:
+  fixes p :: "'a::zero poly"
+  assumes "degree p = 0"
+  obtains a where "p = pCons a 0"
+proof -
+  obtain a q where p: "p = pCons a q"
+    by (cases p)
+  with assms have "q = 0"
+    by (cases "q = 0") simp_all
+  with p have "p = pCons a 0"
+    by simp
+  then show thesis ..
+qed
+
+
+subsection \<open>Quickcheck generator for polynomials\<close>
+
+quickcheck_generator poly constructors: "0 :: _ poly", pCons
+
+
+subsection \<open>List-style syntax for polynomials\<close>
+
+syntax "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
+translations
+  "[:x, xs:]" \<rightleftharpoons> "CONST pCons x [:xs:]"
+  "[:x:]" \<rightleftharpoons> "CONST pCons x 0"
+  "[:x:]" \<leftharpoondown> "CONST pCons x (_constrain 0 t)"
+
+
+subsection \<open>Representation of polynomials by lists of coefficients\<close>
+
+primrec Poly :: "'a::zero list \<Rightarrow> 'a poly"
+  where
+    [code_post]: "Poly [] = 0"
+  | [code_post]: "Poly (a # as) = pCons a (Poly as)"
+
+lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0"
+  by (induct n) simp_all
+
+lemma Poly_eq_0: "Poly as = 0 \<longleftrightarrow> (\<exists>n. as = replicate n 0)"
+  by (induct as) (auto simp add: Cons_replicate_eq)
+
+lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as"
+  by (induct as) simp_all
+
+lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as"
+  using Poly_append_replicate_zero [of as 1] by simp
+
+lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)"
+  by (simp add: cCons_def)
+
+lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \<Longrightarrow> Poly (rev (tl as)) = Poly (rev as)"
+  by (cases as) simp_all
+
+lemma degree_Poly: "degree (Poly xs) \<le> length xs"
+  by (induct xs) simp_all
+
+lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs"
+  by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits)
+
+definition coeffs :: "'a poly \<Rightarrow> 'a::zero list"
+  where "coeffs p = (if p = 0 then [] else map (\<lambda>i. coeff p i) [0 ..< Suc (degree p)])"
+
+lemma coeffs_eq_Nil [simp]: "coeffs p = [] \<longleftrightarrow> p = 0"
+  by (simp add: coeffs_def)
+
+lemma not_0_coeffs_not_Nil: "p \<noteq> 0 \<Longrightarrow> coeffs p \<noteq> []"
+  by simp
+
+lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []"
+  by simp
+
+lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p"
+proof -
+  have *: "\<forall>m\<in>set ms. m > 0 \<Longrightarrow> map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)"
+    for ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a"
+    by (induct ms) (auto split: nat.split)
+  show ?thesis
+    by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc)
+qed
+
+lemma length_coeffs: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = degree p + 1"
+  by (simp add: coeffs_def)
+
+lemma coeffs_nth: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeffs p ! n = coeff p n"
+  by (auto simp: coeffs_def simp del: upt_Suc)
+
+lemma coeff_in_coeffs: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeff p n \<in> set (coeffs p)"
+  using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs)
+
+lemma not_0_cCons_eq [simp]: "p \<noteq> 0 \<Longrightarrow> a ## coeffs p = a # coeffs p"
+  by (simp add: cCons_def)
+
+lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p"
+  by (induct p) auto
+
+lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as"
+proof (induct as)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons a as)
+  from replicate_length_same [of as 0] have "(\<forall>n. as \<noteq> replicate n 0) \<longleftrightarrow> (\<exists>a\<in>set as. a \<noteq> 0)"
+    by (auto dest: sym [of _ as])
+  with Cons show ?case by auto
+qed
+
+lemma no_trailing_coeffs [simp]:
+  "no_trailing (HOL.eq 0) (coeffs p)"
+  by (induct p)  auto
+
+lemma strip_while_coeffs [simp]:
+  "strip_while (HOL.eq 0) (coeffs p) = coeffs p"
+  by simp
+
+lemma coeffs_eq_iff: "p = q \<longleftrightarrow> coeffs p = coeffs q"
+  (is "?P \<longleftrightarrow> ?Q")
+proof
+  assume ?P
+  then show ?Q by simp
+next
+  assume ?Q
+  then have "Poly (coeffs p) = Poly (coeffs q)" by simp
+  then show ?P by simp
+qed
+
+lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p"
+  by (simp add: fun_eq_iff coeff_Poly_eq [symmetric])
+
+lemma [code]: "coeff p = nth_default 0 (coeffs p)"
+  by (simp add: nth_default_coeffs_eq)
+
+lemma coeffs_eqI:
+  assumes coeff: "\<And>n. coeff p n = nth_default 0 xs n"
+  assumes zero: "no_trailing (HOL.eq 0) xs"
+  shows "coeffs p = xs"
+proof -
+  from coeff have "p = Poly xs"
+    by (simp add: poly_eq_iff)
+  with zero show ?thesis by simp
+qed
+
+lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1"
+  by (simp add: coeffs_def)
+
+lemma length_coeffs_degree: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = Suc (degree p)"
+  by (induct p) (auto simp: cCons_def)
+
+lemma [code abstract]: "coeffs 0 = []"
+  by (fact coeffs_0_eq_Nil)
+
+lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p"
+  by (fact coeffs_pCons_eq_cCons)
+
+instantiation poly :: ("{zero, equal}") equal
+begin
+
+definition [code]: "HOL.equal (p::'a poly) q \<longleftrightarrow> HOL.equal (coeffs p) (coeffs q)"
+
+instance
+  by standard (simp add: equal equal_poly_def coeffs_eq_iff)
+
+end
+
+lemma [code nbe]: "HOL.equal (p :: _ poly) p \<longleftrightarrow> True"
+  by (fact equal_refl)
+
+definition is_zero :: "'a::zero poly \<Rightarrow> bool"
+  where [code]: "is_zero p \<longleftrightarrow> List.null (coeffs p)"
+
+lemma is_zero_null [code_abbrev]: "is_zero p \<longleftrightarrow> p = 0"
+  by (simp add: is_zero_def null_def)
+
+
+subsubsection \<open>Reconstructing the polynomial from the list\<close>
+  \<comment> \<open>contributed by Sebastiaan J.C. Joosten and René Thiemann\<close>
+
+definition poly_of_list :: "'a::comm_monoid_add list \<Rightarrow> 'a poly"
+  where [simp]: "poly_of_list = Poly"
+
+lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as"
+  by simp
+
+
+subsection \<open>Fold combinator for polynomials\<close>
+
+definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
+  where "fold_coeffs f p = foldr f (coeffs p)"
+
+lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id"
+  by (simp add: fold_coeffs_def)
+
+lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p"
+  by (simp add: fold_coeffs_def cCons_def fun_eq_iff)
+
+lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id"
+  by (simp add: fold_coeffs_def)
+
+lemma fold_coeffs_pCons_coeff_not_0_eq [simp]:
+  "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p"
+  by (simp add: fold_coeffs_def)
+
+lemma fold_coeffs_pCons_not_0_0_eq [simp]:
+  "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"
+  where "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" \<comment> \<open>The Horner Schema\<close>
+
+lemma poly_0 [simp]: "poly 0 x = 0"
+  by (simp add: poly_def)
+
+lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x"
+  by (cases "p = 0 \<and> a = 0") (auto simp add: poly_def)
+
+lemma poly_altdef: "poly p x = (\<Sum>i\<le>degree p. coeff p i * x ^ i)"
+  for x :: "'a::{comm_semiring_0,semiring_1}"
+proof (induction p rule: pCons_induct)
+  case 0
+  then show ?case
+    by simp
+next
+  case (pCons a p)
+  show ?case
+  proof (cases "p = 0")
+    case True
+    then show ?thesis by simp
+  next
+    case False
+    let ?p' = "pCons a p"
+    note poly_pCons[of a p x]
+    also note pCons.IH
+    also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) =
+        coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)"
+      by (simp add: field_simps sum_distrib_left coeff_pCons)
+    also note sum_atMost_Suc_shift[symmetric]
+    also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric]
+    finally show ?thesis .
+  qed
+qed
+
+lemma poly_0_coeff_0: "poly p 0 = coeff p 0"
+  by (cases p) (auto simp: poly_altdef)
+
+
+subsection \<open>Monomials\<close>
+
+lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly"
+  is "\<lambda>a m n. if m = n then a else 0"
+  by (simp add: MOST_iff_cofinite)
+
+lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)"
+  by transfer rule
+
+lemma monom_0: "monom a 0 = pCons a 0"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma monom_eq_0 [simp]: "monom 0 n = 0"
+  by (rule poly_eqI) simp
+
+lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0"
+  by (simp add: poly_eq_iff)
+
+lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b"
+  by (simp add: poly_eq_iff)
+
+lemma degree_monom_le: "degree (monom a n) \<le> n"
+  by (rule degree_le, simp)
+
+lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n"
+  apply (rule order_antisym [OF degree_monom_le])
+  apply (rule le_degree)
+  apply simp
+  done
+
+lemma coeffs_monom [code abstract]:
+  "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])"
+  by (induct n) (simp_all add: monom_0 monom_Suc)
+
+lemma fold_coeffs_monom [simp]: "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (monom a n) = f 0 ^^ n \<circ> f a"
+  by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff)
+
+lemma poly_monom: "poly (monom a n) x = a * x ^ n"
+  for a x :: "'a::comm_semiring_1"
+  by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_def)
+
+lemma monom_eq_iff': "monom c n = monom d m \<longleftrightarrow>  c = d \<and> (c = 0 \<or> n = m)"
+  by (auto simp: poly_eq_iff)
+
+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
+begin
+
+lift_definition plus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  is "\<lambda>p q n. coeff p n + coeff q n"
+proof -
+  fix q p :: "'a poly"
+  show "\<forall>\<^sub>\<infinity>n. coeff p n + coeff q n = 0"
+    using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp
+qed
+
+lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n"
+  by (simp add: plus_poly.rep_eq)
+
+instance
+proof
+  fix p q r :: "'a poly"
+  show "(p + q) + r = p + (q + r)"
+    by (simp add: poly_eq_iff add.assoc)
+  show "p + q = q + p"
+    by (simp add: poly_eq_iff add.commute)
+  show "0 + p = p"
+    by (simp add: poly_eq_iff)
+qed
+
+end
+
+instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add
+begin
+
+lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  is "\<lambda>p q n. coeff p n - coeff q n"
+proof -
+  fix q p :: "'a poly"
+  show "\<forall>\<^sub>\<infinity>n. coeff p n - coeff q n = 0"
+    using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp
+qed
+
+lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n"
+  by (simp add: minus_poly.rep_eq)
+
+instance
+proof
+  fix p q r :: "'a poly"
+  show "p + q - p = q"
+    by (simp add: poly_eq_iff)
+  show "p - q - r = p - (q + r)"
+    by (simp add: poly_eq_iff diff_diff_eq)
+qed
+
+end
+
+instantiation poly :: (ab_group_add) ab_group_add
+begin
+
+lift_definition uminus_poly :: "'a poly \<Rightarrow> 'a poly"
+  is "\<lambda>p n. - coeff p n"
+proof -
+  fix p :: "'a poly"
+  show "\<forall>\<^sub>\<infinity>n. - coeff p n = 0"
+    using MOST_coeff_eq_0 by simp
+qed
+
+lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n"
+  by (simp add: uminus_poly.rep_eq)
+
+instance
+proof
+  fix p q :: "'a poly"
+  show "- p + p = 0"
+    by (simp add: poly_eq_iff)
+  show "p - q = p + - q"
+    by (simp add: poly_eq_iff)
+qed
+
+end
+
+lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)"
+  by (rule degree_le) (auto simp add: coeff_eq_0)
+
+lemma degree_add_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p + q) \<le> n"
+  by (auto intro: order_trans degree_add_le_max)
+
+lemma degree_add_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p + q) < n"
+  by (auto intro: le_less_trans degree_add_le_max)
+
+lemma degree_add_eq_right: "degree p < degree q \<Longrightarrow> degree (p + q) = degree q"
+  apply (cases "q = 0")
+   apply simp
+  apply (rule order_antisym)
+   apply (simp add: degree_add_le)
+  apply (rule le_degree)
+  apply (simp add: coeff_eq_0)
+  done
+
+lemma degree_add_eq_left: "degree q < degree p \<Longrightarrow> degree (p + q) = degree p"
+  using degree_add_eq_right [of q p] by (simp add: add.commute)
+
+lemma degree_minus [simp]: "degree (- p) = degree p"
+  by (simp add: degree_def)
+
+lemma lead_coeff_add_le: "degree p < degree q \<Longrightarrow> lead_coeff (p + q) = lead_coeff q"
+  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: "degree (p - q) \<le> max (degree p) (degree q)"
+  for p q :: "'a::ab_group_add poly"
+  using degree_add_le [where p=p and q="-q"] by simp
+
+lemma degree_diff_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p - q) \<le> n"
+  for p q :: "'a::ab_group_add poly"
+  using degree_add_le [of p n "- q"] by simp
+
+lemma degree_diff_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p - q) < n"
+  for p q :: "'a::ab_group_add poly"
+  using degree_add_less [of p n "- q"] by simp
+
+lemma add_monom: "monom a n + monom b n = monom (a + b) n"
+  by (rule poly_eqI) simp
+
+lemma diff_monom: "monom a n - monom b n = monom (a - b) n"
+  by (rule poly_eqI) simp
+
+lemma minus_monom: "- monom a n = monom (- a) n"
+  by (rule poly_eqI) simp
+
+lemma coeff_sum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)"
+  by (induct A rule: infinite_finite_induct) simp_all
+
+lemma monom_sum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)"
+  by (rule poly_eqI) (simp add: coeff_sum)
+
+fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+  where
+    "plus_coeffs xs [] = xs"
+  | "plus_coeffs [] ys = ys"
+  | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys"
+
+lemma coeffs_plus_eq_plus_coeffs [code abstract]:
+  "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)"
+proof -
+  have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n"
+    for xs ys :: "'a list" and n
+  proof (induct xs ys arbitrary: n rule: plus_coeffs.induct)
+    case (3 x xs y ys n)
+    then show ?case
+      by (cases n) (auto simp add: cCons_def)
+  qed simp_all
+  have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)"
+    if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys"
+    for xs ys :: "'a list"
+    using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def)
+  show ?thesis
+    by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **)
+qed
+
+lemma coeffs_uminus [code abstract]:
+  "coeffs (- p) = map uminus (coeffs p)"
+proof -
+  have eq_0: "HOL.eq 0 \<circ> uminus = HOL.eq (0::'a)"
+    by (simp add: fun_eq_iff)
+  show ?thesis
+    by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0)
+qed
+
+lemma [code]: "p - q = p + - q"
+  for p q :: "'a::ab_group_add poly"
+  by (fact diff_conv_add_uminus)
+
+lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x"
+  apply (induct p arbitrary: q)
+   apply simp
+  apply (case_tac q, simp, simp add: algebra_simps)
+  done
+
+lemma poly_minus [simp]: "poly (- p) x = - poly p x"
+  for x :: "'a::comm_ring"
+  by (induct p) simp_all
+
+lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x"
+  for x :: "'a::comm_ring"
+  using poly_add [of p "- q" x] by simp
+
+lemma poly_sum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
+  by (induct A rule: infinite_finite_induct) simp_all
+
+lemma degree_sum_le: "finite S \<Longrightarrow> (\<And>p. p \<in> S \<Longrightarrow> degree (f p) \<le> n) \<Longrightarrow> degree (sum f S) \<le> n"
+proof (induct S rule: finite_induct)
+  case empty
+  then show ?case by simp
+next
+  case (insert p S)
+  then have "degree (sum f S) \<le> n" "degree (f p) \<le> n"
+    by auto
+  then show ?case
+    unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le)
+qed
+
+lemma poly_as_sum_of_monoms':
+  assumes "degree p \<le> n"
+  shows "(\<Sum>i\<le>n. monom (coeff p i) i) = p"
+proof -
+  have eq: "\<And>i. {..n} \<inter> {i} = (if i \<le> n then {i} else {})"
+    by auto
+  from assms show ?thesis
+    by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq
+        if_distrib[where f="\<lambda>x. x * a" for a])
+qed
+
+lemma poly_as_sum_of_monoms: "(\<Sum>i\<le>degree p. monom (coeff p i) i) = p"
+  by (intro poly_as_sum_of_monoms' order_refl)
+
+lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)"
+  by (induct xs) (simp_all add: monom_0 monom_Suc)
+
+
+subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close>
+
+lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  is "\<lambda>a p n. a * coeff p n"
+proof -
+  fix a :: 'a and p :: "'a poly"
+  show "\<forall>\<^sub>\<infinity> i. a * coeff p i = 0"
+    using MOST_coeff_eq_0[of p] by eventually_elim simp
+qed
+
+lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n"
+  by (simp add: smult.rep_eq)
+
+lemma degree_smult_le: "degree (smult a p) \<le> degree p"
+  by (rule degree_le) (simp add: coeff_eq_0)
+
+lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p"
+  by (rule poly_eqI) (simp add: mult.assoc)
+
+lemma smult_0_right [simp]: "smult a 0 = 0"
+  by (rule poly_eqI) simp
+
+lemma smult_0_left [simp]: "smult 0 p = 0"
+  by (rule poly_eqI) simp
+
+lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p"
+  by (rule poly_eqI) simp
+
+lemma smult_add_right: "smult a (p + q) = smult a p + smult a q"
+  by (rule poly_eqI) (simp add: algebra_simps)
+
+lemma smult_add_left: "smult (a + b) p = smult a p + smult b p"
+  by (rule poly_eqI) (simp add: algebra_simps)
+
+lemma smult_minus_right [simp]: "smult a (- p) = - smult a p"
+  for a :: "'a::comm_ring"
+  by (rule poly_eqI) simp
+
+lemma smult_minus_left [simp]: "smult (- a) p = - smult a p"
+  for a :: "'a::comm_ring"
+  by (rule poly_eqI) simp
+
+lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q"
+  for a :: "'a::comm_ring"
+  by (rule poly_eqI) (simp add: algebra_simps)
+
+lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p"
+  for a b :: "'a::comm_ring"
+  by (rule poly_eqI) (simp add: algebra_simps)
+
+lemmas smult_distribs =
+  smult_add_left smult_add_right
+  smult_diff_left smult_diff_right
+
+lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)"
+  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
+
+lemma smult_monom: "smult a (monom b n) = monom (a * b) n"
+  by (induct n) (simp_all add: monom_0 monom_Suc)
+
+lemma smult_Poly: "smult c (Poly xs) = Poly (map (op * c) xs)"
+  by (auto simp: poly_eq_iff nth_default_def)
+
+lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)"
+  for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
+  by (cases "a = 0") (simp_all add: degree_def)
+
+lemma smult_eq_0_iff [simp]: "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0"
+  for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
+  by (simp add: poly_eq_iff)
+
+lemma coeffs_smult [code abstract]:
+  "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))"
+  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+proof -
+  have eq_0: "HOL.eq 0 \<circ> times a = HOL.eq (0::'a)" if "a \<noteq> 0"
+    using that by (simp add: fun_eq_iff)
+  show ?thesis
+    by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0)
+qed  
+
+lemma smult_eq_iff:
+  fixes b :: "'a :: field"
+  assumes "b \<noteq> 0"
+  shows "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?lhs
+  also from assms have "smult (inverse b) \<dots> = q"
+    by simp
+  finally show ?rhs
+    by (simp add: field_simps)
+next
+  assume ?rhs
+  with assms show ?lhs by auto
+qed
+
+instantiation poly :: (comm_semiring_0) comm_semiring_0
+begin
+
+definition "p * q = fold_coeffs (\<lambda>a p. smult a q + pCons 0 p) p 0"
+
+lemma mult_poly_0_left: "(0::'a poly) * q = 0"
+  by (simp add: times_poly_def)
+
+lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)"
+  by (cases "p = 0 \<and> a = 0") (auto simp add: times_poly_def)
+
+lemma mult_poly_0_right: "p * (0::'a poly) = 0"
+  by (induct p) (simp_all add: mult_poly_0_left)
+
+lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)"
+  by (induct p) (simp_all add: mult_poly_0_left algebra_simps)
+
+lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right
+
+lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)"
+  by (induct p) (simp_all add: mult_poly_0 smult_add_right)
+
+lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)"
+  by (induct q) (simp_all add: mult_poly_0 smult_add_right)
+
+lemma mult_poly_add_left: "(p + q) * r = p * r + q * r"
+  for p q r :: "'a poly"
+  by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps)
+
+instance
+proof
+  fix p q r :: "'a poly"
+  show 0: "0 * p = 0"
+    by (rule mult_poly_0_left)
+  show "p * 0 = 0"
+    by (rule mult_poly_0_right)
+  show "(p + q) * r = p * r + q * r"
+    by (rule mult_poly_add_left)
+  show "(p * q) * r = p * (q * r)"
+    by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left)
+  show "p * q = q * p"
+    by (induct p) (simp_all add: mult_poly_0)
+qed
+
+end
+
+lemma coeff_mult_degree_sum:
+  "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)"
+  by (induct p) (simp_all add: coeff_eq_0)
+
+instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors
+proof
+  fix p q :: "'a poly"
+  assume "p \<noteq> 0" and "q \<noteq> 0"
+  have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)"
+    by (rule coeff_mult_degree_sum)
+  also from \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "coeff p (degree p) * coeff q (degree q) \<noteq> 0"
+    by simp
+  finally have "\<exists>n. coeff (p * q) n \<noteq> 0" ..
+  then show "p * q \<noteq> 0"
+    by (simp add: poly_eq_iff)
+qed
+
+instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel ..
+
+lemma coeff_mult: "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))"
+proof (induct p arbitrary: n)
+  case 0
+  show ?case by simp
+next
+  case (pCons a p n)
+  then show ?case
+    by (cases n) (simp_all add: sum_atMost_Suc_shift del: sum_atMost_Suc)
+qed
+
+lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q"
+  apply (rule degree_le)
+  apply (induct p)
+   apply simp
+  apply (simp add: coeff_eq_0 coeff_pCons split: nat.split)
+  done
+
+lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)"
+  by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc)
+
+instantiation poly :: (comm_semiring_1) comm_semiring_1
+begin
+
+definition one_poly_def: "1 = pCons 1 0"
+
+instance
+proof
+  show "1 * p = p" for p :: "'a poly"
+    by (simp add: one_poly_def)
+  show "0 \<noteq> (1::'a poly)"
+    by (simp add: one_poly_def)
+qed
+
+end
+
+instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors ..
+instance poly :: (comm_ring) comm_ring ..
+instance poly :: (comm_ring_1) comm_ring_1 ..
+instance poly :: (comm_ring_1) comm_semiring_1_cancel ..
+
+lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)"
+  by (simp add: one_poly_def coeff_pCons split: nat.split)
+
+lemma monom_eq_1 [simp]: "monom 1 0 = 1"
+  by (simp add: monom_0 one_poly_def)
+
+lemma monom_eq_1_iff: "monom c n = 1 \<longleftrightarrow> c = 1 \<and> n = 0"
+  using monom_eq_const_iff[of c n 1] by (auto simp: one_poly_def)
+
+lemma monom_altdef: "monom c n = smult c ([:0, 1:]^n)"
+  by (induct n) (simp_all add: monom_0 monom_Suc one_poly_def)
+
+lemma degree_1 [simp]: "degree 1 = 0"
+  unfolding one_poly_def by (rule degree_pCons_0)
+
+lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]"
+  by (simp add: one_poly_def)
+
+lemma degree_power_le: "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 (induct n) (simp_all add: coeff_mult)
+
+lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x"
+  by (induct p) (simp_all add: algebra_simps)
+
+lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x"
+  by (induct p) (simp_all add: algebra_simps)
+
+lemma poly_1 [simp]: "poly 1 x = 1"
+  by (simp add: one_poly_def)
+
+lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n"
+  for p :: "'a::comm_semiring_1 poly"
+  by (induct n) simp_all
+
+lemma poly_prod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)"
+  by (induct A rule: infinite_finite_induct) simp_all
+
+lemma degree_prod_sum_le: "finite S \<Longrightarrow> degree (prod f S) \<le> sum (degree o f) S"
+proof (induct S rule: finite_induct)
+  case empty
+  then show ?case by simp
+next
+  case (insert a S)
+  show ?case
+    unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)]
+    by (rule le_trans[OF degree_mult_le]) (use insert in auto)
+qed
+
+lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
+  by (induct 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 "p = monom 1 n * r"
+    by (rule dvdE)
+  then show "\<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
+  then have 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 (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero)
+  then show "monom 1 n dvd p" by simp
+qed
+
+
+subsection \<open>Mapping polynomials\<close>
+
+definition map_poly :: "('a :: zero \<Rightarrow> 'b :: zero) \<Rightarrow> 'a poly \<Rightarrow> 'b poly"
+  where "map_poly f p = Poly (map f (coeffs p))"
+
+lemma map_poly_0 [simp]: "map_poly f 0 = 0"
+  by (simp add: map_poly_def)
+
+lemma map_poly_1: "map_poly f 1 = [:f 1:]"
+  by (simp add: map_poly_def)
+
+lemma map_poly_1' [simp]: "f 1 = 1 \<Longrightarrow> map_poly f 1 = 1"
+  by (simp add: map_poly_def one_poly_def)
+
+lemma coeff_map_poly:
+  assumes "f 0 = 0"
+  shows "coeff (map_poly f p) n = f (coeff p n)"
+  by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0
+      simp del: upt_Suc)
+
+lemma coeffs_map_poly [code abstract]:
+  "coeffs (map_poly f p) = strip_while (op = 0) (map f (coeffs p))"
+  by (simp add: map_poly_def)
+
+lemma coeffs_map_poly':
+  assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
+  shows "coeffs (map_poly f p) = map f (coeffs p)"
+  using assms by (simp add: coeffs_map_poly no_trailing_map strip_while_idem_iff)
+    (metis comp_def no_leading_def no_trailing_coeffs)
+
+lemma set_coeffs_map_poly:
+  "(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)"
+  by (simp add: coeffs_map_poly')
+
+lemma degree_map_poly:
+  assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
+  shows "degree (map_poly f p) = degree p"
+  by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms)
+
+lemma map_poly_eq_0_iff:
+  assumes "f 0 = 0" "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
+  shows "map_poly f p = 0 \<longleftrightarrow> p = 0"
+proof -
+  have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n
+  proof -
+    have "coeff (map_poly f p) n = f (coeff p n)"
+      by (simp add: coeff_map_poly assms)
+    also have "\<dots> = 0 \<longleftrightarrow> coeff p n = 0"
+    proof (cases "n < length (coeffs p)")
+      case True
+      then have "coeff p n \<in> set (coeffs p)"
+        by (auto simp: coeffs_def simp del: upt_Suc)
+      with assms show "f (coeff p n) = 0 \<longleftrightarrow> coeff p n = 0"
+        by auto
+    next
+      case False
+      then show ?thesis
+        by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def)
+    qed
+    finally show ?thesis .
+  qed
+  then show ?thesis by (auto simp: poly_eq_iff)
+qed
+
+lemma map_poly_smult:
+  assumes "f 0 = 0""\<And>c x. f (c * x) = f c * f x"
+  shows "map_poly f (smult c p) = smult (f c) (map_poly f p)"
+  by (intro poly_eqI) (simp_all add: assms coeff_map_poly)
+
+lemma map_poly_pCons:
+  assumes "f 0 = 0"
+  shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)"
+  by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits)
+
+lemma map_poly_map_poly:
+  assumes "f 0 = 0" "g 0 = 0"
+  shows "map_poly f (map_poly g p) = map_poly (f \<circ> g) p"
+  by (intro poly_eqI) (simp add: coeff_map_poly assms)
+
+lemma map_poly_id [simp]: "map_poly id p = p"
+  by (simp add: map_poly_def)
+
+lemma map_poly_id' [simp]: "map_poly (\<lambda>x. x) p = p"
+  by (simp add: map_poly_def)
+
+lemma map_poly_cong:
+  assumes "(\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = g x)"
+  shows "map_poly f p = map_poly g p"
+proof -
+  from assms have "map f (coeffs p) = map g (coeffs p)"
+    by (intro map_cong) simp_all
+  then show ?thesis
+    by (simp only: coeffs_eq_iff coeffs_map_poly)
+qed
+
+lemma map_poly_monom: "f 0 = 0 \<Longrightarrow> map_poly f (monom c n) = monom (f c) n"
+  by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+lemma map_poly_idI:
+  assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
+  shows "map_poly f p = p"
+  using map_poly_cong[OF assms, of _ id] by simp
+
+lemma map_poly_idI':
+  assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
+  shows "p = map_poly f p"
+  using map_poly_cong[OF assms, of _ id] by simp
+
+lemma smult_conv_map_poly: "smult c p = map_poly (\<lambda>x. c * x) p"
+  by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+
+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 (induct n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  then have "of_nat (Suc n) = 1 + (of_nat n :: 'a poly)"
+    by simp
+  also have "(of_nat n :: 'a poly) = [: of_nat n :]"
+    by (subst Suc) (rule refl)
+  also have "1 = [:1:]"
+    by (simp add: one_poly_def)
+  finally show ?case
+    by (subst (asm) add_pCons) simp
+qed
+
+lemma degree_of_nat [simp]: "degree (of_nat n) = 0"
+  by (simp add: of_nat_poly)
+
+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>
+
+lemma dvd_smult:
+  assumes "p dvd q"
+  shows "p dvd smult a q"
+proof -
+  from assms obtain k where "q = p * k" ..
+  then have "smult a q = p * smult a k" by simp
+  then show "p dvd smult a q" ..
+qed
+
+lemma dvd_smult_cancel: "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q"
+  for a :: "'a::field"
+  by (drule dvd_smult [where a="inverse a"]) simp
+
+lemma dvd_smult_iff: "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q"
+  for a :: "'a::field"
+  by (safe elim!: dvd_smult dvd_smult_cancel)
+
+lemma smult_dvd_cancel:
+  assumes "smult a p dvd q"
+  shows "p dvd q"
+proof -
+  from assms obtain k where "q = smult a p * k" ..
+  then have "q = p * smult a k" by simp
+  then show "p dvd q" ..
+qed
+
+lemma smult_dvd: "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q"
+  for a :: "'a::field"
+  by (rule smult_dvd_cancel [where a="inverse a"]) simp
+
+lemma smult_dvd_iff: "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)"
+  for a :: "'a::field"
+  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 *: "[:c:] * p dvd 1"
+    then show "p dvd 1"
+      by (rule dvd_mult_right)
+    from * obtain q where q: "1 = [:c:] * p * q"
+      by (rule 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 q [symmetric]
+    finally have "c dvd coeff 1 0" .
+    then show "c dvd 1" by simp
+  next
+    assume "c dvd 1" "p dvd 1"
+    from this(1) obtain d where "1 = c * d"
+      by (rule dvdE)
+    then have "1 = [:c:] * [:d:]"
+      by (simp add: one_poly_def mult_ac)
+    then have "[: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>
+
+instance poly :: (idom) idom ..
+
+lemma degree_mult_eq: "p \<noteq> 0 \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree (p * q) = degree p + degree q"
+  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum)
+
+lemma degree_mult_eq_0:
+  "degree (p * q) = 0 \<longleftrightarrow> p = 0 \<or> q = 0 \<or> (p \<noteq> 0 \<and> q \<noteq> 0 \<and> degree p = 0 \<and> degree q = 0)"
+  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  by (auto simp: degree_mult_eq)
+
+lemma degree_mult_right_le:
+  fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  assumes "q \<noteq> 0"
+  shows "degree p \<le> degree (p * q)"
+  using assms by (cases "p = 0") (simp_all add: degree_mult_eq)
+
+lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)"
+  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  by (cases "p = 0 \<or> q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac)
+
+lemma dvd_imp_degree_le: "p dvd q \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree p \<le> degree q"
+  for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  by (erule dvdE, hypsubst, subst degree_mult_eq) auto
+
+lemma divides_degree:
+  fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  assumes "p dvd q"
+  shows "degree p \<le> degree q \<or> q = 0"
+  by (metis dvd_imp_degree_le assms)
+
+lemma const_poly_dvd_iff:
+  fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
+  shows "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)"
+proof (cases "c = 0 \<or> p = 0")
+  case True
+  then show ?thesis
+    by (auto intro!: poly_eqI)
+next
+  case False
+  show ?thesis
+  proof
+    assume "[:c:] dvd p"
+    then show "\<forall>n. c dvd coeff p n"
+      by (auto elim!: dvdE simp: coeffs_def)
+  next
+    assume *: "\<forall>n. c dvd coeff p n"
+    define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a
+    have mydiv: "x = y * mydiv x y" if "y dvd x" for x y
+      using that unfolding mydiv_def dvd_def by (rule someI_ex)
+    define q where "q = Poly (map (\<lambda>a. mydiv a c) (coeffs p))"
+    from False * have "p = q * [:c:]"
+      by (intro poly_eqI)
+        (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth
+          intro!: coeff_eq_0 mydiv)
+    then show "[:c:] dvd p"
+      by (simp only: dvd_triv_right)
+  qed
+qed
+
+lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \<longleftrightarrow> a dvd b"
+  for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
+  by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits)
+
+lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
+  for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly"
+  by (cases "p = 0 \<or> q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq)
+
+lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p"
+  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+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) = lead_coeff p ^ n"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  by (induct n) (simp_all add: lead_coeff_mult)
+
+
+subsection \<open>Polynomials form an ordered integral domain\<close>
+
+definition pos_poly :: "'a::linordered_semidom poly \<Rightarrow> bool"
+  where "pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)"
+
+lemma pos_poly_pCons: "pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)"
+  by (simp add: pos_poly_def)
+
+lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0"
+  by (simp add: pos_poly_def)
+
+lemma pos_poly_add: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p + q)"
+  apply (induct p arbitrary: q)
+   apply simp
+  apply (case_tac q)
+  apply (force simp add: pos_poly_pCons add_pos_pos)
+  done
+
+lemma pos_poly_mult: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p * q)"
+  unfolding pos_poly_def
+  apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0")
+   apply (simp add: degree_mult_eq coeff_mult_degree_sum)
+  apply auto
+  done
+
+lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)"
+  for p :: "'a::linordered_idom poly"
+  by (induct p) (auto simp: pos_poly_pCons)
+
+lemma last_coeffs_eq_coeff_degree: "p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)"
+  by (simp add: coeffs_def)
+
+lemma pos_poly_coeffs [code]: "pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs
+  then show ?lhs
+    by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree)
+next
+  assume ?lhs
+  then have *: "0 < coeff p (degree p)"
+    by (simp add: pos_poly_def)
+  then have "p \<noteq> 0"
+    by auto
+  with * show ?rhs
+    by (simp add: last_coeffs_eq_coeff_degree)
+qed
+
+instantiation poly :: (linordered_idom) linordered_idom
+begin
+
+definition "x < y \<longleftrightarrow> pos_poly (y - x)"
+
+definition "x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)"
+
+definition "\<bar>x::'a poly\<bar> = (if x < 0 then - x else x)"
+
+definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
+
+instance
+proof
+  fix x y z :: "'a poly"
+  show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
+    unfolding less_eq_poly_def less_poly_def
+    apply safe
+     apply simp
+    apply (drule (1) pos_poly_add)
+    apply simp
+    done
+  show "x \<le> x"
+    by (simp add: less_eq_poly_def)
+  show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
+    unfolding less_eq_poly_def
+    apply safe
+    apply (drule (1) pos_poly_add)
+    apply (simp add: algebra_simps)
+    done
+  show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
+    unfolding less_eq_poly_def
+    apply safe
+    apply (drule (1) pos_poly_add)
+    apply simp
+    done
+  show "x \<le> y \<Longrightarrow> z + x \<le> z + y"
+    unfolding less_eq_poly_def
+    apply safe
+    apply (simp add: algebra_simps)
+    done
+  show "x \<le> y \<or> y \<le> x"
+    unfolding less_eq_poly_def
+    using pos_poly_total [of "x - y"]
+    by auto
+  show "x < y \<Longrightarrow> 0 < z \<Longrightarrow> z * x < z * y"
+    by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult)
+  show "\<bar>x\<bar> = (if x < 0 then - x else x)"
+    by (rule abs_poly_def)
+  show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
+    by (rule sgn_poly_def)
+qed
+
+end
+
+text \<open>TODO: Simplification rules for comparisons\<close>
+
+
+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>
+
+definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a"
+  where "synthetic_divmod p c = fold_coeffs (\<lambda>a (q, r). (pCons r q, a + c * r)) p (0, 0)"
+
+definition synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
+  where "synthetic_div p c = fst (synthetic_divmod p c)"
+
+lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)"
+  by (simp add: synthetic_divmod_def)
+
+lemma synthetic_divmod_pCons [simp]:
+  "synthetic_divmod (pCons a p) c = (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)"
+  by (cases "p = 0 \<and> a = 0") (auto simp add: synthetic_divmod_def)
+
+lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0"
+  by (simp add: synthetic_div_def)
+
+lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0"
+  by (induct p arbitrary: a) simp_all
+
+lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c"
+  by (induct p) (simp_all add: split_def)
+
+lemma synthetic_div_pCons [simp]:
+  "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)"
+  by (simp add: synthetic_div_def split_def snd_synthetic_divmod)
+
+lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0"
+proof (induct p)
+  case 0
+  then show ?case by simp
+next
+  case (pCons a p)
+  then show ?case by (cases p) simp
+qed
+
+lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1"
+  by (induct p) (simp_all add: synthetic_div_eq_0_iff)
+
+lemma synthetic_div_correct:
+  "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)"
+  by (induct p) simp_all
+
+lemma synthetic_div_unique: "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c"
+  apply (induct p arbitrary: q r)
+   apply simp
+   apply (frule synthetic_div_unique_lemma)
+   apply simp
+  apply (case_tac q, force)
+  done
+
+lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p"
+  for c :: "'a::comm_ring_1"
+  using synthetic_div_correct [of p c] by (simp add: algebra_simps)
+
+
+subsubsection \<open>Polynomial roots\<close>
+
+lemma poly_eq_0_iff_dvd: "poly p c = 0 \<longleftrightarrow> [:- c, 1:] dvd p"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+  for c :: "'a::comm_ring_1"
+proof
+  assume ?lhs
+  with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp
+  then show ?rhs ..
+next
+  assume ?rhs
+  then obtain k where "p = [:-c, 1:] * k" by (rule dvdE)
+  then show ?lhs by simp
+qed
+
+lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \<longleftrightarrow> poly p (- c) = 0"
+  for c :: "'a::comm_ring_1"
+  by (simp add: poly_eq_0_iff_dvd)
+
+lemma poly_roots_finite: "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
+  for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly"
+proof (induct n \<equiv> "degree p" arbitrary: p)
+  case 0
+  then obtain a where "a \<noteq> 0" and "p = [:a:]"
+    by (cases p) (simp split: if_splits)
+  then show "finite {x. poly p x = 0}"
+    by simp
+next
+  case (Suc n)
+  show "finite {x. poly p x = 0}"
+  proof (cases "\<exists>x. poly p x = 0")
+    case False
+    then show "finite {x. poly p x = 0}" by simp
+  next
+    case True
+    then obtain a where "poly p a = 0" ..
+    then have "[:-a, 1:] dvd p"
+      by (simp only: poly_eq_0_iff_dvd)
+    then obtain k where k: "p = [:-a, 1:] * k" ..
+    with \<open>p \<noteq> 0\<close> have "k \<noteq> 0"
+      by auto
+    with k have "degree p = Suc (degree k)"
+      by (simp add: degree_mult_eq del: mult_pCons_left)
+    with \<open>Suc n = degree p\<close> have "n = degree k"
+      by simp
+    from this \<open>k \<noteq> 0\<close> have "finite {x. poly k x = 0}"
+      by (rule Suc.hyps)
+    then have "finite (insert a {x. poly k x = 0})"
+      by simp
+    then show "finite {x. poly p x = 0}"
+      by (simp add: k Collect_disj_eq del: mult_pCons_left)
+  qed
+qed
+
+lemma poly_eq_poly_eq_iff: "poly p = poly q \<longleftrightarrow> p = q"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+  for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly"
+proof
+  assume ?rhs
+  then show ?lhs by simp
+next
+  assume ?lhs
+  have "poly p = poly 0 \<longleftrightarrow> p = 0" for p :: "'a poly"
+    apply (cases "p = 0")
+     apply simp_all
+    apply (drule poly_roots_finite)
+    apply (auto simp add: infinite_UNIV_char_0)
+    done
+  from \<open>?lhs\<close> and this [of "p - q"] show ?rhs
+    by auto
+qed
+
+lemma poly_all_0_iff_0: "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0"
+  for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly"
+  by (auto simp add: poly_eq_poly_eq_iff [symmetric])
+
+
+subsubsection \<open>Order of polynomial roots\<close>
+
+definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat"
+  where "order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)"
+
+lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1"
+  for a :: "'a::comm_semiring_1"
+  apply (induct n)
+   apply simp_all
+  apply (subst coeff_eq_0)
+   apply (auto intro: le_less_trans degree_power_le)
+  done
+
+lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n"
+  for a :: "'a::comm_semiring_1"
+  apply (rule order_antisym)
+   apply (rule ord_le_eq_trans [OF degree_power_le])
+   apply simp
+  apply (rule le_degree)
+  apply (simp add: coeff_linear_power)
+  done
+
+lemma order_1: "[:-a, 1:] ^ order a p dvd p"
+  apply (cases "p = 0")
+   apply simp
+  apply (cases "order a p")
+   apply simp
+  apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)")
+   apply (drule not_less_Least)
+   apply simp
+  apply (fold order_def)
+  apply simp
+  done
+
+lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
+  unfolding order_def
+  apply (rule LeastI_ex)
+  apply (rule_tac x="degree p" in exI)
+  apply (rule notI)
+  apply (drule (1) dvd_imp_degree_le)
+  apply (simp only: degree_linear_power)
+  done
+
+lemma order: "p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
+  by (rule conjI [OF order_1 order_2])
+
+lemma order_degree:
+  assumes p: "p \<noteq> 0"
+  shows "order a p \<le> degree p"
+proof -
+  have "order a p = degree ([:-a, 1:] ^ order a p)"
+    by (simp only: degree_linear_power)
+  also from order_1 p have "\<dots> \<le> degree p"
+    by (rule dvd_imp_degree_le)
+  finally show ?thesis .
+qed
+
+lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0"
+  apply (cases "p = 0")
+   apply simp_all
+  apply (rule iffI)
+   apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right)
+  unfolding poly_eq_0_iff_dvd
+  apply (metis dvd_power dvd_trans order_1)
+  done
+
+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"
+    by (simp add: t_def 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 True
+  then show ?thesis
+    by simp
+next
+  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 have "order x [:c:] = 0"
+    by (rule order_0I) (use assms in auto)
+  finally show ?thesis
+    by simp
+qed
+
+(* 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
+  then show ?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 notI)
+    assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
+    then have "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:])"
+      by (rule dvd_imp_degree_le) auto
+    then show False
+      by auto
+  next
+    fix y
+    assume *: "\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
+    show "1 \<le> y"
+    proof (rule ccontr)
+      assume "\<not> 1 \<le> y"
+      then have "y = 0" by auto
+      then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto
+      with * show False by auto
+    qed
+  qed
+  ultimately show ?case
+    using Suc by auto
+qed
+
+lemma order_0_monom [simp]: "c \<noteq> 0 \<Longrightarrow> order 0 (monom c n) = n"
+  using 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. \<open>f / gcd f f'\<close>.\<close>
+
+lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"
+  apply (cases "p = 0")
+  apply 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, 1:] ^ order a p dvd p"
+    and **: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p"
+    by (auto dest: order)
+  from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" ..
+  with ** 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
+  with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q]
+  have "\<not> [:- a, 1:] dvd q" by auto
+  with q show ?thesis by blast
+qed
+
+lemma monom_1_dvd_iff: "p \<noteq> 0 \<Longrightarrow> monom 1 n dvd p \<longleftrightarrow> n \<le> order 0 p"
+  using order_divides[of 0 n p] by (simp add: monom_altdef)
+
+
+subsection \<open>Additional induction rules on polynomials\<close>
+
+text \<open>
+  An induction rule for induction over the roots of a polynomial with a certain property.
+  (e.g. all positive roots)
+\<close>
+lemma poly_root_induct [case_names 0 no_roots root]:
+  fixes p :: "'a :: idom poly"
+  assumes "Q 0"
+    and "\<And>p. (\<And>a. P a \<Longrightarrow> poly p a \<noteq> 0) \<Longrightarrow> Q p"
+    and "\<And>a p. P a \<Longrightarrow> Q p \<Longrightarrow> Q ([:a, -1:] * p)"
+  shows "Q p"
+proof (induction "degree p" arbitrary: p rule: less_induct)
+  case (less p)
+  show ?case
+  proof (cases "p = 0")
+    case True
+    with assms(1) show ?thesis by simp
+  next
+    case False
+    show ?thesis
+    proof (cases "\<exists>a. P a \<and> poly p a = 0")
+      case False
+      then show ?thesis by (intro assms(2)) blast
+    next
+      case True
+      then obtain a where a: "P a" "poly p a = 0"
+        by blast
+      then have "-[:-a, 1:] dvd p"
+        by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd)
+      then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp
+      with False have "q \<noteq> 0" by auto
+      have "degree p = Suc (degree q)"
+        by (subst q, subst degree_mult_eq) (simp_all add: \<open>q \<noteq> 0\<close>)
+      then have "Q q" by (intro less) simp
+      with a(1) have "Q ([:a, -1:] * q)"
+        by (rule assms(3))
+      with q show ?thesis by simp
+    qed
+  qed
+qed
+
+lemma dropWhile_replicate_append:
+  "dropWhile (op = a) (replicate n a @ ys) = dropWhile (op = a) ys"
+  by (induct n) simp_all
+
+lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs"
+  by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append)
+
+text \<open>
+  An induction rule for simultaneous induction over two polynomials,
+  prepending one coefficient in each step.
+\<close>
+lemma poly_induct2 [case_names 0 pCons]:
+  assumes "P 0 0" "\<And>a p b q. P p q \<Longrightarrow> P (pCons a p) (pCons b q)"
+  shows "P p q"
+proof -
+  define n where "n = max (length (coeffs p)) (length (coeffs q))"
+  define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)"
+  define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)"
+  have "length xs = length ys"
+    by (simp add: xs_def ys_def n_def)
+  then have "P (Poly xs) (Poly ys)"
+    by (induct rule: list_induct2) (simp_all add: assms)
+  also have "Poly xs = p"
+    by (simp add: xs_def Poly_append_replicate_0)
+  also have "Poly ys = q"
+    by (simp add: ys_def Poly_append_replicate_0)
+  finally show ?thesis .
+qed
+
+
+subsection \<open>Composition of polynomials\<close>
+
+(* Several lemmas contributed by René Thiemann and Akihisa Yamada *)
+
+definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  where "pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0"
+
+notation pcompose (infixl "\<circ>\<^sub>p" 71)
+
+lemma pcompose_0 [simp]: "pcompose 0 q = 0"
+  by (simp add: pcompose_def)
+
+lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q"
+  by (cases "p = 0 \<and> a = 0") (auto simp add: pcompose_def)
+
+lemma pcompose_1: "pcompose 1 p = 1"
+  for p :: "'a::comm_semiring_1 poly"
+  by (auto simp: one_poly_def pcompose_pCons)
+
+lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)"
+  by (induct p) (simp_all add: pcompose_pCons)
+
+lemma degree_pcompose_le: "degree (pcompose p q) \<le> degree p * degree q"
+  apply (induct p)
+   apply simp
+  apply (simp add: pcompose_pCons)
+  apply clarify
+  apply (rule degree_add_le)
+   apply simp
+  apply (rule order_trans [OF degree_mult_le])
+  apply simp
+  done
+
+lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r"
+  for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly"
+proof (induction p q rule: poly_induct2)
+  case 0
+  then show ?case by simp
+next
+  case (pCons a p b q)
+  have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r"
+    by (simp_all add: pcompose_pCons pCons.IH algebra_simps)
+  also have "[:a + b:] = [:a:] + [:b:]" by simp
+  also have "\<dots> + r * pcompose p r + r * pcompose q r =
+    pcompose (pCons a p) r + pcompose (pCons b q) r"
+    by (simp only: pcompose_pCons add_ac)
+  finally show ?case .
+qed
+
+lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r"
+  for p r :: "'a::comm_ring poly"
+  by (induct p) (simp_all add: pcompose_pCons)
+
+lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r"
+  for p q r :: "'a::comm_ring poly"
+  using pcompose_add[of p "-q"] by (simp add: pcompose_uminus)
+
+lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)"
+  for p r :: "'a::comm_semiring_0 poly"
+  by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right)
+
+lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r"
+  for p q r :: "'a::comm_semiring_0 poly"
+  by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps)
+
+lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r"
+  for p q r :: "'a::comm_semiring_0 poly"
+  by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult)
+
+lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p"
+  for p :: "'a::comm_semiring_1 poly"
+  by (induct p) (simp_all add: pcompose_pCons)
+
+lemma pcompose_sum: "pcompose (sum f A) p = sum (\<lambda>i. pcompose (f i) p) A"
+  by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add)
+
+lemma pcompose_prod: "pcompose (prod f A) p = prod (\<lambda>i. pcompose (f i) p) A"
+  by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult)
+
+lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]"
+  by (subst pcompose_pCons) simp
+
+lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]"
+  by (induct p) (auto simp add: pcompose_pCons)
+
+lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q"
+  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+proof (induct p)
+  case 0
+  then show ?case by auto
+next
+  case (pCons a p)
+  consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0"
+    by blast
+  then show ?case
+  proof cases
+    case prems: 1
+    show ?thesis
+    proof (cases "p = 0")
+      case True
+      then show ?thesis by auto
+    next
+      case False
+      from prems have "degree q = 0 \<or> pcompose p q = 0"
+        by (auto simp add: degree_mult_eq_0)
+      moreover have False if "pcompose p q = 0" "degree q \<noteq> 0"
+      proof -
+        from pCons.hyps(2) that have "degree p = 0"
+          by auto
+        then obtain a1 where "p = [:a1:]"
+          by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases)
+        with \<open>pcompose p q = 0\<close> \<open>p \<noteq> 0\<close> show False
+          by auto
+      qed
+      ultimately have "degree (pCons a p) * degree q = 0"
+        by auto
+      moreover have "degree (pcompose (pCons a p) q) = 0"
+      proof -
+        from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))"
+          by simp
+        also have "\<dots> \<ge> degree ([:a:] + q * pcompose p q)"
+          by (rule degree_add_le_max)
+        finally show ?thesis
+          by (auto simp add: pcompose_pCons)
+      qed
+      ultimately show ?thesis by simp
+    qed
+  next
+    case prems: 2
+    then have "p \<noteq> 0" "q \<noteq> 0" "pcompose p q \<noteq> 0"
+      by auto
+    from prems degree_add_eq_right [of "[:a:]"]
+    have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)"
+      by (auto simp: pcompose_pCons)
+    with pCons.hyps(2) degree_mult_eq[OF \<open>q\<noteq>0\<close> \<open>pcompose p q\<noteq>0\<close>] show ?thesis
+      by auto
+  qed
+qed
+
+lemma pcompose_eq_0:
+  fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  assumes "pcompose p q = 0" "degree q > 0"
+  shows "p = 0"
+proof -
+  from assms degree_pcompose [of p q] have "degree p = 0"
+    by auto
+  then obtain a where "p = [:a:]"
+    by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases)
+  with assms(1) have "a = 0"
+    by auto
+  with \<open>p = [:a:]\<close> show ?thesis
+    by simp
+qed
+
+lemma lead_coeff_comp:
+  fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  assumes "degree q > 0"
+  shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)"
+proof (induct p)
+  case 0
+  then show ?case by auto
+next
+  case (pCons a p)
+  consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0"
+    by blast
+  then show ?case
+  proof cases
+    case prems: 1
+    then have "pcompose p q = 0"
+      by (metis assms degree_0 degree_mult_eq_0 neq0_conv)
+    with pcompose_eq_0[OF _ \<open>degree q > 0\<close>] have "p = 0"
+      by simp
+    then show ?thesis
+      by auto
+  next
+    case prems: 2
+    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 "\<dots> = 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 "\<dots> = lead_coeff p * lead_coeff q ^ (degree p + 1)"
+      by (auto simp: mult_ac)
+    finally show ?thesis by auto
+  qed
+qed
+
+
+subsection \<open>Shifting polynomials\<close>
+
+definition poly_shift :: "nat \<Rightarrow> 'a::zero poly \<Rightarrow> 'a poly"
+  where "poly_shift n p = Abs_poly (\<lambda>i. coeff p (i + n))"
+
+lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)"
+  by (auto simp add: nth_default_def add_ac)
+
+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)
+  then have "\<forall>k>m. coeff p (k + n) = 0"
+    by auto
+  then have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0"
+    by (auto simp: MOST_nat)
+  then show ?thesis
+    by (simp add: poly_shift_def poly.Abs_poly_inverse)
+qed
+
+lemma poly_shift_id [simp]: "poly_shift 0 = (\<lambda>x. x)"
+  by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift)
+
+lemma poly_shift_0 [simp]: "poly_shift n 0 = 0"
+  by (simp add: poly_eq_iff coeff_poly_shift)
+
+lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)"
+  by (simp add: poly_eq_iff coeff_poly_shift)
+
+lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \<ge> n then monom c (m - n) else 0)"
+  by (auto simp add: poly_eq_iff coeff_poly_shift)
+
+lemma coeffs_shift_poly [code abstract]:
+  "coeffs (poly_shift n p) = drop n (coeffs p)"
+proof (cases "p = 0")
+  case True
+  then show ?thesis by simp
+next
+  case False
+  then show ?thesis
+    by (intro coeffs_eqI)
+      (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq)
+qed
+
+
+subsection \<open>Truncating polynomials\<close>
+
+definition poly_cutoff
+  where "poly_cutoff n p = Abs_poly (\<lambda>k. if k < n then coeff p k else 0)"
+
+lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)"
+  unfolding poly_cutoff_def
+  by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n])
+
+lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0"
+  by (simp add: poly_eq_iff coeff_poly_cutoff)
+
+lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)"
+  by (simp add: poly_eq_iff coeff_poly_cutoff)
+
+lemma coeffs_poly_cutoff [code abstract]:
+  "coeffs (poly_cutoff n p) = strip_while (op = 0) (take n (coeffs p))"
+proof (cases "strip_while (op = 0) (take n (coeffs p)) = []")
+  case True
+  then have "coeff (poly_cutoff n p) k = 0" for k
+    unfolding coeff_poly_cutoff
+    by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth)
+  then have "poly_cutoff n p = 0"
+    by (simp add: poly_eq_iff)
+  then show ?thesis
+    by (subst True) simp_all
+next
+  case False
+  have "no_trailing (op = 0) (strip_while (op = 0) (take n (coeffs p)))"
+    by simp
+  with False have "last (strip_while (op = 0) (take n (coeffs p))) \<noteq> 0"
+    unfolding no_trailing_unfold by auto
+  then show ?thesis
+    by (intro coeffs_eqI)
+      (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq)
+qed
+
+
+subsection \<open>Reflecting polynomials\<close>
+
+definition reflect_poly :: "'a::zero poly \<Rightarrow> 'a poly"
+  where "reflect_poly p = Poly (rev (coeffs p))"
+
+lemma coeffs_reflect_poly [code abstract]:
+  "coeffs (reflect_poly p) = rev (dropWhile (op = 0) (coeffs p))"
+  by (simp add: reflect_poly_def)
+
+lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0"
+  by (simp add: reflect_poly_def)
+
+lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1"
+  by (simp add: reflect_poly_def one_poly_def)
+
+lemma coeff_reflect_poly:
+  "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))"
+  by (cases "p = 0")
+    (auto simp add: reflect_poly_def nth_default_def
+      rev_nth degree_eq_length_coeffs coeffs_nth not_less
+      dest: le_imp_less_Suc)
+
+lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0"
+  by (simp add: coeff_reflect_poly)
+
+lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0"
+  by (simp add: coeff_reflect_poly poly_0_coeff_0)
+
+lemma reflect_poly_pCons':
+  "p \<noteq> 0 \<Longrightarrow> reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))"
+  by (intro poly_eqI)
+    (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split)
+
+lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]"
+  by (cases "a = 0") (simp_all add: reflect_poly_def)
+
+lemma poly_reflect_poly_nz:
+  "x \<noteq> 0 \<Longrightarrow> poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)"
+  for x :: "'a::field"
+  by (induct 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)
+
+lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p"
+  by (simp add: poly_0_coeff_0)
+
+lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> reflect_poly (reflect_poly p) = p"
+  by (cases p rule: pCons_cases) (simp add: reflect_poly_def )
+
+lemma degree_reflect_poly_le: "degree (reflect_poly p) \<le> degree p"
+  by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono)
+
+lemma reflect_poly_pCons: "a \<noteq> 0 \<Longrightarrow> reflect_poly (pCons a p) = Poly (rev (a # coeffs p))"
+  by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly)
+
+lemma degree_reflect_poly_eq [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> degree (reflect_poly p) = degree p"
+  by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs)
+
+(* TODO: does this work with zero divisors as well? Probably not. *)
+lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q"
+  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+proof (cases "p = 0 \<or> q = 0")
+  case False
+  then have [simp]: "p \<noteq> 0" "q \<noteq> 0" by auto
+  show ?thesis
+  proof (rule poly_eqI)
+    show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i
+    proof (cases "i \<le> degree (p * q)")
+      case True
+      define A where "A = {..i} \<inter> {i - degree q..degree p}"
+      define B where "B = {..degree p} \<inter> {degree p - i..degree (p*q) - i}"
+      let ?f = "\<lambda>j. degree p - j"
+
+      from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)"
+        by (simp add: coeff_reflect_poly)
+      also have "\<dots> = (\<Sum>j\<le>degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))"
+        by (simp add: coeff_mult)
+      also have "\<dots> = (\<Sum>j\<in>B. coeff p j * coeff q (degree (p * q) - i - j))"
+        by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0)
+      also from True have "\<dots> = (\<Sum>j\<in>A. coeff p (degree p - j) * coeff q (degree q - (i - j)))"
+        by (intro sum.reindex_bij_witness[of _ ?f ?f])
+          (auto simp: A_def B_def degree_mult_eq add_ac)
+      also have "\<dots> =
+        (\<Sum>j\<le>i.
+          if j \<in> {i - degree q..degree p}
+          then coeff p (degree p - j) * coeff q (degree q - (i - j))
+          else 0)"
+        by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def)
+      also have "\<dots> = coeff (reflect_poly p * reflect_poly q) i"
+        by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong)
+      finally show ?thesis .
+    qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral)
+  qed
+qed auto
+
+lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)"
+  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  using reflect_poly_mult[of "[:c:]" p] by simp
+
+lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  by (induct n) (simp_all add: reflect_poly_mult)
+
+lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\<lambda>x. reflect_poly (f x)) A"
+  for f :: "_ \<Rightarrow> _::{comm_semiring_0,semiring_no_zero_divisors} poly"
+  by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult)
+
+lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)"
+  for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list"
+  by (induct xs) (simp_all add: reflect_poly_mult)
+
+lemma reflect_poly_Poly_nz:
+  "no_trailing (HOL.eq 0) xs \<Longrightarrow> reflect_poly (Poly xs) = Poly (rev xs)"
+  by (simp add: reflect_poly_def)
+
+lemmas reflect_poly_simps =
+  reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult
+  reflect_poly_power reflect_poly_prod reflect_poly_prod_list
+
+
+subsection \<open>Derivatives\<close>
+
+function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \<Rightarrow> 'a poly"
+  where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))"
+  by (auto intro: pCons_cases)
+
+termination pderiv
+  by (relation "measure degree") simp_all
+
+declare pderiv.simps[simp del]
+
+lemma pderiv_0 [simp]: "pderiv 0 = 0"
+  using pderiv.simps [of 0 0] by simp
+
+lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)"
+  by (simp add: pderiv.simps)
+
+lemma pderiv_1 [simp]: "pderiv 1 = 0"
+  by (simp add: one_poly_def pderiv_pCons)
+
+lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0"
+  and pderiv_numeral [simp]: "pderiv (numeral m) = 0"
+  by (simp_all add: of_nat_poly numeral_poly pderiv_pCons)
+
+lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)"
+  by (induct p arbitrary: n)
+    (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split)
+
+fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \<Rightarrow> 'a list \<Rightarrow> 'a list"
+  where
+    "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)"
+  | "pderiv_coeffs_code f [] = []"
+
+definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \<Rightarrow> 'a list"
+  where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)"
+
+(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *)
+lemma pderiv_coeffs_code:
+  "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n"
+proof (induct xs arbitrary: f n)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons x xs)
+  show ?case
+  proof (cases n)
+    case 0
+    then show ?thesis
+      by (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0") (auto simp: cCons_def)
+  next
+    case n: (Suc m)
+    show ?thesis
+    proof (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0")
+      case False
+      then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n =
+          nth_default 0 (pderiv_coeffs_code (f + 1) xs) m"
+        by (auto simp: cCons_def n)
+      also have "\<dots> = (f + of_nat n) * nth_default 0 xs m"
+        by (simp add: Cons n add_ac)
+      finally show ?thesis
+        by (simp add: n)
+    next
+      case True
+      have empty: "pderiv_coeffs_code g xs = [] \<Longrightarrow> g + of_nat m = 0 \<or> nth_default 0 xs m = 0" for g
+      proof (induct xs arbitrary: g m)
+        case Nil
+        then show ?case by simp
+      next
+        case (Cons x xs)
+        from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \<or> x = 0"
+          by (auto simp: cCons_def split: if_splits)
+        note IH = Cons(1)[OF empty]
+        from IH[of m] IH[of "m - 1"] g show ?case
+          by (cases m) (auto simp: field_simps)
+      qed
+      from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0"
+        by (auto simp: cCons_def n)
+      moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0"
+        by (simp add: n) (use empty[of "f+1"] in \<open>auto simp: field_simps\<close>)
+      ultimately show ?thesis by simp
+    qed
+  qed
+qed
+
+lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda>i. f (Suc i)) [0 ..< n]"
+  by (induct n arbitrary: f) auto
+
+lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)"
+  unfolding pderiv_coeffs_def
+proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases)
+  case (1 n)
+  have id: "coeff p (Suc n) = nth_default 0 (map (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n"
+    by (cases "n < degree p") (auto simp: nth_default_def coeff_eq_0)
+  show ?case
+    unfolding coeffs_def map_upt_Suc by (auto simp: id)
+next
+  case 2
+  obtain n :: 'a and xs where defs: "tl (coeffs p) = xs" "1 = n"
+    by simp
+  from 2 show ?case
+    unfolding defs by (induct xs arbitrary: n) (auto simp: cCons_def)
+qed
+
+lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
+  apply (rule iffI)
+   apply (cases p)
+   apply simp
+   apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc)
+  apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0)
+  done
+
+lemma degree_pderiv: "degree (pderiv p) = degree p - 1"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
+  apply (rule order_antisym [OF degree_le])
+   apply (simp add: coeff_pderiv coeff_eq_0)
+  apply (cases "degree p")
+   apply simp
+  apply (rule le_degree)
+  apply (simp add: coeff_pderiv del: of_nat_Suc)
+  apply (metis degree_0 leading_coeff_0_iff nat.distinct(1))
+  done
+
+lemma not_dvd_pderiv:
+  fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
+  assumes "degree p \<noteq> 0"
+  shows "\<not> p dvd pderiv p"
+proof
+  assume dvd: "p dvd pderiv p"
+  then obtain q where p: "pderiv p = p * q"
+    unfolding dvd_def by auto
+  from dvd have le: "degree p \<le> degree (pderiv p)"
+    by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff)
+  from assms and this [unfolded degree_pderiv]
+    show False by auto
+qed
+
+lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \<longleftrightarrow> degree p = 0"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
+  using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric])
+
+lemma pderiv_singleton [simp]: "pderiv [:a:] = 0"
+  by (simp add: pderiv_pCons)
+
+lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q"
+  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
+
+lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p"
+  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
+
+lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q"
+  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
+
+lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)"
+  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
+
+lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p"
+  by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps)
+
+lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p"
+  apply (induct n)
+   apply simp
+  apply (subst power_Suc)
+  apply (subst pderiv_mult)
+  apply (erule ssubst)
+  apply (simp only: of_nat_Suc smult_add_left smult_1_left)
+  apply (simp add: algebra_simps)
+  done
+
+lemma pderiv_prod: "pderiv (prod f (as)) = (\<Sum>a\<in>as. prod f (as - {a}) * pderiv (f a))"
+proof (induct as rule: infinite_finite_induct)
+  case (insert a as)
+  then have id: "prod f (insert a as) = f a * prod f as"
+    "\<And>g. sum g (insert a as) = g a + sum g as"
+    "insert a as - {a} = as"
+    by auto
+  have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \<in> as" for b
+  proof -
+    from \<open>a \<notin> as\<close> that have *: "insert a as - {b} = insert a (as - {b})"
+      by auto
+    show ?thesis
+      unfolding * by (subst prod.insert) (use insert in auto)
+  qed
+  then show ?case
+    unfolding id pderiv_mult insert(3) sum_distrib_left
+    by (auto simp add: ac_simps intro!: sum.cong)
+qed auto
+
+lemma DERIV_pow2: "DERIV (\<lambda>x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"
+  by (rule DERIV_cong, rule DERIV_pow) simp
+declare DERIV_pow2 [simp] DERIV_pow [simp]
+
+lemma DERIV_add_const: "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. a + f x :: 'a::real_normed_field) x :> D"
+  by (rule DERIV_cong, rule DERIV_add) auto
+
+lemma poly_DERIV [simp]: "DERIV (\<lambda>x. poly p x) x :> poly (pderiv p) x"
+  by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons)
+
+lemma continuous_on_poly [continuous_intros]:
+  fixes p :: "'a :: {real_normed_field} poly"
+  assumes "continuous_on A f"
+  shows "continuous_on A (\<lambda>x. poly p (f x))"
+proof -
+  have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))"
+    by (intro continuous_intros assms)
+  also have "\<dots> = (\<lambda>x. poly p (f x))"
+    by (rule ext) (simp add: poly_altdef mult_ac)
+  finally show ?thesis .
+qed
+
+text \<open>Consequences of the derivative theorem above.\<close>
+
+lemma poly_differentiable[simp]: "(\<lambda>x. poly p x) differentiable (at x)"
+  for x :: real
+  by (simp add: real_differentiable_def) (blast intro: poly_DERIV)
+
+lemma poly_isCont[simp]: "isCont (\<lambda>x. poly p x) x"
+  for x :: real
+  by (rule poly_DERIV [THEN DERIV_isCont])
+
+lemma poly_IVT_pos: "a < b \<Longrightarrow> poly p a < 0 \<Longrightarrow> 0 < poly p b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0"
+  for a b :: real
+  using IVT_objl [of "poly p" a 0 b] by (auto simp add: order_le_less)
+
+lemma poly_IVT_neg: "a < b \<Longrightarrow> 0 < poly p a \<Longrightarrow> poly p b < 0 \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0"
+  for a b :: real
+  using poly_IVT_pos [where p = "- p"] by simp
+
+lemma poly_IVT: "a < b \<Longrightarrow> poly p a * poly p b < 0 \<Longrightarrow> \<exists>x>a. x < b \<and> poly p x = 0"
+  for p :: "real poly"
+  by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos)
+
+lemma poly_MVT: "a < b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p b - poly p a = (b - a) * poly (pderiv p) x"
+  for a b :: real
+  using MVT [of a b "poly p"]
+  apply auto
+  apply (rule_tac x = z in exI)
+  apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique])
+  done
+
+lemma poly_MVT':
+  fixes a b :: real
+  assumes "{min a b..max a b} \<subseteq> A"
+  shows "\<exists>x\<in>A. poly p b - poly p a = (b - a) * poly (pderiv p) x"
+proof (cases a b rule: linorder_cases)
+  case less
+  from poly_MVT[OF less, of p] guess x by (elim exE conjE)
+  then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms])
+next
+  case greater
+  from poly_MVT[OF greater, of p] guess x by (elim exE conjE)
+  then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms])
+qed (use assms in auto)
+
+lemma poly_pinfty_gt_lc:
+  fixes p :: "real poly"
+  assumes "lead_coeff p > 0"
+  shows "\<exists>n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p"
+  using assms
+proof (induct p)
+  case 0
+  then show ?case by auto
+next
+  case (pCons a p)
+  from this(1) consider "a \<noteq> 0" "p = 0" | "p \<noteq> 0" by auto
+  then show ?case
+  proof cases
+    case 1
+    then show ?thesis by auto
+  next
+    case 2
+    with pCons obtain n1 where gte_lcoeff: "\<forall>x\<ge>n1. lead_coeff p \<le> poly p x"
+      by auto
+    from pCons(3) \<open>p \<noteq> 0\<close> have gt_0: "lead_coeff p > 0" by auto
+    define n where "n = max n1 (1 + \<bar>a\<bar> / lead_coeff p)"
+    have "lead_coeff (pCons a p) \<le> poly (pCons a p) x" if "n \<le> x" for x
+    proof -
+      from gte_lcoeff that have "lead_coeff p \<le> poly p x"
+        by (auto simp: n_def)
+      with gt_0 have "\<bar>a\<bar> / lead_coeff p \<ge> \<bar>a\<bar> / poly p x" and "poly p x > 0"
+        by (auto intro: frac_le)
+      with \<open>n \<le> x\<close>[unfolded n_def] have "x \<ge> 1 + \<bar>a\<bar> / poly p x"
+        by auto
+      with \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x > 0\<close> \<open>p \<noteq> 0\<close>
+      show "lead_coeff (pCons a p) \<le> poly (pCons a p) x"
+        by (auto simp: field_simps)
+    qed
+    then show ?thesis by blast
+  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)"
+  by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons)
+
+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: "\<not> [:- a, 1:] dvd q"
+  shows "n = Suc (order a (pderiv p))"
+proof -
+  from assms have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0"
+    by auto
+  from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
+    by (cases n) auto
+  have *: "k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l" for k 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
+    show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
+      apply (subst lemma_order_pderiv1)
+      apply (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one)
+      done
+  qed
+  then show ?thesis
+    by (metis \<open>n = Suc n'\<close> pe)
+qed
+
+lemma order_pderiv: "pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a p = Suc (order a (pderiv p))"
+  for p :: "'a::field_char_0 poly"
+  apply (cases "p = 0")
+   apply 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:
+  fixes p :: "'a::field_char_0 poly"
+  assumes "pderiv p \<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: "\<not> ?thesis"
+  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)
+  from \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> have "order a (pderiv p) = order a e + order a d"
+    by (simp add: order_mult)
+  from \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> have "order a p = Suc (order a (pderiv p))"
+    by (rule order_pderiv)
+  from \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> have "d \<noteq> 0"
+    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
+  with \<open>d \<noteq> 0\<close> have "order a (pderiv p) \<le> order a d"
+    by (simp add: order_divides)
+  show ?thesis
+    using \<open>order a p = order a q + order a d\<close>
+      and \<open>order a (pderiv p) = order a e + order a d\<close>
+      and \<open>order a p = Suc (order a (pderiv p))\<close>
+      and \<open>order a (pderiv p) \<le> order a d\<close>
+    by auto
+qed
+
+lemma poly_squarefree_decomp_order2:
+  "pderiv p \<noteq> 0 \<Longrightarrow> p = q * d \<Longrightarrow> pderiv p = e * d \<Longrightarrow>
+    d = r * p + s * pderiv p \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+  for p :: "'a::field_char_0 poly"
+  by (blast intro: poly_squarefree_decomp_order)
+
+lemma order_pderiv2:
+  "pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a (pderiv p) = n \<longleftrightarrow> order a p = Suc n"
+  for p :: "'a::field_char_0 poly"
+  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:]"
+  for p :: "'a::{semidom,semiring_char_0} poly"
+  by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits)
+
+lemma rsquarefree_roots: "rsquarefree p \<longleftrightarrow> (\<forall>a. \<not> (poly p a = 0 \<and> poly (pderiv p) a = 0))"
+  for p :: "'a::field_char_0 poly"
+  apply (simp add: rsquarefree_def)
+  apply (case_tac "p = 0")
+   apply simp
+  apply 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:
+  fixes p :: "'a::field_char_0 poly"
+  assumes "pderiv p \<noteq> 0"
+    and "p = q * d"
+    and "pderiv p = e * d"
+    and "d = r * p + s * pderiv p"
+  shows "rsquarefree q \<and> (\<forall>a. poly q a = 0 \<longleftrightarrow> 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
+  from assms have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+    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>
+
+text \<open>
+  Algebraic numbers can be defined in two equivalent ways: all real numbers that are
+  roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry
+  uses the rational definition, but we need the integer definition.
+
+  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>
+
+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)"
+
+lemma algebraicI: "(\<And>i. coeff p i \<in> \<int>) \<Longrightarrow> p \<noteq> 0 \<Longrightarrow> poly p x = 0 \<Longrightarrow> algebraic x"
+  unfolding algebraic_def by blast
+
+lemma algebraicE:
+  assumes "algebraic x"
+  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 algebraic_altdef: "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
+  for p :: "'a::field_char_0 poly"
+proof safe
+  fix p
+  assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0"
+  define cs where "cs = coeffs p"
+  from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'"
+    unfolding Rats_def by blast
+  then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i
+    by (subst (asm) bchoice_iff) blast
+  define cs' where "cs' = map (quotient_of \<circ> f) (coeffs p)"
+  define d where "d = Lcm (set (map snd cs'))"
+  define p' where "p' = smult (of_int d) p"
+
+  have "coeff p' n \<in> \<int>" for n
+  proof (cases "n \<le> degree p")
+    case True
+    define c where "c = coeff p n"
+    define a where "a = fst (quotient_of (f (coeff p n)))"
+    define b where "b = snd (quotient_of (f (coeff p n)))"
+    have b_pos: "b > 0"
+      unfolding b_def using quotient_of_denom_pos' by simp
+    have "coeff p' n = of_int d * coeff p n"
+      by (simp add: p'_def)
+    also have "coeff p n = of_rat (of_int a / of_int b)"
+      unfolding a_def b_def
+      by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric])
+    also have "of_int d * \<dots> = of_rat (of_int (a*d) / of_int b)"
+      by (simp add: of_rat_mult of_rat_divide)
+    also from nz True have "b \<in> snd ` set cs'"
+      by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc)
+    then have "b dvd (a * d)"
+      by (simp add: d_def)
+    then have "of_int (a * d) / of_int b \<in> (\<int> :: rat set)"
+      by (rule of_int_divide_in_Ints)
+    then have "of_rat (of_int (a * d) / of_int b) \<in> \<int>" by (elim Ints_cases) auto
+    finally show ?thesis .
+  next
+    case False
+    then show ?thesis
+      by (auto simp: p'_def not_le coeff_eq_0)
+  qed
+  moreover have "set (map snd cs') \<subseteq> {0<..}"
+    unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc)
+  then have "d \<noteq> 0"
+    unfolding d_def by (induct cs') simp_all
+  with nz have "p' \<noteq> 0" by (simp add: p'_def)
+  moreover from root have "poly p' x = 0"
+    by (simp add: p'_def)
+  ultimately show "algebraic x"
+    unfolding algebraic_def by blast
+next
+  assume "algebraic x"
+  then obtain p where p: "coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0" for i
+    by (force simp: algebraic_def)
+  moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i
+    by (elim Ints_cases) simp
+  ultimately show "\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0" by auto
+qed
+
+
+subsection \<open>Content and primitive part of a polynomial\<close>
+
+definition content :: "'a::semiring_gcd poly \<Rightarrow> 'a"
+  where "content p = gcd_list (coeffs p)"
+
+lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0"
+  by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps)
+
+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: "[:c:] dvd p \<longleftrightarrow> c dvd content p"
+  for c :: "'a::semiring_gcd"
+proof (cases "p = 0")
+  case True
+  then show ?thesis by simp
+next
+  case 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"
+    then show "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_fin_iff dvd_mult_unit_iff)
+  finally show ?thesis .
+qed
+
+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"
+proof (cases "p = 0")
+  case True
+  then show ?thesis
+    by simp
+next
+  case False
+  then show ?thesis
+    by (cases "n \<le> degree p")
+      (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd)
+qed
+
+lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
+  by (simp add: content_def Gcd_fin_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)"
+  then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
+  then show "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_fin_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 poly \<Rightarrow> 'a poly"
+  where "primitive_part p = 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]: "smult (content p) (primitive_part p) = p"
+  for p :: "'a :: semiring_gcd poly"
+proof (cases "p = 0")
+  case True
+  then show ?thesis by simp
+next
+  case False
+  then show ?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
+
+lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
+proof (cases "p = 0")
+  case True
+  then show ?thesis by simp
+next
+  case False
+  then have "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
+
+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 (primitive_part p) * content p"
+    by (simp del: content_times_primitive_part add: ac_simps)
+  finally have "1 * content p = content (primitive_part p) * content p"
+    by simp
+  then have "1 * content p div content p = content (primitive_part p) * content p div content p"
+    by simp
+  with assms show ?thesis
+    by simp
+qed
+
+lemma content_decompose:
+  obtains p' :: "'a::semiring_gcd poly" where "p = smult (content p) p'" "content p' = 1"
+proof (cases "p = 0")
+  case True
+  then show ?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 (rule 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 True
+  then show ?thesis by simp
+next
+  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
+
+
+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'"
+    and "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> d * r = 0"
+  shows "q' = q + r"
+  using assms(3-)
+proof (induct n arbitrary: q r dr)
+  case (Suc n)
+  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)
+  from Suc(4) have dr: "dr = n + degree d" by auto
+  from d have lc: "lc \<noteq> 0" by auto
+  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
+  proof (cases "?qq = 0")
+    case True
+    then show ?thesis by simp
+  next
+    case False
+    then have n: "n = degree b"
+      by (simp add: degree_monom_eq)
+    show ?thesis
+      unfolding n dr by (simp add: coeff_mult_degree_sum)
+  qed
+  also have "\<dots> = lc * coeff b n"
+    by (simp add: d)
+  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
+    with Suc(2) show ?thesis
+      unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps)
+  next
+    case False
+    from dr Suc(2) have "degree r \<le> n"
+      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
+    then have right: "lc * coeff r n = 0"
+      by (simp add: coeff_eq_0)
+    have "coeff (d * r) dr = coeff (d * r) (degree d + n)"
+      by (simp add: dr ac_simps)
+    also from r_n have "\<dots> = 0"
+      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
+      by (simp only: 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]
+  from Suc(4) have dr: "dr = n + degree d" by auto
+  from Suc(2) have deg_rr: "degree ?rr \<le> dr" 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
+  next
+    case _: Suc
+    with Suc(4) show ?thesis by auto
+  qed
+  from IH[OF deg_rrr this] show ?case
+    by simp
+next
+  case 0
+  show ?case
+  proof (cases "r = 0")
+    case True
+    with 0 show ?thesis by auto
+  next
+    case False
+    from d False have "degree (d * r) = degree d + degree r"
+      by (subst degree_mult_eq) auto
+    with 0 d show ?thesis 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 0
+  then show ?case by simp
+next
+  case Suc
+  show ?case
+    unfolding divide_poly_main.simps[of _ _ r] Let_def
+    by (simp add: Suc del: divide_poly_main.simps)
+qed
+
+lemma divide_poly:
+  assumes g: "g \<noteq> 0"
+  shows "(f * g) div g = (f :: 'a poly)"
+proof -
+  have len: "length (coeffs f) = Suc (degree f)" if "f \<noteq> 0" for f :: "'a poly"
+    using that unfolding degree_eq_length_coeffs by auto
+  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"
+    by (simp add: divide_poly_def Let_def ac_simps)
+  note main = divide_poly_main[OF g refl le_refl 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 simp
+qed
+
+lemma divide_poly_0: "f div 0 = 0"
+  for f :: "'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 True
+  then show ?thesis
+    by (auto intro!: poly_eqI simp: coeff_map_poly)
+next
+  case False
+  from assms obtain q where p: "p = [:c:] * q" by (rule dvdE)
+  moreover {
+    have "smult c q = [:c:] * q"
+      by simp
+    also have "\<dots> div [:c:] = q"
+      by (rule nonzero_mult_div_cancel_left) (use False in auto)
+    finally have "smult c q div [:c:] = q" .
+  }
+  ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
+qed
+
+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: "a \<noteq> 0 \<Longrightarrow> is_unit [:a:]"
+  for a :: "'a::field"
+  by (simp add: is_unit_monom_0 monom_0 [symmetric])
+
+lemma is_unit_iff_degree:
+  fixes p :: "'a::field poly"
+  assumes "p \<noteq> 0"
+  shows "is_unit p \<longleftrightarrow> degree p = 0"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs
+  then obtain a where "p = [:a:]"
+    by (rule degree_eq_zeroE)
+  with assms show ?lhs
+    by (simp add: is_unit_triv)
+next
+  assume ?lhs
+  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 ?rhs by simp
+qed
+
+lemma is_unit_pCons_iff: "is_unit (pCons a p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
+  for p :: "'a::field poly"
+  by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree)
+
+lemma is_unit_monom_trival: "is_unit p \<Longrightarrow> monom (coeff p (degree p)) 0 = p"
+  for p :: "'a::field poly"
+  by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
+
+lemma is_unit_const_poly_iff: "[:c:] dvd 1 \<longleftrightarrow> c dvd 1"
+  for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
+  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:]" .
+  with \<open>p dvd 1\<close> have "c dvd 1"
+    by (simp add: is_unit_const_poly_iff)
+  with c show thesis ..
+qed
+
+lemma is_unit_polyE':
+  fixes p :: "'a::field poly"
+  assumes "is_unit p"
+  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: "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
+  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
+  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')"
+    and "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 assms(3-)
+proof (induct n arbitrary: q r dr)
+  case 0
+  then show ?case by auto
+next
+  case (Suc n)
+  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)
+  from Suc(4) have dr: "dr = n + degree d" by auto
+  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
+  proof (cases "?qq = 0")
+    case True
+    then show ?thesis by auto
+  next
+    case False
+    then have n: "n = degree b"
+      by (simp add: degree_monom_eq)
+    show ?thesis
+      unfolding n dr by (simp add: coeff_mult_degree_sum)
+  qed
+  also have "\<dots> = lc * coeff b n" by (simp add: d)
+  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
+  from Suc(4) have dr: "dr = n + degree d" by auto
+  have deg_rr: "degree ?rr \<le> dr"
+    using Suc(2) 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
+    then have "\<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
+  next
+    case _: Suc
+    with Suc(4) show ?thesis by auto
+  qed
+  note IH = Suc(1)[OF deg_rrr res this]
+  show ?case
+  proof (intro conjI)
+    from IH show "r' = 0 \<or> degree r' < degree d"
+      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
+
+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]
+  from g 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"
+    by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs)
+  note main' = main[OF this]
+  then 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,
+        cases "f = 0"; cases "coeffs g", use g in 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_all 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 :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly"
+  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"
+  from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)"
+    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"
+    by (auto simp: a_def)
+  show "r = 0 \<or> degree r < degree g" by fact
+  from g have "a \<noteq> 0"
+    by (auto simp: a_def)
+  with id show "\<exists>a q. a \<noteq> 0 \<and> smult a f = g * q + r"
+    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 \<dots> = 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"
+    by (simp only: Suc[unfolded divide_poly_main.simps Let_def])
+  also have "\<dots> = 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
+
+
+subsubsection \<open>Division in polynomials over fields\<close>
+
+lemma pseudo_divmod_field:
+  fixes g :: "'a::field poly"
+  assumes g: "g \<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"
+    by (auto simp: c_def)
+  from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r"
+    by auto
+  also have "smult (1 / c) \<dots> = 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:
+  fixes d :: "'a::field poly"
+  assumes d: "d \<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 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"
+    by (auto simp: f'_def)
+  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 :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom
+begin
+
+definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
+  where "unit_factor_poly p = [:unit_factor (lead_coeff p):]"
+
+definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
+  where "normalize p = p div [:unit_factor (lead_coeff p):]"
+
+instance
+proof
+  fix p :: "'a poly"
+  show "unit_factor p * normalize p = p"
+  proof (cases "p = 0")
+    case True
+    then show ?thesis
+      by (simp add: unit_factor_poly_def normalize_poly_def)
+  next
+    case False
+    then have "lead_coeff p \<noteq> 0"
+      by simp
+    then have *: "unit_factor (lead_coeff p) \<noteq> 0"
+      using unit_factor_is_unit [of "lead_coeff p"] by auto
+    then have "unit_factor (lead_coeff p) dvd 1"
+      by (auto intro: unit_factor_is_unit)
+    then have **: "unit_factor (lead_coeff p) dvd c" for c
+      by (rule dvd_trans) simp
+    have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c
+    proof -
+      from ** obtain b where "c = unit_factor (lead_coeff p) * b" ..
+      with False * show ?thesis by simp
+    qed
+    have "p div [:unit_factor (lead_coeff p):] =
+      map_poly (\<lambda>c. c div unit_factor (lead_coeff p)) p"
+      by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **)
+    then show ?thesis
+      by (simp add: normalize_poly_def unit_factor_poly_def
+        smult_conv_map_poly map_poly_map_poly o_def ***)
+  qed
+next
+  fix p :: "'a poly"
+  assume "is_unit p"
+  then obtain c where p: "p = [:c:]" "c dvd 1"
+    by (auto simp: is_unit_poly_iff)
+  then show "unit_factor p = p"
+    by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor)
+next
+  fix p :: "'a poly"
+  assume "p \<noteq> 0"
+  then show "is_unit (unit_factor p)"
+    by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit)
+qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
+
+end
+
+lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
+proof -
+  have "[:unit_factor (lead_coeff p):] dvd p"
+    by (metis unit_factor_poly_def unit_factor_self)
+  then show ?thesis
+    by (simp add: normalize_poly_def div_const_poly_conv_map_poly)
+qed
+
+lemma coeff_normalize [simp]:
+  "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)"
+  by (simp add: normalize_poly_eq_map_poly coeff_map_poly)
+
+class field_unit_factor = field + unit_factor +
+  assumes unit_factor_field [simp]: "unit_factor = id"
+begin
+
+subclass semidom_divide_unit_factor
+proof
+  fix a
+  assume "a \<noteq> 0"
+  then have "1 = a * inverse a" by simp
+  then have "a dvd 1" ..
+  then show "unit_factor a dvd 1" by simp
+qed simp_all
+
+end
+
+lemma unit_factor_pCons:
+  "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] 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_eq_map_poly degree_monom_eq)
+
+lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]"
+  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_eq_map_poly 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
+
+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
+
+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)"
+  by (simp add: eucl_rel_poly_iff)
+
+lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)"
+  by (simp add: eucl_rel_poly_iff)
+
+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 -
+  from assms have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
+    by (simp_all add: eucl_rel_poly_iff)
+  from b x have "?x = ?q * y + ?r" by simp
+  moreover
+  have "?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)
+      from r show "degree (pCons a r) \<le> degree y"
+        by auto
+      show "degree (smult b y) \<le> degree y"
+        by (rule degree_smult_le)
+    qed
+    from \<open>y \<noteq> 0\<close> show "coeff ?r (degree y) = 0"
+      by (simp add: b)
+  qed
+  ultimately 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 classical)
+    assume "\<not> ?thesis"
+    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 from \<open>q1 \<noteq> q2\<close> have "\<dots> = degree ((q1 - q2) * y)"
+      by (simp add: degree_mult_eq)
+    also from q3 have "\<dots> = degree (r2 - r1)"
+      by simp
+    finally have "degree (r2 - r1) < degree (r2 - r1)" .
+    then show ?thesis 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) semidom_modulo
+begin
+
+definition modulo_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+  where mod_poly_def: "f mod g =
+    (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)"
+
+instance
+proof
+  fix x y :: "'a poly"
+  show "x div y * y + x mod y = x"
+  proof (cases "y = 0")
+    case True
+    then show ?thesis
+      by (simp add: divide_poly_0 mod_poly_def)
+  next
+    case False
+    then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y =
+        (x div y, x mod y)"
+      by (simp add: divide_poly_field mod_poly_def pseudo_mod_def)
+    with False pseudo_divmod [OF False this] show ?thesis
+      by (simp add: power_mult_distrib [symmetric] ac_simps)
+  qed
+qed
+
+end
+
+lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)"
+  unfolding eucl_rel_poly_iff
+proof
+  show "x = x div y * y + x mod y"
+    by (simp add: div_mult_mod_eq)
+  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
+      by (simp add: mod_poly_def)
+  qed
+qed
+
+lemma div_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x div y = q"
+  for x :: "'a::field poly"
+  by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly])
+
+lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x mod y = r"
+  for x :: "'a::field poly"
+  by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly])
+
+instance poly :: (field) ring_div
+proof
+  fix x y z :: "'a poly"
+  assume "y \<noteq> 0"
+  with eucl_rel_poly [of x y] have "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
+    by (simp add: eucl_rel_poly_iff distrib_right)
+  then show "(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: 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
+
+lemma div_pCons_eq:
+  "pCons a p div q =
+    (if q = 0 then 0
+     else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))"
+  using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p]
+  by (auto intro: div_poly_eq)
+
+lemma mod_pCons_eq:
+  "pCons a p mod q =
+    (if q = 0 then pCons a p
+     else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)"
+  using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p]
+  by (auto intro: mod_poly_eq)
+
+lemma div_mod_fold_coeffs:
+  "(p div q, p mod 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))"
+  by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def)
+
+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:
+  fixes x :: "'a::field poly"
+  assumes "degree x < degree y"
+  shows "x div y = 0"
+proof -
+  from assms have "eucl_rel_poly x y (0, x)"
+    by (simp add: eucl_rel_poly_iff)
+  then show "x div y = 0"
+    by (rule div_poly_eq)
+qed
+
+lemma mod_poly_less:
+  assumes "degree x < degree y"
+  shows "x mod y = x"
+proof -
+  from assms have "eucl_rel_poly x y (0, x)"
+    by (simp add: eucl_rel_poly_iff)
+  then show "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)"
+  by (simp add: eucl_rel_poly_iff smult_add_right)
+
+lemma div_smult_left: "(smult a x) div y = smult a (x div y)"
+  for x y :: "'a::field poly"
+  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]: "(- x) div y = - (x div y)"
+  for x y :: "'a::field poly"
+  using div_smult_left [of "- 1::'a"] by simp
+
+lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)"
+  for x y :: "'a::field poly"
+  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: algebra_simps degree_add_less)
+
+lemma poly_div_add_left: "(x + y) div z = x div z + y div z"
+  for x y z :: "'a::field poly"
+  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
+  by (rule div_poly_eq)
+
+lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z"
+  for x y z :: "'a::field poly"
+  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
+  by (rule mod_poly_eq)
+
+lemma poly_div_diff_left: "(x - y) div z = x div z - y div z"
+  for x y z :: "'a::field poly"
+  by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left)
+
+lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z"
+  for x y z :: "'a::field poly"
+  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)"
+  by (simp add: eucl_rel_poly_iff)
+
+lemma div_smult_right: "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
+  for x y :: "'a::field poly"
+  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]: "x div (- y) = - (x div y)"
+  for x y :: "'a::field poly"
+  using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq)
+
+lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y"
+  for x y :: "'a::field poly"
+  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: "x div (y * z) = (x div y) div z"
+  for x y z :: "'a::field poly"
+  by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
+
+lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y"
+  for x y z :: "'a::field poly"
+  by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
+
+lemma mod_pCons:
+  fixes a :: "'a::field"
+    and x y :: "'a::field poly"
+  assumes y: "y \<noteq> 0"
+  defines "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_def
+  by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl])
+
+
+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"
+  for 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) y)"
+  for a :: "'a::ring"
+  by(cases "a = 0") (simp_all add: minus_zero_does_nothing)
+
+lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b"
+  for a :: "'a::comm_semiring_1 list"
+  by (induct a) (auto simp: monom_0 monom_Suc)
+
+lemma minus_poly_rev_list: "length p \<ge> length q \<Longrightarrow>
+  Poly (rev (minus_poly_rev_list (rev p) (rev q))) =
+    Poly p - monom 1 (length p - length q) * Poly q"
+  for p q :: "'a :: comm_ring_1 list"
+proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct)
+  case (1 x xs y ys)
+  then have "length (rev q) \<le> length (rev p)"
+    by simp
+  from this[folded 1(2,3)] have ys_xs: "length ys \<le> length xs"
+    by simp
+  then have *: "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 from ys_xs have "\<dots> =
+    Poly (rev xs) + monom x (length xs) -
+      (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length 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 * diff_monom[symmetric] by simp
+  finally show ?case
+    by (simp add: 1(2,3)[symmetric] 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)) r) (map (op * (hd r)) (rev d))) = 0"
+  for d r :: "'a::comm_ring list"
+proof (induct r)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons a rs)
+  then show ?case by (cases "rev d") (simp_all add: ac_simps)
+qed
+
+lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)"
+proof (induct p)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons x xs)
+  then show ?case by (cases "Poly xs = 0") auto
+qed
+
+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 "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)
+  from Suc.prems have *: "\<not> Suc (length r) \<le> length d"
+    by simp
+  with \<open>d \<noteq> []\<close> have "r \<noteq> []"
+    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)"
+  from * Suc(3) have n: "n = (1 + length r - length d - 1)"
+    by simp
+  from * have rr_val:"(length ?rrr) = (length r - 1)"
+    by auto
+  with \<open>r \<noteq> []\<close> * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)"
+    by auto
+  from * have id: "Suc (length r) - length d = Suc (length r - length d)"
+    by auto
+  from Suc.prems *
+  have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')"
+    by (simp add: Let_def if_0_minus_poly_rev_list id)
+  with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')"
+    by auto
+  from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)"
+    using Suc_diff_le not_less_eq_eq by blast
+  from Suc(3) \<open>r \<noteq> []\<close> have n_ok : "n = 1 + (length ?rrr) - length d"
+    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 \<open>r \<noteq> []\<close>] 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
+      by (simp add: monom_Suc hd_rev[symmetric] 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 * \<open>d \<noteq> []\<close>, auto)
+      from * 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 g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))"
+    for f g :: "'a::comm_ring_1 poly"
+proof (cases "g = 0")
+  case False
+  then have "last (coeffs g) \<noteq> 0"
+    and "last (coeffs g) = lead_coeff g"
+    and "coeffs g \<noteq> []"
+    by (simp_all add: last_coeffs_eq_coeff_degree)
+  moreover 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
+  ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g
+    (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))"
+    by (subst pseudo_divmod_main_list_invar [symmetric]) auto
+  moreover have "pseudo_divmod_main_list
+    (hd (rev (coeffs g))) []
+    (rev (coeffs f)) (rev (coeffs g))
+    (1 + length (coeffs f) -
+    length (coeffs g)) = (q, r)"
+    using qr hd_rev [OF \<open>coeffs g \<noteq> []\<close>] by simp
+  ultimately show ?thesis
+    by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def)
+next
+  case True
+  then show ?thesis
+    by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def)
+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> (p div q, p mod q) = (r, s)"
+  by (metis eucl_rel_poly eucl_rel_poly_unique)
+
+lemma pdivmod_via_pseudo_divmod:
+  "(f div g, f mod 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 True
+  then show ?thesis by simp
+next
+  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"
+    by (auto simp: h_def lc_def)
+  then have 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)"
+    by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p)
+  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
+  from f r h0 have "eucl_rel_poly f h (q, r)"
+    by (auto simp: eucl_rel_poly_iff)
+  then have "(f div h, f mod h) = (q, r)"
+    by (simp add: pdivmod_pdivmodrel)
+  with lc have "(f div g, f mod g) = (smult lc q, r)"
+    by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc])
+  with id show ?thesis
+    by auto
+qed
+
+lemma pdivmod_via_pseudo_divmod_list:
+  "(f div g, f mod 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
+    with d show ?thesis by auto
+  next
+    case False
+    define ilc where "ilc = inverse (coeff g (degree g))"
+    from False have ilc: "ilc \<noteq> 0"
+      by (auto simp: ilc_def)
+    with False have id: "g = 0 \<longleftrightarrow> False" "coeffs g = [] \<longleftrightarrow> False"
+      "last (coeffs g) = coeff g (degree g)"
+      "coeffs (smult ilc g) = [] \<longleftrightarrow> 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)
+  have *: "map (op * 1) xs = xs" for xs :: "'a list"
+    by (induct xs) auto
+  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 = 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 "_ = ?rhs")
+proof -
+  have "snd (f div g, f mod g) = ?rhs"
+    unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil]
+    by (auto split: prod.splits)
+  then show ?thesis by simp
+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 "fst (f div g, f mod g) = div_field_poly_impl f g"
+    unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def
+    by (auto split: prod.splits)
+  then show "f div g =  div_field_poly_impl f g"
+    by simp
+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)
+  from Suc.prems have ifCond: "\<not> Suc (length r) \<le> length d"
+    by simp
+  with d have r: "r \<noteq> []"
+    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
+    with r d show ?thesis
+      unfolding Suc(2)[symmetric]
+      by (auto simp add: Let_def nth_default_append)
+  next
+    case True
+    with r d have id:
+      "?thesis \<longleftrightarrow>
+        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"
+      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)
+      then show ?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 by (auto simp: d True)
+  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 "lcg \<noteq> 0" by auto
+    from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = 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 \<open>lcg \<noteq> 0\<close>)
+        (auto simp: lcg ltp, simp add: degree_eq_length_coeffs)
+  qed
+qed
+
+no_notation cCons (infixr "##" 65)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Polynomial_FPS.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,313 @@
+(*  Title:      HOL/Library/Polynomial_FPS.thy
+    Author:     Manuel Eberl, TU München
+  
+Converting polynomials to formal power series.
+*)
+
+section \<open>Converting polynomials to formal power series\<close>
+
+theory Polynomial_FPS
+  imports Polynomial Formal_Power_Series
+begin
+
+definition fps_of_poly where
+  "fps_of_poly p = Abs_fps (coeff p)"
+
+lemma fps_of_poly_eq_iff: "fps_of_poly p = fps_of_poly q \<longleftrightarrow> p = q"
+  by (simp add: fps_of_poly_def poly_eq_iff fps_eq_iff)
+
+lemma fps_of_poly_nth [simp]: "fps_of_poly p $ n = coeff p n"
+  by (simp add: fps_of_poly_def)
+  
+lemma fps_of_poly_const: "fps_of_poly [:c:] = fps_const c"
+proof (subst fps_eq_iff, clarify)
+  fix n :: nat show "fps_of_poly [:c:] $ n = fps_const c $ n"
+    by (cases n) (auto simp: fps_of_poly_def)
+qed
+
+lemma fps_of_poly_0 [simp]: "fps_of_poly 0 = 0"
+  by (subst fps_const_0_eq_0 [symmetric], subst fps_of_poly_const [symmetric]) simp
+
+lemma fps_of_poly_1 [simp]: "fps_of_poly 1 = 1"
+  by (subst fps_const_1_eq_1 [symmetric], subst fps_of_poly_const [symmetric])
+     (simp add: one_poly_def)
+
+lemma fps_of_poly_1' [simp]: "fps_of_poly [:1:] = 1"
+  by (subst fps_const_1_eq_1 [symmetric], subst fps_of_poly_const [symmetric])
+     (simp add: one_poly_def)
+
+lemma fps_of_poly_numeral [simp]: "fps_of_poly (numeral n) = numeral n"
+  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
+
+lemma fps_of_poly_numeral' [simp]: "fps_of_poly [:numeral n:] = numeral n"
+  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
+
+lemma fps_of_poly_X [simp]: "fps_of_poly [:0, 1:] = X"
+  by (auto simp add: fps_of_poly_def fps_eq_iff coeff_pCons split: nat.split)
+
+lemma fps_of_poly_add: "fps_of_poly (p + q) = fps_of_poly p + fps_of_poly q"
+  by (simp add: fps_of_poly_def plus_poly.rep_eq fps_plus_def)
+
+lemma fps_of_poly_diff: "fps_of_poly (p - q) = fps_of_poly p - fps_of_poly q"
+  by (simp add: fps_of_poly_def minus_poly.rep_eq fps_minus_def)
+
+lemma fps_of_poly_uminus: "fps_of_poly (-p) = -fps_of_poly p"
+  by (simp add: fps_of_poly_def uminus_poly.rep_eq fps_uminus_def)
+
+lemma fps_of_poly_mult: "fps_of_poly (p * q) = fps_of_poly p * fps_of_poly q"
+  by (simp add: fps_of_poly_def fps_times_def fps_eq_iff coeff_mult atLeast0AtMost)
+
+lemma fps_of_poly_smult: 
+  "fps_of_poly (smult c p) = fps_const c * fps_of_poly p"
+  using fps_of_poly_mult[of "[:c:]" p] by (simp add: fps_of_poly_mult fps_of_poly_const)
+  
+lemma fps_of_poly_sum: "fps_of_poly (sum f A) = sum (\<lambda>x. fps_of_poly (f x)) A"
+  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_add)
+
+lemma fps_of_poly_sum_list: "fps_of_poly (sum_list xs) = sum_list (map fps_of_poly xs)"
+  by (induction xs) (simp_all add: fps_of_poly_add)
+  
+lemma fps_of_poly_prod: "fps_of_poly (prod f A) = prod (\<lambda>x. fps_of_poly (f x)) A"
+  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_mult)
+  
+lemma fps_of_poly_prod_list: "fps_of_poly (prod_list xs) = prod_list (map fps_of_poly xs)"
+  by (induction xs) (simp_all add: fps_of_poly_mult)
+
+lemma fps_of_poly_pCons: 
+  "fps_of_poly (pCons (c :: 'a :: semiring_1) p) = fps_const c + fps_of_poly p * X"
+  by (subst fps_mult_X_commute [symmetric], intro fps_ext) 
+     (auto simp: fps_of_poly_def coeff_pCons split: nat.split)
+  
+lemma fps_of_poly_pderiv: "fps_of_poly (pderiv p) = fps_deriv (fps_of_poly p)"
+  by (intro fps_ext) (simp add: fps_of_poly_nth coeff_pderiv)
+
+lemma fps_of_poly_power: "fps_of_poly (p ^ n) = fps_of_poly p ^ n"
+  by (induction n) (simp_all add: fps_of_poly_mult)
+  
+lemma fps_of_poly_monom: "fps_of_poly (monom (c :: 'a :: comm_ring_1) n) = fps_const c * X ^ n"
+  by (intro fps_ext) simp_all
+
+lemma fps_of_poly_monom': "fps_of_poly (monom (1 :: 'a :: comm_ring_1) n) = X ^ n"
+  by (simp add: fps_of_poly_monom)
+
+lemma fps_of_poly_div:
+  assumes "(q :: 'a :: field poly) dvd p"
+  shows   "fps_of_poly (p div q) = fps_of_poly p / fps_of_poly q"
+proof (cases "q = 0")
+  case False
+  from False fps_of_poly_eq_iff[of q 0] have nz: "fps_of_poly q \<noteq> 0" by simp 
+  from assms have "p = (p div q) * q" by simp
+  also have "fps_of_poly \<dots> = fps_of_poly (p div q) * fps_of_poly q" 
+    by (simp add: fps_of_poly_mult)
+  also from nz have "\<dots> / fps_of_poly q = fps_of_poly (p div q)"
+    by (intro nonzero_mult_div_cancel_right) (auto simp: fps_of_poly_0)
+  finally show ?thesis ..
+qed simp
+
+lemma fps_of_poly_divide_numeral:
+  "fps_of_poly (smult (inverse (numeral c :: 'a :: field)) p) = fps_of_poly p / numeral c"
+proof -
+  have "smult (inverse (numeral c)) p = [:inverse (numeral c):] * p" by simp
+  also have "fps_of_poly \<dots> = fps_of_poly p / numeral c"
+    by (subst fps_of_poly_mult) (simp add: numeral_fps_const fps_of_poly_pCons)
+  finally show ?thesis by simp
+qed
+
+
+lemma subdegree_fps_of_poly:
+  assumes "p \<noteq> 0"
+  defines "n \<equiv> Polynomial.order 0 p"
+  shows   "subdegree (fps_of_poly p) = n"
+proof (rule subdegreeI)
+  from assms have "monom 1 n dvd p" by (simp add: monom_1_dvd_iff)
+  thus zero: "fps_of_poly p $ i = 0" if "i < n" for i
+    using that by (simp add: monom_1_dvd_iff')
+    
+  from assms have "\<not>monom 1 (Suc n) dvd p"
+    by (auto simp: monom_1_dvd_iff simp del: power_Suc)
+  then obtain k where k: "k \<le> n" "fps_of_poly p $ k \<noteq> 0" 
+    by (auto simp: monom_1_dvd_iff' less_Suc_eq_le)
+  with zero[of k] have "k = n" by linarith
+  with k show "fps_of_poly p $ n \<noteq> 0" by simp
+qed
+
+lemma fps_of_poly_dvd:
+  assumes "p dvd q"
+  shows   "fps_of_poly (p :: 'a :: field poly) dvd fps_of_poly q"
+proof (cases "p = 0 \<or> q = 0")
+  case False
+  with assms fps_of_poly_eq_iff[of p 0] fps_of_poly_eq_iff[of q 0] show ?thesis
+    by (auto simp: fps_dvd_iff subdegree_fps_of_poly dvd_imp_order_le)
+qed (insert assms, auto)
+
+
+lemmas fps_of_poly_simps =
+  fps_of_poly_0 fps_of_poly_1 fps_of_poly_numeral fps_of_poly_const fps_of_poly_X
+  fps_of_poly_add fps_of_poly_diff fps_of_poly_uminus fps_of_poly_mult fps_of_poly_smult
+  fps_of_poly_sum fps_of_poly_sum_list fps_of_poly_prod fps_of_poly_prod_list
+  fps_of_poly_pCons fps_of_poly_pderiv fps_of_poly_power fps_of_poly_monom
+  fps_of_poly_divide_numeral
+
+lemma fps_of_poly_pcompose:
+  assumes "coeff q 0 = (0 :: 'a :: idom)"
+  shows   "fps_of_poly (pcompose p q) = fps_compose (fps_of_poly p) (fps_of_poly q)"
+  using assms by (induction p rule: pCons_induct)
+                 (auto simp: pcompose_pCons fps_of_poly_simps fps_of_poly_pCons 
+                             fps_compose_add_distrib fps_compose_mult_distrib)
+  
+lemmas reify_fps_atom =
+  fps_of_poly_0 fps_of_poly_1' fps_of_poly_numeral' fps_of_poly_const fps_of_poly_X
+
+
+text \<open>
+  The following simproc can reduce the equality of two polynomial FPSs two equality of the
+  respective polynomials. A polynomial FPS is one that only has finitely many non-zero 
+  coefficients and can therefore be written as @{term "fps_of_poly p"} for some 
+  polynomial \<open>p\<close>.
+  
+  This may sound trivial, but it covers a number of annoying side conditions like 
+  @{term "1 + X \<noteq> 0"} that would otherwise not be solved automatically.
+\<close>
+
+ML \<open>
+
+(* TODO: Support for division *)
+signature POLY_FPS = sig
+
+val reify_conv : conv
+val eq_conv : conv
+val eq_simproc : cterm -> thm option
+
+end
+
+
+structure Poly_Fps = struct
+
+fun const_binop_conv s conv ct =
+  case Thm.term_of ct of
+    (Const (s', _) $ _ $ _) => 
+      if s = s' then 
+        Conv.binop_conv conv ct 
+      else 
+        raise CTERM ("const_binop_conv", [ct])
+  | _ => raise CTERM ("const_binop_conv", [ct])
+
+fun reify_conv ct = 
+  let
+    val rewr = Conv.rewrs_conv o map (fn thm => thm RS @{thm eq_reflection})
+    val un = Conv.arg_conv reify_conv
+    val bin = Conv.binop_conv reify_conv
+  in
+    case Thm.term_of ct of
+      (Const (@{const_name "fps_of_poly"}, _) $ _) => ct |> Conv.all_conv
+    | (Const (@{const_name "Groups.plus"}, _) $ _ $ _) => ct |> (
+        bin then_conv rewr @{thms fps_of_poly_add [symmetric]})
+    | (Const (@{const_name "Groups.uminus"}, _) $ _) => ct |> (
+        un then_conv rewr @{thms fps_of_poly_uminus [symmetric]})
+    | (Const (@{const_name "Groups.minus"}, _) $ _ $ _) => ct |> (
+        bin then_conv rewr @{thms fps_of_poly_diff [symmetric]})
+    | (Const (@{const_name "Groups.times"}, _) $ _ $ _) => ct |> (
+        bin then_conv rewr @{thms fps_of_poly_mult [symmetric]})
+    | (Const (@{const_name "Rings.divide"}, _) $ _ $ (Const (@{const_name "Num.numeral"}, _) $ _))
+        => ct |> (Conv.fun_conv (Conv.arg_conv reify_conv)
+             then_conv rewr @{thms fps_of_poly_divide_numeral [symmetric]})
+    | (Const (@{const_name "Power.power"}, _) $ Const (@{const_name "X"},_) $ _) => ct |> (
+        rewr @{thms fps_of_poly_monom' [symmetric]}) 
+    | (Const (@{const_name "Power.power"}, _) $ _ $ _) => ct |> (
+        Conv.fun_conv (Conv.arg_conv reify_conv) 
+        then_conv rewr @{thms fps_of_poly_power [symmetric]})
+    | _ => ct |> (
+        rewr @{thms reify_fps_atom [symmetric]})
+  end
+    
+
+fun eq_conv ct =
+  case Thm.term_of ct of
+    (Const (@{const_name "HOL.eq"}, _) $ _ $ _) => ct |> (
+      Conv.binop_conv reify_conv
+      then_conv Conv.rewr_conv @{thm fps_of_poly_eq_iff[THEN eq_reflection]})
+  | _ => raise CTERM ("poly_fps_eq_conv", [ct])
+
+val eq_simproc = try eq_conv
+
+end
+\<close> 
+
+simproc_setup poly_fps_eq ("(f :: 'a fps) = g") = \<open>K (K Poly_Fps.eq_simproc)\<close>
+
+lemma fps_of_poly_linear: "fps_of_poly [:a,1 :: 'a :: field:] = X + fps_const a"
+  by simp
+
+lemma fps_of_poly_linear': "fps_of_poly [:1,a :: 'a :: field:] = 1 + fps_const a * X"
+  by simp
+
+lemma fps_of_poly_cutoff [simp]: 
+  "fps_of_poly (poly_cutoff n p) = fps_cutoff n (fps_of_poly p)"
+  by (simp add: fps_eq_iff coeff_poly_cutoff)
+
+lemma fps_of_poly_shift [simp]: "fps_of_poly (poly_shift n p) = fps_shift n (fps_of_poly p)"
+  by (simp add: fps_eq_iff coeff_poly_shift)
+
+
+definition poly_subdegree :: "'a::zero poly \<Rightarrow> nat" where
+  "poly_subdegree p = subdegree (fps_of_poly p)"
+
+lemma coeff_less_poly_subdegree:
+  "k < poly_subdegree p \<Longrightarrow> coeff p k = 0"
+  unfolding poly_subdegree_def using nth_less_subdegree_zero[of k "fps_of_poly p"] by simp
+
+(* TODO: Move ? *)
+definition prefix_length :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat" where
+  "prefix_length P xs = length (takeWhile P xs)"
+
+primrec prefix_length_aux :: "('a \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> nat" where
+  "prefix_length_aux P acc [] = acc"
+| "prefix_length_aux P acc (x#xs) = (if P x then prefix_length_aux P (Suc acc) xs else acc)"
+
+lemma prefix_length_aux_correct: "prefix_length_aux P acc xs = prefix_length P xs + acc"
+  by (induction xs arbitrary: acc) (simp_all add: prefix_length_def)
+
+lemma prefix_length_code [code]: "prefix_length P xs = prefix_length_aux P 0 xs"
+  by (simp add: prefix_length_aux_correct)
+
+lemma prefix_length_le_length: "prefix_length P xs \<le> length xs"
+  by (induction xs) (simp_all add: prefix_length_def)
+  
+lemma prefix_length_less_length: "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> prefix_length P xs < length xs"
+  by (induction xs) (simp_all add: prefix_length_def)
+
+lemma nth_prefix_length:
+  "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> \<not>P (xs ! prefix_length P xs)"
+  by (induction xs) (simp_all add: prefix_length_def)
+  
+lemma nth_less_prefix_length:
+  "n < prefix_length P xs \<Longrightarrow> P (xs ! n)"
+  by (induction xs arbitrary: n) 
+     (auto simp: prefix_length_def nth_Cons split: if_splits nat.splits)
+(* END TODO *)
+  
+lemma poly_subdegree_code [code]: "poly_subdegree p = prefix_length (op = 0) (coeffs p)"
+proof (cases "p = 0")
+  case False
+  note [simp] = this
+  define n where "n = prefix_length (op = 0) (coeffs p)"
+  from False have "\<exists>k. coeff p k \<noteq> 0" by (auto simp: poly_eq_iff)
+  hence ex: "\<exists>x\<in>set (coeffs p). x \<noteq> 0" by (auto simp: coeffs_def)
+  hence n_less: "n < length (coeffs p)" and nonzero: "coeffs p ! n \<noteq> 0" 
+    unfolding n_def by (auto intro!: prefix_length_less_length nth_prefix_length)
+  show ?thesis unfolding poly_subdegree_def
+  proof (intro subdegreeI)
+    from n_less have "fps_of_poly p $ n = coeffs p ! n"
+      by (subst coeffs_nth) (simp_all add: degree_eq_length_coeffs)
+    with nonzero show "fps_of_poly p $ prefix_length (op = 0) (coeffs p) \<noteq> 0"
+      unfolding n_def by simp
+  next
+    fix k assume A: "k < prefix_length (op = 0) (coeffs p)"
+    also have "\<dots> \<le> length (coeffs p)" by (rule prefix_length_le_length)
+    finally show "fps_of_poly p $ k = 0"
+      using nth_less_prefix_length[OF A]
+      by (simp add: coeffs_nth degree_eq_length_coeffs)
+  qed
+qed (simp_all add: poly_subdegree_def prefix_length_def)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,992 @@
+(*  Title:      HOL/Library/Polynomial_Factorial.thy
+    Author:     Brian Huffman
+    Author:     Clemens Ballarin
+    Author:     Amine Chaieb
+    Author:     Florian Haftmann
+    Author:     Manuel Eberl
+*)
+
+theory Polynomial_Factorial
+imports 
+  Complex_Main
+  Polynomial
+  Normalized_Fraction
+  Field_as_Ring
+begin
+
+subsection \<open>Various facts about polynomials\<close>
+
+lemma prod_mset_const_poly: " (\<Prod>x\<in>#A. [:f x:]) = [:prod_mset (image_mset f A):]"
+  by (induct A) (simp_all add: one_poly_def ac_simps)
+
+lemma irreducible_const_poly_iff:
+  fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
+  shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
+proof
+  assume A: "irreducible c"
+  show "irreducible [:c:]"
+  proof (rule irreducibleI)
+    fix a b assume ab: "[:c:] = a * b"
+    hence "degree [:c:] = degree (a * b)" by (simp only: )
+    also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
+    hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
+    finally have "degree a = 0" "degree b = 0" by auto
+    then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
+    from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
+    hence "c = a' * b'" by (simp add: ab' mult_ac)
+    from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
+    with ab' show "a dvd 1 \<or> b dvd 1" by (auto simp: one_poly_def)
+  qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
+next
+  assume A: "irreducible [:c:]"
+  show "irreducible c"
+  proof (rule irreducibleI)
+    fix a b assume ab: "c = a * b"
+    hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
+    from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
+    thus "a dvd 1 \<or> b dvd 1" by (simp add: one_poly_def)
+  qed (insert A, auto simp: irreducible_def one_poly_def)
+qed
+
+
+subsection \<open>Lifting elements into the field of fractions\<close>
+
+definition to_fract :: "'a :: idom \<Rightarrow> 'a fract" where "to_fract x = Fract x 1"
+  \<comment> \<open>FIXME: name \<open>of_idom\<close>, abbreviation\<close>
+
+lemma to_fract_0 [simp]: "to_fract 0 = 0"
+  by (simp add: to_fract_def eq_fract Zero_fract_def)
+
+lemma to_fract_1 [simp]: "to_fract 1 = 1"
+  by (simp add: to_fract_def eq_fract One_fract_def)
+
+lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y"
+  by (simp add: to_fract_def)
+
+lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y"
+  by (simp add: to_fract_def)
+  
+lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x"
+  by (simp add: to_fract_def)
+  
+lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y"
+  by (simp add: to_fract_def)
+
+lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \<longleftrightarrow> x = y"
+  by (simp add: to_fract_def eq_fract)
+  
+lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \<longleftrightarrow> x = 0"
+  by (simp add: to_fract_def Zero_fract_def eq_fract)
+
+lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \<noteq> 0"
+  by transfer simp
+
+lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x"
+  by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp)
+
+lemma to_fract_quot_of_fract:
+  assumes "snd (quot_of_fract x) = 1"
+  shows   "to_fract (fst (quot_of_fract x)) = x"
+proof -
+  have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp
+  also note assms
+  finally show ?thesis by (simp add: to_fract_def)
+qed
+
+lemma snd_quot_of_fract_Fract_whole:
+  assumes "y dvd x"
+  shows   "snd (quot_of_fract (Fract x y)) = 1"
+  using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd)
+  
+lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b"
+  by (simp add: to_fract_def)
+
+lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)"
+  unfolding to_fract_def by transfer (simp add: normalize_quot_def)
+
+lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \<longleftrightarrow> x = 0"
+  by transfer simp
+ 
+lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1"
+  unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all
+
+lemma coprime_quot_of_fract:
+  "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))"
+  by transfer (simp add: coprime_normalize_quot)
+
+lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1"
+  using quot_of_fract_in_normalized_fracts[of x] 
+  by (simp add: normalized_fracts_def case_prod_unfold)  
+
+lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \<Longrightarrow> normalize x = x"
+  by (subst (2) normalize_mult_unit_factor [symmetric, of x])
+     (simp del: normalize_mult_unit_factor)
+  
+lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)"
+  by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract)
+
+
+subsection \<open>Lifting polynomial coefficients to the field of fractions\<close>
+
+abbreviation (input) fract_poly 
+  where "fract_poly \<equiv> map_poly to_fract"
+
+abbreviation (input) unfract_poly 
+  where "unfract_poly \<equiv> map_poly (fst \<circ> quot_of_fract)"
+  
+lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)"
+  by (simp add: smult_conv_map_poly map_poly_map_poly o_def)
+
+lemma fract_poly_0 [simp]: "fract_poly 0 = 0"
+  by (simp add: poly_eqI coeff_map_poly)
+
+lemma fract_poly_1 [simp]: "fract_poly 1 = 1"
+  by (simp add: one_poly_def map_poly_pCons)
+
+lemma fract_poly_add [simp]:
+  "fract_poly (p + q) = fract_poly p + fract_poly q"
+  by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+lemma fract_poly_diff [simp]:
+  "fract_poly (p - q) = fract_poly p - fract_poly q"
+  by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\<lambda>x. to_fract (f x)) A"
+  by (cases "finite A", induction A rule: finite_induct) simp_all 
+
+lemma fract_poly_mult [simp]:
+  "fract_poly (p * q) = fract_poly p * fract_poly q"
+  by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult)
+
+lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \<longleftrightarrow> p = q"
+  by (auto simp: poly_eq_iff coeff_map_poly)
+
+lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \<longleftrightarrow> p = 0"
+  using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff)
+
+lemma fract_poly_dvd: "p dvd q \<Longrightarrow> fract_poly p dvd fract_poly q"
+  by (auto elim!: dvdE)
+
+lemma prod_mset_fract_poly: 
+  "(\<Prod>x\<in>#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))"
+  by (induct A) (simp_all add: ac_simps)
+  
+lemma is_unit_fract_poly_iff:
+  "p dvd 1 \<longleftrightarrow> fract_poly p dvd 1 \<and> content p = 1"
+proof safe
+  assume A: "p dvd 1"
+  with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)"
+    by simp
+  from A show "content p = 1"
+    by (auto simp: is_unit_poly_iff normalize_1_iff)
+next
+  assume A: "fract_poly p dvd 1" and B: "content p = 1"
+  from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff)
+  {
+    fix n :: nat assume "n > 0"
+    have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly)
+    also note c
+    also from \<open>n > 0\<close> have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits)
+    finally have "coeff p n = 0" by simp
+  }
+  hence "degree p \<le> 0" by (intro degree_le) simp_all
+  with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE)
+qed
+  
+lemma fract_poly_is_unit: "p dvd 1 \<Longrightarrow> fract_poly p dvd 1"
+  using fract_poly_dvd[of p 1] by simp
+
+lemma fract_poly_smult_eqE:
+  fixes c :: "'a :: {idom_divide,ring_gcd} fract"
+  assumes "fract_poly p = smult c (fract_poly q)"
+  obtains a b 
+    where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a"
+proof -
+  define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)"
+  have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)"
+    by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms)
+  hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff)
+  hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff)
+  moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b"
+    by (simp_all add: a_def b_def coprime_quot_of_fract gcd.commute
+          normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric])
+  ultimately show ?thesis by (intro that[of a b])
+qed
+
+
+subsection \<open>Fractional content\<close>
+
+abbreviation (input) Lcm_coeff_denoms 
+    :: "'a :: {semiring_Gcd,idom_divide,ring_gcd} fract poly \<Rightarrow> 'a"
+  where "Lcm_coeff_denoms p \<equiv> Lcm (snd ` quot_of_fract ` set (coeffs p))"
+  
+definition fract_content :: 
+      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a fract" where
+  "fract_content p = 
+     (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" 
+
+definition primitive_part_fract :: 
+      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a poly" where
+  "primitive_part_fract p = 
+     primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))"
+
+lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0"
+  by (simp add: primitive_part_fract_def)
+
+lemma fract_content_eq_0_iff [simp]:
+  "fract_content p = 0 \<longleftrightarrow> p = 0"
+  unfolding fract_content_def Let_def Zero_fract_def
+  by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff)
+
+lemma content_primitive_part_fract [simp]: "p \<noteq> 0 \<Longrightarrow> content (primitive_part_fract p) = 1"
+  unfolding primitive_part_fract_def
+  by (rule content_primitive_part)
+     (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff)  
+
+lemma content_times_primitive_part_fract:
+  "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p"
+proof -
+  define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)"
+  have "fract_poly p' = 
+          map_poly (to_fract \<circ> fst \<circ> quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)"
+    unfolding primitive_part_fract_def p'_def 
+    by (subst map_poly_map_poly) (simp_all add: o_assoc)
+  also have "\<dots> = smult (to_fract (Lcm_coeff_denoms p)) p"
+  proof (intro map_poly_idI, unfold o_apply)
+    fix c assume "c \<in> set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))"
+    then obtain c' where c: "c' \<in> set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'"
+      by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits)
+    note c(2)
+    also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))"
+      by simp
+    also have "to_fract (Lcm_coeff_denoms p) * \<dots> = 
+                 Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))"
+      unfolding to_fract_def by (subst mult_fract) simp_all
+    also have "snd (quot_of_fract \<dots>) = 1"
+      by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto)
+    finally show "to_fract (fst (quot_of_fract c)) = c"
+      by (rule to_fract_quot_of_fract)
+  qed
+  also have "p' = smult (content p') (primitive_part p')" 
+    by (rule content_times_primitive_part [symmetric])
+  also have "primitive_part p' = primitive_part_fract p"
+    by (simp add: primitive_part_fract_def p'_def)
+  also have "fract_poly (smult (content p') (primitive_part_fract p)) = 
+               smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp
+  finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) =
+                      smult (to_fract (Lcm_coeff_denoms p)) p" .
+  thus ?thesis
+    by (subst (asm) smult_eq_iff)
+       (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def)
+qed
+
+lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)"
+proof -
+  have "Lcm_coeff_denoms (fract_poly p) = 1"
+    by (auto simp: set_coeffs_map_poly)
+  hence "fract_content (fract_poly p) = 
+           to_fract (content (map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p))"
+    by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff)
+  also have "map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p = p"
+    by (intro map_poly_idI) simp_all
+  finally show ?thesis .
+qed
+
+lemma content_decompose_fract:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly"
+  obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1"
+proof (cases "p = 0")
+  case True
+  hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all
+  thus ?thesis ..
+next
+  case False
+  thus ?thesis
+    by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract])
+qed
+
+
+subsection \<open>More properties of content and primitive part\<close>
+
+lemma lift_prime_elem_poly:
+  assumes "prime_elem (c :: 'a :: semidom)"
+  shows   "prime_elem [:c:]"
+proof (rule prime_elemI)
+  fix a b assume *: "[:c:] dvd a * b"
+  from * have dvd: "c dvd coeff (a * b) n" for n
+    by (subst (asm) const_poly_dvd_iff) blast
+  {
+    define m where "m = (GREATEST m. \<not>c dvd coeff b m)"
+    assume "\<not>[:c:] dvd b"
+    hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast
+    have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i < Suc (degree b)"
+      by (auto intro: le_degree simp: less_Suc_eq_le)
+    have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex[OF A B])
+    have "i \<le> m" if "\<not>c dvd coeff b i" for i
+      unfolding m_def by (rule Greatest_le[OF that B])
+    hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force
+
+    have "c dvd coeff a i" for i
+    proof (induction i rule: nat_descend_induct[of "degree a"])
+      case (base i)
+      thus ?case by (simp add: coeff_eq_0)
+    next
+      case (descend i)
+      let ?A = "{..i+m} - {i}"
+      have "c dvd coeff (a * b) (i + m)" by (rule dvd)
+      also have "coeff (a * b) (i + m) = (\<Sum>k\<le>i + m. coeff a k * coeff b (i + m - k))"
+        by (simp add: coeff_mult)
+      also have "{..i+m} = insert i ?A" by auto
+      also have "(\<Sum>k\<in>\<dots>. coeff a k * coeff b (i + m - k)) =
+                   coeff a i * coeff b m + (\<Sum>k\<in>?A. coeff a k * coeff b (i + m - k))"
+        (is "_ = _ + ?S")
+        by (subst sum.insert) simp_all
+      finally have eq: "c dvd coeff a i * coeff b m + ?S" .
+      moreover have "c dvd ?S"
+      proof (rule dvd_sum)
+        fix k assume k: "k \<in> {..i+m} - {i}"
+        show "c dvd coeff a k * coeff b (i + m - k)"
+        proof (cases "k < i")
+          case False
+          with k have "c dvd coeff a k" by (intro descend.IH) simp
+          thus ?thesis by simp
+        next
+          case True
+          hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp
+          thus ?thesis by simp
+        qed
+      qed
+      ultimately have "c dvd coeff a i * coeff b m"
+        by (simp add: dvd_add_left_iff)
+      with assms coeff_m show "c dvd coeff a i"
+        by (simp add: prime_elem_dvd_mult_iff)
+    qed
+    hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast
+  }
+  thus "[:c:] dvd a \<or> [:c:] dvd b" by blast
+qed (insert assms, simp_all add: prime_elem_def one_poly_def)
+
+lemma prime_elem_const_poly_iff:
+  fixes c :: "'a :: semidom"
+  shows   "prime_elem [:c:] \<longleftrightarrow> prime_elem c"
+proof
+  assume A: "prime_elem [:c:]"
+  show "prime_elem c"
+  proof (rule prime_elemI)
+    fix a b assume "c dvd a * b"
+    hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac)
+    from A and this have "[:c:] dvd [:a:] \<or> [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD)
+    thus "c dvd a \<or> c dvd b" by simp
+  qed (insert A, auto simp: prime_elem_def is_unit_poly_iff)
+qed (auto intro: lift_prime_elem_poly)
+
+context
+begin
+
+private lemma content_1_mult:
+  fixes f g :: "'a :: {semiring_Gcd,factorial_semiring} poly"
+  assumes "content f = 1" "content g = 1"
+  shows   "content (f * g) = 1"
+proof (cases "f * g = 0")
+  case False
+  from assms have "f \<noteq> 0" "g \<noteq> 0" by auto
+
+  hence "f * g \<noteq> 0" by auto
+  {
+    assume "\<not>is_unit (content (f * g))"
+    with False have "\<exists>p. p dvd content (f * g) \<and> prime p"
+      by (intro prime_divisor_exists) simp_all
+    then obtain p where "p dvd content (f * g)" "prime p" by blast
+    from \<open>p dvd content (f * g)\<close> have "[:p:] dvd f * g"
+      by (simp add: const_poly_dvd_iff_dvd_content)
+    moreover from \<open>prime p\<close> have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly)
+    ultimately have "[:p:] dvd f \<or> [:p:] dvd g"
+      by (simp add: prime_elem_dvd_mult_iff)
+    with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content)
+    with \<open>prime p\<close> have False by simp
+  }
+  hence "is_unit (content (f * g))" by blast
+  hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content)
+  thus ?thesis by simp
+qed (insert assms, auto)
+
+lemma content_mult:
+  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
+  shows "content (p * q) = content p * content q"
+proof -
+  from content_decompose[of p] guess p' . note p = this
+  from content_decompose[of q] guess q' . note q = this
+  have "content (p * q) = content p * content q * content (p' * q')"
+    by (subst p, subst q) (simp add: mult_ac normalize_mult)
+  also from p q have "content (p' * q') = 1" by (intro content_1_mult)
+  finally show ?thesis by simp
+qed
+
+lemma primitive_part_mult:
+  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+  shows "primitive_part (p * q) = primitive_part p * primitive_part q"
+proof -
+  have "primitive_part (p * q) = p * q div [:content (p * q):]"
+    by (simp add: primitive_part_def div_const_poly_conv_map_poly)
+  also have "\<dots> = (p div [:content p:]) * (q div [:content q:])"
+    by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac)
+  also have "\<dots> = primitive_part p * primitive_part q"
+    by (simp add: primitive_part_def div_const_poly_conv_map_poly)
+  finally show ?thesis .
+qed
+
+lemma primitive_part_smult:
+  fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+  shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)"
+proof -
+  have "smult a p = [:a:] * p" by simp
+  also have "primitive_part \<dots> = smult (unit_factor a) (primitive_part p)"
+    by (subst primitive_part_mult) simp_all
+  finally show ?thesis .
+qed  
+
+lemma primitive_part_dvd_primitive_partI [intro]:
+  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+  shows "p dvd q \<Longrightarrow> primitive_part p dvd primitive_part q"
+  by (auto elim!: dvdE simp: primitive_part_mult)
+
+lemma content_prod_mset: 
+  fixes A :: "'a :: {factorial_semiring, semiring_Gcd} poly multiset"
+  shows "content (prod_mset A) = prod_mset (image_mset content A)"
+  by (induction A) (simp_all add: content_mult mult_ac)
+
+lemma fract_poly_dvdD:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
+  assumes "fract_poly p dvd fract_poly q" "content p = 1"
+  shows   "p dvd q"
+proof -
+  from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE)
+  from content_decompose_fract[of r] guess c r' . note r' = this
+  from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp  
+  from fract_poly_smult_eqE[OF this] guess a b . note ab = this
+  have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2))
+  hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4))
+  have "1 = gcd a (normalize b)" by (simp add: ab)
+  also note eq'
+  also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4))
+  finally have [simp]: "a = 1" by simp
+  from eq ab have "q = p * ([:b:] * r')" by simp
+  thus ?thesis by (rule dvdI)
+qed
+
+lemma content_prod_eq_1_iff: 
+  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
+  shows "content (p * q) = 1 \<longleftrightarrow> content p = 1 \<and> content q = 1"
+proof safe
+  assume A: "content (p * q) = 1"
+  {
+    fix p q :: "'a poly" assume "content p * content q = 1"
+    hence "1 = content p * content q" by simp
+    hence "content p dvd 1" by (rule dvdI)
+    hence "content p = 1" by simp
+  } note B = this
+  from A B[of p q] B [of q p] show "content p = 1" "content q = 1" 
+    by (simp_all add: content_mult mult_ac)
+qed (auto simp: content_mult)
+
+end
+
+
+subsection \<open>Polynomials over a field are a Euclidean ring\<close>
+
+definition unit_factor_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
+  "unit_factor_field_poly p = [:lead_coeff p:]"
+
+definition normalize_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
+  "normalize_field_poly p = smult (inverse (lead_coeff p)) p"
+
+definition euclidean_size_field_poly :: "'a :: field poly \<Rightarrow> nat" where
+  "euclidean_size_field_poly p = (if p = 0 then 0 else 2 ^ degree p)" 
+
+lemma dvd_field_poly: "dvd.dvd (op * :: 'a :: field poly \<Rightarrow> _) = op dvd"
+  by (intro ext) (simp_all add: dvd.dvd_def dvd_def)
+
+interpretation field_poly: 
+  unique_euclidean_ring where zero = "0 :: 'a :: field poly"
+    and one = 1 and plus = plus and uminus = uminus and minus = minus
+    and times = times
+    and normalize = normalize_field_poly and unit_factor = unit_factor_field_poly
+    and euclidean_size = euclidean_size_field_poly
+    and uniqueness_constraint = top
+    and divide = divide and modulo = modulo
+proof (standard, unfold dvd_field_poly)
+  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)
+next
+  fix p :: "'a poly" assume "is_unit p"
+  then show "unit_factor_field_poly p = p"
+    by (elim is_unit_polyE) (auto simp: unit_factor_field_poly_def monom_0 one_poly_def field_simps)
+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 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"
+  ultimately show "euclidean_size_field_poly (p * s) < euclidean_size_field_poly (q * s)"
+    by (auto simp add: euclidean_size_field_poly_def degree_mult_eq)
+next
+  fix p q r :: "'a poly" assume "p \<noteq> 0"
+  moreover assume "euclidean_size_field_poly r < euclidean_size_field_poly p"
+  ultimately show "(q * p + r) div p = q"
+    by (cases "r = 0")
+      (auto simp add: unit_factor_field_poly_def euclidean_size_field_poly_def div_poly_less)
+qed (auto simp: unit_factor_field_poly_def normalize_field_poly_def lead_coeff_mult 
+       euclidean_size_field_poly_def Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le)
+
+lemma field_poly_irreducible_imp_prime:
+  assumes "irreducible (p :: 'a :: field poly)"
+  shows   "prime_elem p"
+proof -
+  have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
+  from field_poly.irreducible_imp_prime_elem[of p] assms
+    show ?thesis unfolding irreducible_def prime_elem_def dvd_field_poly
+      comm_semiring_1.irreducible_def[OF A] comm_semiring_1.prime_elem_def[OF A] by blast
+qed
+
+lemma field_poly_prod_mset_prime_factorization:
+  assumes "(x :: 'a :: field poly) \<noteq> 0"
+  shows   "prod_mset (field_poly.prime_factorization x) = normalize_field_poly x"
+proof -
+  have A: "class.comm_monoid_mult op * (1 :: 'a poly)" ..
+  have "comm_monoid_mult.prod_mset op * (1 :: 'a poly) = prod_mset"
+    by (intro ext) (simp add: comm_monoid_mult.prod_mset_def[OF A] prod_mset_def)
+  with field_poly.prod_mset_prime_factorization[OF assms] show ?thesis by simp
+qed
+
+lemma field_poly_in_prime_factorization_imp_prime:
+  assumes "(p :: 'a :: field poly) \<in># field_poly.prime_factorization x"
+  shows   "prime_elem p"
+proof -
+  have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
+  have B: "class.normalization_semidom op div op + op - (0 :: 'a poly) op * 1 
+             unit_factor_field_poly normalize_field_poly" ..
+  from field_poly.in_prime_factors_imp_prime [of p x] assms
+    show ?thesis unfolding prime_elem_def dvd_field_poly
+      comm_semiring_1.prime_elem_def[OF A] normalization_semidom.prime_def[OF B] by blast
+qed
+
+
+subsection \<open>Primality and irreducibility in polynomial rings\<close>
+
+lemma nonconst_poly_irreducible_iff:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
+  assumes "degree p \<noteq> 0"
+  shows   "irreducible p \<longleftrightarrow> irreducible (fract_poly p) \<and> content p = 1"
+proof safe
+  assume p: "irreducible p"
+
+  from content_decompose[of p] guess p' . note p' = this
+  hence "p = [:content p:] * p'" by simp
+  from p this have "[:content p:] dvd 1 \<or> p' dvd 1" by (rule irreducibleD)
+  moreover have "\<not>p' dvd 1"
+  proof
+    assume "p' dvd 1"
+    hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff)
+    with assms show False by contradiction
+  qed
+  ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff)
+  
+  show "irreducible (map_poly to_fract p)"
+  proof (rule irreducibleI)
+    have "fract_poly p = 0 \<longleftrightarrow> p = 0" by (intro map_poly_eq_0_iff) auto
+    with assms show "map_poly to_fract p \<noteq> 0" by auto
+  next
+    show "\<not>is_unit (fract_poly p)"
+    proof
+      assume "is_unit (map_poly to_fract p)"
+      hence "degree (map_poly to_fract p) = 0"
+        by (auto simp: is_unit_poly_iff)
+      hence "degree p = 0" by (simp add: degree_map_poly)
+      with assms show False by contradiction
+   qed
+ next
+   fix q r assume qr: "fract_poly p = q * r"
+   from content_decompose_fract[of q] guess cg q' . note q = this
+   from content_decompose_fract[of r] guess cr r' . note r = this
+   from qr q r p have nz: "cg \<noteq> 0" "cr \<noteq> 0" by auto
+   from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))"
+     by (simp add: q r)
+   from fract_poly_smult_eqE[OF this] guess a b . note ab = this
+   hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:)
+   with ab(4) have a: "a = normalize b" by (simp add: content_mult q r)
+   hence "normalize b = gcd a b" by simp
+   also from ab(3) have "\<dots> = 1" .
+   finally have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff)
+   
+   note eq
+   also from ab(1) \<open>a = 1\<close> have "cr * cg = to_fract b" by simp
+   also have "smult \<dots> (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp
+   finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult)
+   from p and this have "([:b:] * q') dvd 1 \<or> r' dvd 1" by (rule irreducibleD)
+   hence "q' dvd 1 \<or> r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left)
+   hence "fract_poly q' dvd 1 \<or> fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit)
+   with q r show "is_unit q \<or> is_unit r"
+     by (auto simp add: is_unit_smult_iff dvd_field_iff nz)
+ qed
+
+next
+
+  assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1"
+  show "irreducible p"
+  proof (rule irreducibleI)
+    from irred show "p \<noteq> 0" by auto
+  next
+    from irred show "\<not>p dvd 1"
+      by (auto simp: irreducible_def dest: fract_poly_is_unit)
+  next
+    fix q r assume qr: "p = q * r"
+    hence "fract_poly p = fract_poly q * fract_poly r" by simp
+    from irred and this have "fract_poly q dvd 1 \<or> fract_poly r dvd 1" 
+      by (rule irreducibleD)
+    with primitive qr show "q dvd 1 \<or> r dvd 1"
+      by (auto simp:  content_prod_eq_1_iff is_unit_fract_poly_iff)
+  qed
+qed
+
+context
+begin
+
+private lemma irreducible_imp_prime_poly:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
+  assumes "irreducible p"
+  shows   "prime_elem p"
+proof (cases "degree p = 0")
+  case True
+  with assms show ?thesis
+    by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff
+             intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE)
+next
+  case False
+  from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1"
+    by (simp_all add: nonconst_poly_irreducible_iff)
+  from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime)
+  show ?thesis
+  proof (rule prime_elemI)
+    fix q r assume "p dvd q * r"
+    hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd)
+    hence "fract_poly p dvd fract_poly q * fract_poly r" by simp
+    from prime and this have "fract_poly p dvd fract_poly q \<or> fract_poly p dvd fract_poly r"
+      by (rule prime_elem_dvd_multD)
+    with primitive show "p dvd q \<or> p dvd r" by (auto dest: fract_poly_dvdD)
+  qed (insert assms, auto simp: irreducible_def)
+qed
+
+
+lemma degree_primitive_part_fract [simp]:
+  "degree (primitive_part_fract p) = degree p"
+proof -
+  have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))"
+    by (simp add: content_times_primitive_part_fract)
+  also have "degree \<dots> = degree (primitive_part_fract p)"
+    by (auto simp: degree_map_poly)
+  finally show ?thesis ..
+qed
+
+lemma irreducible_primitive_part_fract:
+  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
+  assumes "irreducible p"
+  shows   "irreducible (primitive_part_fract p)"
+proof -
+  from assms have deg: "degree (primitive_part_fract p) \<noteq> 0"
+    by (intro notI) 
+       (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff)
+  hence [simp]: "p \<noteq> 0" by auto
+
+  note \<open>irreducible p\<close>
+  also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" 
+    by (simp add: content_times_primitive_part_fract)
+  also have "irreducible \<dots> \<longleftrightarrow> irreducible (fract_poly (primitive_part_fract p))"
+    by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff)
+  finally show ?thesis using deg
+    by (simp add: nonconst_poly_irreducible_iff)
+qed
+
+lemma prime_elem_primitive_part_fract:
+  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
+  shows "irreducible p \<Longrightarrow> prime_elem (primitive_part_fract p)"
+  by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract)
+
+lemma irreducible_linear_field_poly:
+  fixes a b :: "'a::field"
+  assumes "b \<noteq> 0"
+  shows "irreducible [:a,b:]"
+proof (rule irreducibleI)
+  fix p q assume pq: "[:a,b:] = p * q"
+  also from pq assms have "degree \<dots> = degree p + degree q" 
+    by (intro degree_mult_eq) auto
+  finally have "degree p = 0 \<or> degree q = 0" using assms by auto
+  with assms pq show "is_unit p \<or> is_unit q"
+    by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE)
+qed (insert assms, auto simp: is_unit_poly_iff)
+
+lemma prime_elem_linear_field_poly:
+  "(b :: 'a :: field) \<noteq> 0 \<Longrightarrow> prime_elem [:a,b:]"
+  by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly)
+
+lemma irreducible_linear_poly:
+  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
+  shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> irreducible [:a,b:]"
+  by (auto intro!: irreducible_linear_field_poly 
+           simp:   nonconst_poly_irreducible_iff content_def map_poly_pCons)
+
+lemma prime_elem_linear_poly:
+  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
+  shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> prime_elem [:a,b:]"
+  by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly)
+
+end
+
+ 
+subsection \<open>Prime factorisation of polynomials\<close>   
+
+context
+begin 
+
+private lemma poly_prime_factorization_exists_content_1:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
+  assumes "p \<noteq> 0" "content p = 1"
+  shows   "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p"
+proof -
+  let ?P = "field_poly.prime_factorization (fract_poly p)"
+  define c where "c = prod_mset (image_mset fract_content ?P)"
+  define c' where "c' = c * to_fract (lead_coeff p)"
+  define e where "e = prod_mset (image_mset primitive_part_fract ?P)"
+  define A where "A = image_mset (normalize \<circ> primitive_part_fract) ?P"
+  have "content e = (\<Prod>x\<in>#field_poly.prime_factorization (map_poly to_fract p). 
+                      content (primitive_part_fract x))"
+    by (simp add: e_def content_prod_mset multiset.map_comp o_def)
+  also have "image_mset (\<lambda>x. content (primitive_part_fract x)) ?P = image_mset (\<lambda>_. 1) ?P"
+    by (intro image_mset_cong content_primitive_part_fract) auto
+  finally have content_e: "content e = 1"
+    by simp    
+  
+  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 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
+  also have "image_mset id ?P = 
+               image_mset (\<lambda>x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P"
+    by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract)
+  also have "prod_mset \<dots> = smult c (fract_poly e)"
+    by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def)
+  also have "[:to_fract (lead_coeff p):] * \<dots> = smult c' (fract_poly e)"
+    by (simp add: c'_def)
+  finally have eq: "fract_poly p = smult c' (fract_poly e)" .
+  also obtain b where b: "c' = to_fract b" "is_unit b"
+  proof -
+    from fract_poly_smult_eqE[OF eq] guess a b . note ab = this
+    from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: )
+    with assms content_e have "a = normalize b" by (simp add: ab(4))
+    with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff)
+    with ab ab' have "c' = to_fract b" by auto
+    from this and \<open>is_unit b\<close> show ?thesis by (rule that)
+  qed
+  hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp
+  finally have "p = smult b e" by (simp only: fract_poly_eq_iff)
+  hence "p = [:b:] * e" by simp
+  with b have "normalize p = normalize e" 
+    by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff)
+  also have "normalize e = prod_mset A"
+    by (simp add: multiset.map_comp e_def A_def normalize_prod_mset)
+  finally have "prod_mset A = normalize p" ..
+  
+  have "prime_elem p" if "p \<in># A" for p
+    using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible 
+                        dest!: field_poly_in_prime_factorization_imp_prime )
+  from this and \<open>prod_mset A = normalize p\<close> show ?thesis
+    by (intro exI[of _ A]) blast
+qed
+
+lemma poly_prime_factorization_exists:
+  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
+  assumes "p \<noteq> 0"
+  shows   "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p"
+proof -
+  define B where "B = image_mset (\<lambda>x. [:x:]) (prime_factorization (content p))"
+  have "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize (primitive_part p)"
+    by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all)
+  then guess A by (elim exE conjE) note A = this
+  moreover from assms have "prod_mset B = [:content p:]"
+    by (simp add: B_def prod_mset_const_poly prod_mset_prime_factorization)
+  moreover have "\<forall>p. p \<in># B \<longrightarrow> prime_elem p"
+    by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime)
+  ultimately show ?thesis by (intro exI[of _ "B + A"]) auto
+qed
+
+end
+
+
+subsection \<open>Typeclass instances\<close>
+
+instance poly :: (factorial_ring_gcd) factorial_semiring
+  by standard (rule poly_prime_factorization_exists)  
+
+instantiation poly :: (factorial_ring_gcd) factorial_ring_gcd
+begin
+
+definition gcd_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  [code del]: "gcd_poly = gcd_factorial"
+
+definition lcm_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  [code del]: "lcm_poly = lcm_factorial"
+  
+definition Gcd_poly :: "'a poly set \<Rightarrow> 'a poly" where
+ [code del]: "Gcd_poly = Gcd_factorial"
+
+definition Lcm_poly :: "'a poly set \<Rightarrow> 'a poly" where
+ [code del]: "Lcm_poly = Lcm_factorial"
+ 
+instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def)
+
+end
+
+instantiation poly :: ("{field,factorial_ring_gcd}") unique_euclidean_ring
+begin
+
+definition euclidean_size_poly :: "'a poly \<Rightarrow> nat"
+  where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)"
+
+definition uniqueness_constraint_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
+  where [simp]: "uniqueness_constraint_poly = top"
+
+instance 
+  by standard
+   (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq intro!: degree_mod_less' degree_mult_right_le
+    split: if_splits)
+
+end
+
+instance poly :: ("{field,factorial_ring_gcd}") euclidean_ring_gcd
+  by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI)
+    standard
+
+  
+subsection \<open>Polynomial GCD\<close>
+
+lemma gcd_poly_decompose:
+  fixes p q :: "'a :: factorial_ring_gcd poly"
+  shows "gcd p q = 
+           smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"
+proof (rule sym, rule gcdI)
+  have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd
+          [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all
+  thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p"
+    by simp
+next
+  have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd
+          [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all
+  thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q"
+    by simp
+next
+  fix d assume "d dvd p" "d dvd q"
+  hence "[:content d:] * primitive_part d dvd 
+           [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)"
+    by (intro mult_dvd_mono) auto
+  thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"
+    by simp
+qed (auto simp: normalize_smult)
+  
+
+lemma gcd_poly_pseudo_mod:
+  fixes p q :: "'a :: factorial_ring_gcd poly"
+  assumes nz: "q \<noteq> 0" and prim: "content p = 1" "content q = 1"
+  shows   "gcd p q = gcd q (primitive_part (pseudo_mod p q))"
+proof -
+  define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)"
+  define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]"
+  have [simp]: "primitive_part a = unit_factor a"
+    by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0)
+  from nz have [simp]: "a \<noteq> 0" by (auto simp: a_def)
+  
+  have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def)
+  have "gcd (q * r + s) q = gcd q s"
+    using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac)
+  with pseudo_divmod(1)[OF nz rs]
+    have "gcd (p * a) q = gcd q s" by (simp add: a_def)
+  also from prim have "gcd (p * a) q = gcd p q"
+    by (subst gcd_poly_decompose)
+       (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim 
+             simp del: mult_pCons_right )
+  also from prim have "gcd q s = gcd q (primitive_part s)"
+    by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim)
+  also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def)
+  finally show ?thesis .
+qed
+
+lemma degree_pseudo_mod_less:
+  assumes "q \<noteq> 0" "pseudo_mod p q \<noteq> 0"
+  shows   "degree (pseudo_mod p q) < degree q"
+  using pseudo_mod(2)[of q p] assms by auto
+
+function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  "gcd_poly_code_aux p q = 
+     (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" 
+by auto
+termination
+  by (relation "measure ((\<lambda>p. if p = 0 then 0 else Suc (degree p)) \<circ> snd)")
+     (auto simp: degree_pseudo_mod_less)
+
+declare gcd_poly_code_aux.simps [simp del]
+
+lemma gcd_poly_code_aux_correct:
+  assumes "content p = 1" "q = 0 \<or> content q = 1"
+  shows   "gcd_poly_code_aux p q = gcd p q"
+  using assms
+proof (induction p q rule: gcd_poly_code_aux.induct)
+  case (1 p q)
+  show ?case
+  proof (cases "q = 0")
+    case True
+    thus ?thesis by (subst gcd_poly_code_aux.simps) auto
+  next
+    case False
+    hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))"
+      by (subst gcd_poly_code_aux.simps) simp_all
+    also from "1.prems" False 
+      have "primitive_part (pseudo_mod p q) = 0 \<or> 
+              content (primitive_part (pseudo_mod p q)) = 1"
+      by (cases "pseudo_mod p q = 0") auto
+    with "1.prems" False 
+      have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = 
+              gcd q (primitive_part (pseudo_mod p q))"
+      by (intro 1) simp_all
+    also from "1.prems" False 
+      have "\<dots> = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto
+    finally show ?thesis .
+  qed
+qed
+
+definition gcd_poly_code 
+    :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" 
+  where "gcd_poly_code p q = 
+           (if p = 0 then normalize q else if q = 0 then normalize p else
+              smult (gcd (content p) (content q)) 
+                (gcd_poly_code_aux (primitive_part p) (primitive_part q)))"
+
+lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q"
+  by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric])
+
+lemma lcm_poly_code [code]: 
+  fixes p q :: "'a :: factorial_ring_gcd poly"
+  shows "lcm p q = normalize (p * q) div gcd p q"
+  by (fact lcm_gcd)
+
+lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"]
+lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"]
+
+text \<open>Example:
+  @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval}
+\<close>
+  
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Computational_Algebra/Primes.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -0,0 +1,725 @@
+(*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+                Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow, 
+                Manuel Eberl
+
+
+This file deals with properties of primes. Definitions and lemmas are
+proved uniformly for the natural numbers and integers.
+
+This file combines and revises a number of prior developments.
+
+The original theories "GCD" and "Primes" were by Christophe Tabacznyj
+and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
+gcd, lcm, and prime for the natural numbers.
+
+The original theory "IntPrimes" was by Thomas M. Rasmussen, and
+extended gcd, lcm, primes to the integers. Amine Chaieb provided
+another extension of the notions to the integers, and added a number
+of results to "Primes" and "GCD". IntPrimes also defined and developed
+the congruence relations on the integers. The notion was extended to
+the natural numbers by Chaieb.
+
+Jeremy Avigad combined all of these, made everything uniform for the
+natural numbers and the integers, and added a number of new theorems.
+
+Tobias Nipkow cleaned up a lot.
+
+Florian Haftmann and Manuel Eberl put primality and prime factorisation
+onto an algebraic foundation and thus generalised these concepts to 
+other rings, such as polynomials. (see also the Factorial_Ring theory).
+
+There were also previous formalisations of unique factorisation by 
+Thomas Marthedal Rasmussen, Jeremy Avigad, and David Gray.
+*)
+
+
+section \<open>Primes\<close>
+
+theory Primes
+imports "~~/src/HOL/Binomial" Euclidean_Algorithm
+begin
+
+(* As a simp or intro rule,
+
+     prime p \<Longrightarrow> p > 0
+
+   wreaks havoc here. When the premise includes \<forall>x \<in># M. prime x, it
+   leads to the backchaining
+
+     x > 0
+     prime x
+     x \<in># M   which is, unfortunately,
+     count M x > 0
+*)
+
+declare [[coercion int]]
+declare [[coercion_enabled]]
+
+lemma prime_elem_nat_iff:
+  "prime_elem (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
+proof safe
+  assume *: "prime_elem n"
+  from * have "n \<noteq> 0" "n \<noteq> 1" by (intro notI, simp del: One_nat_def)+
+  thus "n > 1" by (cases n) simp_all
+  fix m assume m: "m dvd n" "m \<noteq> n"
+  from * \<open>m dvd n\<close> have "n dvd m \<or> is_unit m"
+    by (intro irreducibleD' prime_elem_imp_irreducible)
+  with m show "m = 1" by (auto dest: dvd_antisym)
+next
+  assume "n > 1" "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n"
+  thus "prime_elem n"
+    by (auto simp: prime_elem_iff_irreducible irreducible_altdef)
+qed
+
+lemma prime_nat_iff_prime_elem: "prime (n :: nat) \<longleftrightarrow> prime_elem n"
+  by (simp add: prime_def)
+
+lemma prime_nat_iff:
+  "prime (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
+  by (simp add: prime_nat_iff_prime_elem prime_elem_nat_iff)
+
+lemma prime_elem_int_nat_transfer: "prime_elem (n::int) \<longleftrightarrow> prime_elem (nat (abs n))"
+proof
+  assume "prime_elem n"
+  thus "prime_elem (nat (abs n))" by (auto simp: prime_elem_def nat_dvd_iff)
+next
+  assume "prime_elem (nat (abs n))"
+  thus "prime_elem n"
+    by (auto simp: dvd_int_unfold_dvd_nat prime_elem_def abs_mult nat_mult_distrib)
+qed
+
+lemma prime_elem_nat_int_transfer [simp]: "prime_elem (int n) \<longleftrightarrow> prime_elem n"
+  by (auto simp: prime_elem_int_nat_transfer)
+
+lemma prime_nat_int_transfer [simp]: "prime (int n) \<longleftrightarrow> prime n"
+  by (auto simp: prime_elem_int_nat_transfer prime_def)
+
+lemma prime_int_nat_transfer: "prime (n::int) \<longleftrightarrow> n \<ge> 0 \<and> prime (nat n)"
+  by (auto simp: prime_elem_int_nat_transfer prime_def)
+
+lemma prime_int_iff:
+  "prime (n::int) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n))"
+proof (intro iffI conjI allI impI; (elim conjE)?)
+  assume *: "prime n"
+  hence irred: "irreducible n" by (simp add: prime_elem_imp_irreducible)
+  from * have "n \<ge> 0" "n \<noteq> 0" "n \<noteq> 1" 
+    by (auto simp: prime_def zabs_def not_less split: if_splits)
+  thus "n > 1" by presburger
+  fix m assume "m dvd n" \<open>m \<ge> 0\<close>
+  with irred have "m dvd 1 \<or> n dvd m" by (auto simp: irreducible_altdef)
+  with \<open>m dvd n\<close> \<open>m \<ge> 0\<close> \<open>n > 1\<close> show "m = 1 \<or> m = n"
+    using associated_iff_dvd[of m n] by auto
+next
+  assume n: "1 < n" "\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n"
+  hence "nat n > 1" by simp
+  moreover have "\<forall>m. m dvd nat n \<longrightarrow> m = 1 \<or> m = nat n"
+  proof (intro allI impI)
+    fix m assume "m dvd nat n"
+    with \<open>n > 1\<close> have "int m dvd n" by (auto simp: int_dvd_iff)
+    with n(2) have "int m = 1 \<or> int m = n" by auto
+    thus "m = 1 \<or> m = nat n" by auto
+  qed
+  ultimately show "prime n" 
+    unfolding prime_int_nat_transfer prime_nat_iff by auto
+qed
+
+
+lemma prime_nat_not_dvd:
+  assumes "prime p" "p > n" "n \<noteq> (1::nat)"
+  shows   "\<not>n dvd p"
+proof
+  assume "n dvd p"
+  from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
+  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
+    by (cases "n = 0") (auto dest!: dvd_imp_le)
+qed
+
+lemma prime_int_not_dvd:
+  assumes "prime p" "p > n" "n > (1::int)"
+  shows   "\<not>n dvd p"
+proof
+  assume "n dvd p"
+  from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
+  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
+    by (auto dest!: zdvd_imp_le)
+qed
+
+lemma prime_odd_nat: "prime p \<Longrightarrow> p > (2::nat) \<Longrightarrow> odd p"
+  by (intro prime_nat_not_dvd) auto
+
+lemma prime_odd_int: "prime p \<Longrightarrow> p > (2::int) \<Longrightarrow> odd p"
+  by (intro prime_int_not_dvd) auto
+
+lemma prime_ge_0_int: "prime p \<Longrightarrow> p \<ge> (0::int)"
+  unfolding prime_int_iff by auto
+
+lemma prime_gt_0_nat: "prime p \<Longrightarrow> p > (0::nat)"
+  unfolding prime_nat_iff by auto
+
+lemma prime_gt_0_int: "prime p \<Longrightarrow> p > (0::int)"
+  unfolding prime_int_iff by auto
+
+lemma prime_ge_1_nat: "prime p \<Longrightarrow> p \<ge> (1::nat)"
+  unfolding prime_nat_iff by auto
+
+lemma prime_ge_Suc_0_nat: "prime p \<Longrightarrow> p \<ge> Suc 0"
+  unfolding prime_nat_iff by auto
+
+lemma prime_ge_1_int: "prime p \<Longrightarrow> p \<ge> (1::int)"
+  unfolding prime_int_iff by auto
+
+lemma prime_gt_1_nat: "prime p \<Longrightarrow> p > (1::nat)"
+  unfolding prime_nat_iff by auto
+
+lemma prime_gt_Suc_0_nat: "prime p \<Longrightarrow> p > Suc 0"
+  unfolding prime_nat_iff by auto
+
+lemma prime_gt_1_int: "prime p \<Longrightarrow> p > (1::int)"
+  unfolding prime_int_iff by auto
+
+lemma prime_ge_2_nat: "prime p \<Longrightarrow> p \<ge> (2::nat)"
+  unfolding prime_nat_iff by auto
+
+lemma prime_ge_2_int: "prime p \<Longrightarrow> p \<ge> (2::int)"
+  unfolding prime_int_iff by auto
+
+lemma prime_int_altdef:
+  "prime p = (1 < p \<and> (\<forall>m::int. m \<ge> 0 \<longrightarrow> m dvd p \<longrightarrow>
+    m = 1 \<or> m = p))"
+  unfolding prime_int_iff by blast
+
+lemma not_prime_eq_prod_nat:
+  assumes "m > 1" "\<not>prime (m::nat)"
+  shows   "\<exists>n k. n = m * k \<and> 1 < m \<and> m < n \<and> 1 < k \<and> k < n"
+  using assms irreducible_altdef[of m]
+  by (auto simp: prime_elem_iff_irreducible prime_def irreducible_altdef)
+
+    
+subsection\<open>Largest exponent of a prime factor\<close>
+text\<open>Possibly duplicates other material, but avoid the complexities of multisets.\<close>
+  
+lemma prime_power_cancel_less:
+  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and less: "k < k'" and "\<not> p dvd m"
+  shows False
+proof -
+  obtain l where l: "k' = k + l" and "l > 0"
+    using less less_imp_add_positive by auto
+  have "m = m * (p ^ k) div (p ^ k)"
+    using \<open>prime p\<close> by simp
+  also have "\<dots> = m' * (p ^ k') div (p ^ k)"
+    using eq by simp
+  also have "\<dots> = m' * (p ^ l) * (p ^ k) div (p ^ k)"
+    by (simp add: l mult.commute mult.left_commute power_add)
+  also have "... = m' * (p ^ l)"
+    using \<open>prime p\<close> by simp
+  finally have "p dvd m"
+    using \<open>l > 0\<close> by simp
+  with assms show False
+    by simp
+qed
+
+lemma prime_power_cancel:
+  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and "\<not> p dvd m" "\<not> p dvd m'"
+  shows "k = k'"
+  using prime_power_cancel_less [OF \<open>prime p\<close>] assms
+  by (metis linorder_neqE_nat)
+
+lemma prime_power_cancel2:
+  assumes "prime p" "m * (p ^ k) = m' * (p ^ k')" "\<not> p dvd m" "\<not> p dvd m'"
+  obtains "m = m'" "k = k'"
+  using prime_power_cancel [OF assms] assms by auto
+
+lemma prime_power_canonical:
+  fixes m::nat
+  assumes "prime p" "m > 0"
+  shows "\<exists>k n. \<not> p dvd n \<and> m = n * p^k"
+using \<open>m > 0\<close>
+proof (induction m rule: less_induct)
+  case (less m)
+  show ?case
+  proof (cases "p dvd m")
+    case True
+    then obtain m' where m': "m = p * m'"
+      using dvdE by blast
+    with \<open>prime p\<close> have "0 < m'" "m' < m"
+      using less.prems prime_nat_iff by auto
+    with m' less show ?thesis
+      by (metis power_Suc mult.left_commute)
+  next
+    case False
+    then show ?thesis
+      by (metis mult.right_neutral power_0)
+  qed
+qed
+
+
+subsubsection \<open>Make prime naively executable\<close>
+
+lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
+  unfolding One_nat_def [symmetric] by (rule not_prime_1)
+
+lemma prime_nat_iff':
+  "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)"
+proof safe
+  assume "p > 1" and *: "\<forall>n\<in>{2..<p}. \<not>n dvd p"
+  show "prime p" unfolding prime_nat_iff
+  proof (intro conjI allI impI)
+    fix m assume "m dvd p"
+    with \<open>p > 1\<close> have "m \<noteq> 0" by (intro notI) auto
+    hence "m \<ge> 1" by simp
+    moreover from \<open>m dvd p\<close> and * have "m \<notin> {2..<p}" by blast
+    with \<open>m dvd p\<close> and \<open>p > 1\<close> have "m \<le> 1 \<or> m = p" by (auto dest: dvd_imp_le)
+    ultimately show "m = 1 \<or> m = p" by simp
+  qed fact+
+qed (auto simp: prime_nat_iff)
+
+lemma prime_int_iff':
+  "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)" (is "?lhs = ?rhs")
+proof
+  assume "?lhs"
+  thus "?rhs"
+      by (auto simp: prime_int_nat_transfer dvd_int_unfold_dvd_nat prime_nat_iff')
+next
+  assume "?rhs"
+  thus "?lhs"
+    by (auto simp: prime_int_nat_transfer zdvd_int prime_nat_iff')
+qed
+
+lemma prime_int_numeral_eq [simp]:
+  "prime (numeral m :: int) \<longleftrightarrow> prime (numeral m :: nat)"
+  by (simp add: prime_int_nat_transfer)
+
+lemma two_is_prime_nat [simp]: "prime (2::nat)"
+  by (simp add: prime_nat_iff')
+
+lemma prime_nat_numeral_eq [simp]:
+  "prime (numeral m :: nat) \<longleftrightarrow>
+    (1::nat) < numeral m \<and>
+    (\<forall>n::nat \<in> set [2..<numeral m]. \<not> n dvd numeral m)"
+  by (simp only: prime_nat_iff' set_upt)  \<comment> \<open>TODO Sieve Of Erathosthenes might speed this up\<close>
+
+
+text\<open>A bit of regression testing:\<close>
+
+lemma "prime(97::nat)" by simp
+lemma "prime(97::int)" by simp
+
+lemma prime_factor_nat: 
+  "n \<noteq> (1::nat) \<Longrightarrow> \<exists>p. prime p \<and> p dvd n"
+  using prime_divisor_exists[of n]
+  by (cases "n = 0") (auto intro: exI[of _ "2::nat"])
+
+
+subsection \<open>Infinitely many primes\<close>
+
+lemma next_prime_bound: "\<exists>p::nat. prime p \<and> n < p \<and> p \<le> fact n + 1"
+proof-
+  have f1: "fact n + 1 \<noteq> (1::nat)" using fact_ge_1 [of n, where 'a=nat] by arith
+  from prime_factor_nat [OF f1]
+  obtain p :: nat where "prime p" and "p dvd fact n + 1" by auto
+  then have "p \<le> fact n + 1" apply (intro dvd_imp_le) apply auto done
+  { assume "p \<le> n"
+    from \<open>prime p\<close> have "p \<ge> 1"
+      by (cases p, simp_all)
+    with \<open>p <= n\<close> have "p dvd fact n"
+      by (intro dvd_fact)
+    with \<open>p dvd fact n + 1\<close> have "p dvd fact n + 1 - fact n"
+      by (rule dvd_diff_nat)
+    then have "p dvd 1" by simp
+    then have "p <= 1" by auto
+    moreover from \<open>prime p\<close> have "p > 1"
+      using prime_nat_iff by blast
+    ultimately have False by auto}
+  then have "n < p" by presburger
+  with \<open>prime p\<close> and \<open>p <= fact n + 1\<close> show ?thesis by auto
+qed
+
+lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
+  using next_prime_bound by auto
+
+lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
+proof
+  assume "finite {(p::nat). prime p}"
+  with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
+    by auto
+  then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
+    by auto
+  with bigger_prime [of b] show False
+    by auto
+qed
+
+subsection\<open>Powers of Primes\<close>
+
+text\<open>Versions for type nat only\<close>
+
+lemma prime_product:
+  fixes p::nat
+  assumes "prime (p * q)"
+  shows "p = 1 \<or> q = 1"
+proof -
+  from assms have
+    "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
+    unfolding prime_nat_iff by auto
+  from \<open>1 < p * q\<close> have "p \<noteq> 0" by (cases p) auto
+  then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
+  have "p dvd p * q" by simp
+  then have "p = 1 \<or> p = p * q" by (rule P)
+  then show ?thesis by (simp add: Q)
+qed
+
+(* TODO: Generalise? *)
+lemma prime_power_mult_nat:
+  fixes p::nat
+  assumes p: "prime p" and xy: "x * y = p ^ k"
+  shows "\<exists>i j. x = p ^i \<and> y = p^ j"
+using xy
+proof(induct k arbitrary: x y)
+  case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
+next
+  case (Suc k x y)
+  from Suc.prems have pxy: "p dvd x*y" by auto
+  from prime_dvd_multD [OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
+  from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)
+  {assume px: "p dvd x"
+    then obtain d where d: "x = p*d" unfolding dvd_def by blast
+    from Suc.prems d  have "p*d*y = p^Suc k" by simp
+    hence th: "d*y = p^k" using p0 by simp
+    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
+    with d have "x = p^Suc i" by simp
+    with ij(2) have ?case by blast}
+  moreover
+  {assume px: "p dvd y"
+    then obtain d where d: "y = p*d" unfolding dvd_def by blast
+    from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult.commute)
+    hence th: "d*x = p^k" using p0 by simp
+    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
+    with d have "y = p^Suc i" by simp
+    with ij(2) have ?case by blast}
+  ultimately show ?case  using pxyc by blast
+qed
+
+lemma prime_power_exp_nat:
+  fixes p::nat
+  assumes p: "prime p" and n: "n \<noteq> 0"
+    and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
+  using n xn
+proof(induct n arbitrary: k)
+  case 0 thus ?case by simp
+next
+  case (Suc n k) hence th: "x*x^n = p^k" by simp
+  {assume "n = 0" with Suc have ?case by simp (rule exI[where x="k"], simp)}
+  moreover
+  {assume n: "n \<noteq> 0"
+    from prime_power_mult_nat[OF p th]
+    obtain i j where ij: "x = p^i" "x^n = p^j"by blast
+    from Suc.hyps[OF n ij(2)] have ?case .}
+  ultimately show ?case by blast
+qed
+
+lemma divides_primepow_nat:
+  fixes p::nat
+  assumes p: "prime p"
+  shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
+proof
+  assume H: "d dvd p^k" then obtain e where e: "d*e = p^k"
+    unfolding dvd_def  apply (auto simp add: mult.commute) by blast
+  from prime_power_mult_nat[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
+  from e ij have "p^(i + j) = p^k" by (simp add: power_add)
+  hence "i + j = k" using p prime_gt_1_nat power_inject_exp[of p "i+j" k] by simp
+  hence "i \<le> k" by arith
+  with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
+next
+  {fix i assume H: "i \<le> k" "d = p^i"
+    then obtain j where j: "k = i + j"
+      by (metis le_add_diff_inverse)
+    hence "p^k = p^j*d" using H(2) by (simp add: power_add)
+    hence "d dvd p^k" unfolding dvd_def by auto}
+  thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
+qed
+
+
+subsection \<open>Chinese Remainder Theorem Variants\<close>
+
+lemma bezout_gcd_nat:
+  fixes a::nat shows "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
+  using bezout_nat[of a b]
+by (metis bezout_nat diff_add_inverse gcd_add_mult gcd.commute
+  gcd_nat.right_neutral mult_0)
+
+lemma gcd_bezout_sum_nat:
+  fixes a::nat
+  assumes "a * x + b * y = d"
+  shows "gcd a b dvd d"
+proof-
+  let ?g = "gcd a b"
+    have dv: "?g dvd a*x" "?g dvd b * y"
+      by simp_all
+    from dvd_add[OF dv] assms
+    show ?thesis by auto
+qed
+
+
+text \<open>A binary form of the Chinese Remainder Theorem.\<close>
+
+(* TODO: Generalise? *)
+lemma chinese_remainder:
+  fixes a::nat  assumes ab: "coprime a b" and a: "a \<noteq> 0" and b: "b \<noteq> 0"
+  shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
+proof-
+  from bezout_add_strong_nat[OF a, of b] bezout_add_strong_nat[OF b, of a]
+  obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"
+    and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
+  then have d12: "d1 = 1" "d2 =1"
+    by (metis ab coprime_nat)+
+  let ?x = "v * a * x1 + u * b * x2"
+  let ?q1 = "v * x1 + u * y2"
+  let ?q2 = "v * y1 + u * x2"
+  from dxy2(3)[simplified d12] dxy1(3)[simplified d12]
+  have "?x = u + ?q1 * a" "?x = v + ?q2 * b"
+    by algebra+
+  thus ?thesis by blast
+qed
+
+text \<open>Primality\<close>
+
+lemma coprime_bezout_strong:
+  fixes a::nat assumes "coprime a b"  "b \<noteq> 1"
+  shows "\<exists>x y. a * x = b * y + 1"
+by (metis assms bezout_nat gcd_nat.left_neutral)
+
+lemma bezout_prime:
+  assumes p: "prime p" and pa: "\<not> p dvd a"
+  shows "\<exists>x y. a*x = Suc (p*y)"
+proof -
+  have ap: "coprime a p"
+    by (metis gcd.commute p pa prime_imp_coprime)
+  moreover from p have "p \<noteq> 1" by auto
+  ultimately have "\<exists>x y. a * x = p * y + 1"
+    by (rule coprime_bezout_strong)
+  then show ?thesis by simp    
+qed
+(* END TODO *)
+
+
+
+subsection \<open>Multiplicity and primality for natural numbers and integers\<close>
+
+lemma prime_factors_gt_0_nat:
+  "p \<in> prime_factors x \<Longrightarrow> p > (0::nat)"
+  by (simp add: in_prime_factors_imp_prime prime_gt_0_nat)
+
+lemma prime_factors_gt_0_int:
+  "p \<in> prime_factors x \<Longrightarrow> p > (0::int)"
+  by (simp add: in_prime_factors_imp_prime prime_gt_0_int)
+
+lemma prime_factors_ge_0_int [elim]: (* FIXME !? *)
+  fixes n :: int
+  shows "p \<in> prime_factors n \<Longrightarrow> p \<ge> 0"
+  by (drule prime_factors_gt_0_int) simp
+  
+lemma prod_mset_prime_factorization_int:
+  fixes n :: int
+  assumes "n > 0"
+  shows   "prod_mset (prime_factorization n) = n"
+  using assms by (simp add: prod_mset_prime_factorization)
+
+lemma prime_factorization_exists_nat:
+  "n > 0 \<Longrightarrow> (\<exists>M. (\<forall>p::nat \<in> set_mset M. prime p) \<and> n = (\<Prod>i \<in># M. i))"
+  using prime_factorization_exists[of n] by (auto simp: prime_def)
+
+lemma prod_mset_prime_factorization_nat [simp]: 
+  "(n::nat) > 0 \<Longrightarrow> prod_mset (prime_factorization n) = n"
+  by (subst prod_mset_prime_factorization) simp_all
+
+lemma prime_factorization_nat:
+    "n > (0::nat) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
+  by (simp add: prod_prime_factors)
+
+lemma prime_factorization_int:
+    "n > (0::int) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
+  by (simp add: prod_prime_factors)
+
+lemma prime_factorization_unique_nat:
+  fixes f :: "nat \<Rightarrow> _"
+  assumes S_eq: "S = {p. 0 < f p}"
+    and "finite S"
+    and S: "\<forall>p\<in>S. prime p" "n = (\<Prod>p\<in>S. p ^ f p)"
+  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
+  using assms by (intro prime_factorization_unique'') auto
+
+lemma prime_factorization_unique_int:
+  fixes f :: "int \<Rightarrow> _"
+  assumes S_eq: "S = {p. 0 < f p}"
+    and "finite S"
+    and S: "\<forall>p\<in>S. prime p" "abs n = (\<Prod>p\<in>S. p ^ f p)"
+  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
+  using assms by (intro prime_factorization_unique'') auto
+
+lemma prime_factors_characterization_nat:
+  "S = {p. 0 < f (p::nat)} \<Longrightarrow>
+    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
+  by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric])
+
+lemma prime_factors_characterization'_nat:
+  "finite {p. 0 < f (p::nat)} \<Longrightarrow>
+    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
+      prime_factors (\<Prod>p | 0 < f p. p ^ f p) = {p. 0 < f p}"
+  by (rule prime_factors_characterization_nat) auto
+
+lemma prime_factors_characterization_int:
+  "S = {p. 0 < f (p::int)} \<Longrightarrow> finite S \<Longrightarrow>
+    \<forall>p\<in>S. prime p \<Longrightarrow> abs n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
+  by (rule prime_factorization_unique_int [THEN conjunct1, symmetric])
+
+(* TODO Move *)
+lemma abs_prod: "abs (prod f A :: 'a :: linordered_idom) = prod (\<lambda>x. abs (f x)) A"
+  by (cases "finite A", induction A rule: finite_induct) (simp_all add: abs_mult)
+
+lemma primes_characterization'_int [rule_format]:
+  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow> \<forall>p. 0 < f p \<longrightarrow> prime p \<Longrightarrow>
+      prime_factors (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = {p. p \<ge> 0 \<and> 0 < f p}"
+  by (rule prime_factors_characterization_int) (auto simp: abs_prod prime_ge_0_int)
+
+lemma multiplicity_characterization_nat:
+  "S = {p. 0 < f (p::nat)} \<Longrightarrow> finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow>
+    n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
+  by (frule prime_factorization_unique_nat [of S f n, THEN conjunct2, rule_format, symmetric]) auto
+
+lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
+    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<longrightarrow> prime p \<longrightarrow>
+      multiplicity p (\<Prod>p | 0 < f p. p ^ f p) = f p"
+  by (intro impI, rule multiplicity_characterization_nat) auto
+
+lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
+    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
+  by (frule prime_factorization_unique_int [of S f n, THEN conjunct2, rule_format, symmetric]) 
+     (auto simp: abs_prod power_abs prime_ge_0_int intro!: prod.cong)
+
+lemma multiplicity_characterization'_int [rule_format]:
+  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow>
+    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> prime p \<Longrightarrow>
+      multiplicity p (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = f p"
+  by (rule multiplicity_characterization_int) (auto simp: prime_ge_0_int)
+
+lemma multiplicity_one_nat [simp]: "multiplicity p (Suc 0) = 0"
+  unfolding One_nat_def [symmetric] by (rule multiplicity_one)
+
+lemma multiplicity_eq_nat:
+  fixes x and y::nat
+  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+  shows "x = y"
+  using multiplicity_eq_imp_eq[of x y] assms by simp
+
+lemma multiplicity_eq_int:
+  fixes x y :: int
+  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+  shows "x = y"
+  using multiplicity_eq_imp_eq[of x y] assms by simp
+
+lemma multiplicity_prod_prime_powers:
+  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> prime x" "prime p"
+  shows   "multiplicity p (\<Prod>p \<in> S. p ^ f p) = (if p \<in> S then f p else 0)"
+proof -
+  define g where "g = (\<lambda>x. if x \<in> S then f x else 0)"
+  define A where "A = Abs_multiset g"
+  have "{x. g x > 0} \<subseteq> S" by (auto simp: g_def)
+  from finite_subset[OF this assms(1)] have [simp]: "g :  multiset"
+    by (simp add: multiset_def)
+  from assms have count_A: "count A x = g x" for x unfolding A_def
+    by simp
+  have set_mset_A: "set_mset A = {x\<in>S. f x > 0}"
+    unfolding set_mset_def count_A by (auto simp: g_def)
+  with assms have prime: "prime x" if "x \<in># A" for x using that by auto
+  from set_mset_A assms have "(\<Prod>p \<in> S. p ^ f p) = (\<Prod>p \<in> S. p ^ g p) "
+    by (intro prod.cong) (auto simp: g_def)
+  also from set_mset_A assms have "\<dots> = (\<Prod>p \<in> set_mset A. p ^ g p)"
+    by (intro prod.mono_neutral_right) (auto simp: g_def set_mset_A)
+  also have "\<dots> = prod_mset A"
+    by (auto simp: prod_mset_multiplicity count_A set_mset_A intro!: prod.cong)
+  also from assms have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
+    by (subst prime_elem_multiplicity_prod_mset_distrib) (auto dest: prime)
+  also from assms have "image_mset (multiplicity p) A = image_mset (\<lambda>x. if x = p then 1 else 0) A"
+    by (intro image_mset_cong) (auto simp: prime_multiplicity_other dest: prime)
+  also have "sum_mset \<dots> = (if p \<in> S then f p else 0)" by (simp add: sum_mset_delta count_A g_def)
+  finally show ?thesis .
+qed
+
+lemma prime_factorization_prod_mset:
+  assumes "0 \<notin># A"
+  shows "prime_factorization (prod_mset A) = \<Union>#(image_mset prime_factorization A)"
+  using assms by (induct A) (auto simp add: prime_factorization_mult)
+
+lemma prime_factors_prod:
+  assumes "finite A" and "0 \<notin> f ` A"
+  shows "prime_factors (prod f A) = UNION A (prime_factors \<circ> f)"
+  using assms by (simp add: prod_unfold_prod_mset prime_factorization_prod_mset)
+
+lemma prime_factors_fact:
+  "prime_factors (fact n) = {p \<in> {2..n}. prime p}" (is "?M = ?N")
+proof (rule set_eqI)
+  fix p
+  { fix m :: nat
+    assume "p \<in> prime_factors m"
+    then have "prime p" and "p dvd m" by auto
+    moreover assume "m > 0" 
+    ultimately have "2 \<le> p" and "p \<le> m"
+      by (auto intro: prime_ge_2_nat dest: dvd_imp_le)
+    moreover assume "m \<le> n"
+    ultimately have "2 \<le> p" and "p \<le> n"
+      by (auto intro: order_trans)
+  } note * = this
+  show "p \<in> ?M \<longleftrightarrow> p \<in> ?N"
+    by (auto simp add: fact_prod prime_factors_prod Suc_le_eq dest!: prime_prime_factors intro: *)
+qed
+
+lemma prime_dvd_fact_iff:
+  assumes "prime p"
+  shows "p dvd fact n \<longleftrightarrow> p \<le> n"
+  using assms
+  by (auto simp add: prime_factorization_subset_iff_dvd [symmetric]
+    prime_factorization_prime prime_factors_fact prime_ge_2_nat)
+
+(* TODO Legacy names *)
+lemmas prime_imp_coprime_nat = prime_imp_coprime[where ?'a = nat]
+lemmas prime_imp_coprime_int = prime_imp_coprime[where ?'a = int]
+lemmas prime_dvd_mult_nat = prime_dvd_mult_iff[where ?'a = nat]
+lemmas prime_dvd_mult_int = prime_dvd_mult_iff[where ?'a = int]
+lemmas prime_dvd_mult_eq_nat = prime_dvd_mult_iff[where ?'a = nat]
+lemmas prime_dvd_mult_eq_int = prime_dvd_mult_iff[where ?'a = int]
+lemmas prime_dvd_power_nat = prime_dvd_power[where ?'a = nat]
+lemmas prime_dvd_power_int = prime_dvd_power[where ?'a = int]
+lemmas prime_dvd_power_nat_iff = prime_dvd_power_iff[where ?'a = nat]
+lemmas prime_dvd_power_int_iff = prime_dvd_power_iff[where ?'a = int]
+lemmas prime_imp_power_coprime_nat = prime_imp_power_coprime[where ?'a = nat]
+lemmas prime_imp_power_coprime_int = prime_imp_power_coprime[where ?'a = int]
+lemmas primes_coprime_nat = primes_coprime[where ?'a = nat]
+lemmas primes_coprime_int = primes_coprime[where ?'a = nat]
+lemmas prime_divprod_pow_nat = prime_elem_divprod_pow[where ?'a = nat]
+lemmas prime_exp = prime_elem_power_iff[where ?'a = nat]
+
+text \<open>Code generation\<close>
+  
+context
+begin
+
+qualified definition prime_nat :: "nat \<Rightarrow> bool"
+  where [simp, code_abbrev]: "prime_nat = prime"
+
+lemma prime_nat_naive [code]:
+  "prime_nat p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
+  by (auto simp add: prime_nat_iff')
+
+qualified definition prime_int :: "int \<Rightarrow> bool"
+  where [simp, code_abbrev]: "prime_int = prime"
+
+lemma prime_int_naive [code]:
+  "prime_int p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
+  by (auto simp add: prime_int_iff')
+
+lemma "prime(997::nat)" by eval
+
+lemma "prime(997::int)" by eval
+  
+end
+
+end
--- a/src/HOL/Isar_Examples/Fibonacci.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Isar_Examples/Fibonacci.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -15,7 +15,7 @@
 section \<open>Fib and Gcd commute\<close>
 
 theory Fibonacci
-  imports "../Number_Theory/Primes"
+  imports "../Computational_Algebra/Primes"
 begin
 
 text_raw \<open>\<^footnote>\<open>Isar version by Gertrud Bauer. Original tactic script by Larry
--- a/src/HOL/Library/Field_as_Ring.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,120 +0,0 @@
-(*  Title:      HOL/Library/Field_as_Ring.thy
-    Author:     Manuel Eberl
-*)
-
-theory Field_as_Ring
-imports 
-  Complex_Main
-  "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
-begin
-
-context field
-begin
-
-subclass idom_divide ..
-
-definition normalize_field :: "'a \<Rightarrow> 'a" 
-  where [simp]: "normalize_field x = (if x = 0 then 0 else 1)"
-definition unit_factor_field :: "'a \<Rightarrow> 'a" 
-  where [simp]: "unit_factor_field x = x"
-definition euclidean_size_field :: "'a \<Rightarrow> nat" 
-  where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)"
-definition mod_field :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
-  where [simp]: "mod_field x y = (if y = 0 then x else 0)"
-
-end
-
-instantiation real :: unique_euclidean_ring
-begin
-
-definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)"
-definition [simp]: "unit_factor_real = (unit_factor_field :: real \<Rightarrow> _)"
-definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \<Rightarrow> _)"
-definition [simp]: "uniqueness_constraint_real = (top :: real \<Rightarrow> real \<Rightarrow> bool)"
-definition [simp]: "modulo_real = (mod_field :: real \<Rightarrow> _)"
-
-instance
-  by standard
-    (simp_all add: dvd_field_iff divide_simps split: if_splits)
-
-end
-
-instantiation real :: euclidean_ring_gcd
-begin
-
-definition gcd_real :: "real \<Rightarrow> real \<Rightarrow> real" where
-  "gcd_real = Euclidean_Algorithm.gcd"
-definition lcm_real :: "real \<Rightarrow> real \<Rightarrow> real" where
-  "lcm_real = Euclidean_Algorithm.lcm"
-definition Gcd_real :: "real set \<Rightarrow> real" where
- "Gcd_real = Euclidean_Algorithm.Gcd"
-definition Lcm_real :: "real set \<Rightarrow> real" where
- "Lcm_real = Euclidean_Algorithm.Lcm"
-
-instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def)
-
-end
-
-instantiation rat :: unique_euclidean_ring
-begin
-
-definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)"
-definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \<Rightarrow> _)"
-definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \<Rightarrow> _)"
-definition [simp]: "uniqueness_constraint_rat = (top :: rat \<Rightarrow> rat \<Rightarrow> bool)"
-definition [simp]: "modulo_rat = (mod_field :: rat \<Rightarrow> _)"
-
-instance
-  by standard
-    (simp_all add: dvd_field_iff divide_simps split: if_splits)
-
-end
-
-instantiation rat :: euclidean_ring_gcd
-begin
-
-definition gcd_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
-  "gcd_rat = Euclidean_Algorithm.gcd"
-definition lcm_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
-  "lcm_rat = Euclidean_Algorithm.lcm"
-definition Gcd_rat :: "rat set \<Rightarrow> rat" where
- "Gcd_rat = Euclidean_Algorithm.Gcd"
-definition Lcm_rat :: "rat set \<Rightarrow> rat" where
- "Lcm_rat = Euclidean_Algorithm.Lcm"
-
-instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def)
-
-end
-
-instantiation complex :: unique_euclidean_ring
-begin
-
-definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)"
-definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \<Rightarrow> _)"
-definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \<Rightarrow> _)"
-definition [simp]: "uniqueness_constraint_complex = (top :: complex \<Rightarrow> complex \<Rightarrow> bool)"
-definition [simp]: "modulo_complex = (mod_field :: complex \<Rightarrow> _)"
-
-instance
-  by standard
-    (simp_all add: dvd_field_iff divide_simps split: if_splits)
-
-end
-
-instantiation complex :: euclidean_ring_gcd
-begin
-
-definition gcd_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
-  "gcd_complex = Euclidean_Algorithm.gcd"
-definition lcm_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
-  "lcm_complex = Euclidean_Algorithm.lcm"
-definition Gcd_complex :: "complex set \<Rightarrow> complex" where
- "Gcd_complex = Euclidean_Algorithm.Gcd"
-definition Lcm_complex :: "complex set \<Rightarrow> complex" where
- "Lcm_complex = Euclidean_Algorithm.Lcm"
-
-instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def)
-
-end
-
-end
--- a/src/HOL/Library/Formal_Power_Series.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4783 +0,0 @@
-(*  Title:      HOL/Library/Formal_Power_Series.thy
-    Author:     Amine Chaieb, University of Cambridge
-*)
-
-section \<open>A formalization of formal power series\<close>
-
-theory Formal_Power_Series
-imports Complex_Main "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
-begin
-
-
-subsection \<open>The type of formal power series\<close>
-
-typedef 'a fps = "{f :: nat \<Rightarrow> 'a. True}"
-  morphisms fps_nth Abs_fps
-  by simp
-
-notation fps_nth (infixl "$" 75)
-
-lemma expand_fps_eq: "p = q \<longleftrightarrow> (\<forall>n. p $ n = q $ n)"
-  by (simp add: fps_nth_inject [symmetric] fun_eq_iff)
-
-lemma fps_ext: "(\<And>n. p $ n = q $ n) \<Longrightarrow> p = q"
-  by (simp add: expand_fps_eq)
-
-lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n"
-  by (simp add: Abs_fps_inverse)
-
-text \<open>Definition of the basic elements 0 and 1 and the basic operations of addition,
-  negation and multiplication.\<close>
-
-instantiation fps :: (zero) zero
-begin
-  definition fps_zero_def: "0 = Abs_fps (\<lambda>n. 0)"
-  instance ..
-end
-
-lemma fps_zero_nth [simp]: "0 $ n = 0"
-  unfolding fps_zero_def by simp
-
-instantiation fps :: ("{one, zero}") one
-begin
-  definition fps_one_def: "1 = Abs_fps (\<lambda>n. if n = 0 then 1 else 0)"
-  instance ..
-end
-
-lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)"
-  unfolding fps_one_def by simp
-
-instantiation fps :: (plus) plus
-begin
-  definition fps_plus_def: "op + = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n + g $ n))"
-  instance ..
-end
-
-lemma fps_add_nth [simp]: "(f + g) $ n = f $ n + g $ n"
-  unfolding fps_plus_def by simp
-
-instantiation fps :: (minus) minus
-begin
-  definition fps_minus_def: "op - = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n - g $ n))"
-  instance ..
-end
-
-lemma fps_sub_nth [simp]: "(f - g) $ n = f $ n - g $ n"
-  unfolding fps_minus_def by simp
-
-instantiation fps :: (uminus) uminus
-begin
-  definition fps_uminus_def: "uminus = (\<lambda>f. Abs_fps (\<lambda>n. - (f $ n)))"
-  instance ..
-end
-
-lemma fps_neg_nth [simp]: "(- f) $ n = - (f $ n)"
-  unfolding fps_uminus_def by simp
-
-instantiation fps :: ("{comm_monoid_add, times}") times
-begin
-  definition fps_times_def: "op * = (\<lambda>f g. Abs_fps (\<lambda>n. \<Sum>i=0..n. f $ i * g $ (n - i)))"
-  instance ..
-end
-
-lemma fps_mult_nth: "(f * g) $ n = (\<Sum>i=0..n. f$i * g$(n - i))"
-  unfolding fps_times_def by simp
-
-lemma fps_mult_nth_0 [simp]: "(f * g) $ 0 = f $ 0 * g $ 0"
-  unfolding fps_times_def by simp
-
-declare atLeastAtMost_iff [presburger]
-declare Bex_def [presburger]
-declare Ball_def [presburger]
-
-lemma mult_delta_left:
-  fixes x y :: "'a::mult_zero"
-  shows "(if b then x else 0) * y = (if b then x * y else 0)"
-  by simp
-
-lemma mult_delta_right:
-  fixes x y :: "'a::mult_zero"
-  shows "x * (if b then y else 0) = (if b then x * y else 0)"
-  by simp
-
-lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
-  by auto
-
-lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
-  by auto
-
-
-subsection \<open>Formal power series form a commutative ring with unity, if the range of sequences
-  they represent is a commutative ring with unity\<close>
-
-instance fps :: (semigroup_add) semigroup_add
-proof
-  fix a b c :: "'a fps"
-  show "a + b + c = a + (b + c)"
-    by (simp add: fps_ext add.assoc)
-qed
-
-instance fps :: (ab_semigroup_add) ab_semigroup_add
-proof
-  fix a b :: "'a fps"
-  show "a + b = b + a"
-    by (simp add: fps_ext add.commute)
-qed
-
-lemma fps_mult_assoc_lemma:
-  fixes k :: nat
-    and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
-  shows "(\<Sum>j=0..k. \<Sum>i=0..j. f i (j - i) (n - j)) =
-         (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
-  by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc)
-
-instance fps :: (semiring_0) semigroup_mult
-proof
-  fix a b c :: "'a fps"
-  show "(a * b) * c = a * (b * c)"
-  proof (rule fps_ext)
-    fix n :: nat
-    have "(\<Sum>j=0..n. \<Sum>i=0..j. a$i * b$(j - i) * c$(n - j)) =
-          (\<Sum>j=0..n. \<Sum>i=0..n - j. a$j * b$i * c$(n - j - i))"
-      by (rule fps_mult_assoc_lemma)
-    then show "((a * b) * c) $ n = (a * (b * c)) $ n"
-      by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc)
-  qed
-qed
-
-lemma fps_mult_commute_lemma:
-  fixes n :: nat
-    and f :: "nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
-  shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
-  by (rule sum.reindex_bij_witness[where i="op - n" and j="op - n"]) auto
-
-instance fps :: (comm_semiring_0) ab_semigroup_mult
-proof
-  fix a b :: "'a fps"
-  show "a * b = b * a"
-  proof (rule fps_ext)
-    fix n :: nat
-    have "(\<Sum>i=0..n. a$i * b$(n - i)) = (\<Sum>i=0..n. a$(n - i) * b$i)"
-      by (rule fps_mult_commute_lemma)
-    then show "(a * b) $ n = (b * a) $ n"
-      by (simp add: fps_mult_nth mult.commute)
-  qed
-qed
-
-instance fps :: (monoid_add) monoid_add
-proof
-  fix a :: "'a fps"
-  show "0 + a = a" by (simp add: fps_ext)
-  show "a + 0 = a" by (simp add: fps_ext)
-qed
-
-instance fps :: (comm_monoid_add) comm_monoid_add
-proof
-  fix a :: "'a fps"
-  show "0 + a = a" by (simp add: fps_ext)
-qed
-
-instance fps :: (semiring_1) monoid_mult
-proof
-  fix a :: "'a fps"
-  show "1 * a = a"
-    by (simp add: fps_ext fps_mult_nth mult_delta_left sum.delta)
-  show "a * 1 = a"
-    by (simp add: fps_ext fps_mult_nth mult_delta_right sum.delta')
-qed
-
-instance fps :: (cancel_semigroup_add) cancel_semigroup_add
-proof
-  fix a b c :: "'a fps"
-  show "b = c" if "a + b = a + c"
-    using that by (simp add: expand_fps_eq)
-  show "b = c" if "b + a = c + a"
-    using that by (simp add: expand_fps_eq)
-qed
-
-instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
-proof
-  fix a b c :: "'a fps"
-  show "a + b - a = b"
-    by (simp add: expand_fps_eq)
-  show "a - b - c = a - (b + c)"
-    by (simp add: expand_fps_eq diff_diff_eq)
-qed
-
-instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
-
-instance fps :: (group_add) group_add
-proof
-  fix a b :: "'a fps"
-  show "- a + a = 0" by (simp add: fps_ext)
-  show "a + - b = a - b" by (simp add: fps_ext)
-qed
-
-instance fps :: (ab_group_add) ab_group_add
-proof
-  fix a b :: "'a fps"
-  show "- a + a = 0" by (simp add: fps_ext)
-  show "a - b = a + - b" by (simp add: fps_ext)
-qed
-
-instance fps :: (zero_neq_one) zero_neq_one
-  by standard (simp add: expand_fps_eq)
-
-instance fps :: (semiring_0) semiring
-proof
-  fix a b c :: "'a fps"
-  show "(a + b) * c = a * c + b * c"
-    by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib)
-  show "a * (b + c) = a * b + a * c"
-    by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib)
-qed
-
-instance fps :: (semiring_0) semiring_0
-proof
-  fix a :: "'a fps"
-  show "0 * a = 0"
-    by (simp add: fps_ext fps_mult_nth)
-  show "a * 0 = 0"
-    by (simp add: fps_ext fps_mult_nth)
-qed
-
-instance fps :: (semiring_0_cancel) semiring_0_cancel ..
-
-instance fps :: (semiring_1) semiring_1 ..
-
-
-subsection \<open>Selection of the nth power of the implicit variable in the infinite sum\<close>
-
-lemma fps_square_nth: "(f^2) $ n = (\<Sum>k\<le>n. f $ k * f $ (n - k))"
-  by (simp add: power2_eq_square fps_mult_nth atLeast0AtMost)
-
-lemma fps_nonzero_nth: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0)"
-  by (simp add: expand_fps_eq)
-
-lemma fps_nonzero_nth_minimal: "f \<noteq> 0 \<longleftrightarrow> (\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m < n. f $ m = 0))"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  let ?n = "LEAST n. f $ n \<noteq> 0"
-  show ?rhs if ?lhs
-  proof -
-    from that have "\<exists>n. f $ n \<noteq> 0"
-      by (simp add: fps_nonzero_nth)
-    then have "f $ ?n \<noteq> 0"
-      by (rule LeastI_ex)
-    moreover have "\<forall>m<?n. f $ m = 0"
-      by (auto dest: not_less_Least)
-    ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
-    then show ?thesis ..
-  qed
-  show ?lhs if ?rhs
-    using that by (auto simp add: expand_fps_eq)
-qed
-
-lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
-  by (rule expand_fps_eq)
-
-lemma fps_sum_nth: "sum f S $ n = sum (\<lambda>k. (f k) $ n) S"
-proof (cases "finite S")
-  case True
-  then show ?thesis by (induct set: finite) auto
-next
-  case False
-  then show ?thesis by simp
-qed
-
-
-subsection \<open>Injection of the basic ring elements and multiplication by scalars\<close>
-
-definition "fps_const c = Abs_fps (\<lambda>n. if n = 0 then c else 0)"
-
-lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)"
-  unfolding fps_const_def by simp
-
-lemma fps_const_0_eq_0 [simp]: "fps_const 0 = 0"
-  by (simp add: fps_ext)
-
-lemma fps_const_1_eq_1 [simp]: "fps_const 1 = 1"
-  by (simp add: fps_ext)
-
-lemma fps_const_neg [simp]: "- (fps_const (c::'a::ring)) = fps_const (- c)"
-  by (simp add: fps_ext)
-
-lemma fps_const_add [simp]: "fps_const (c::'a::monoid_add) + fps_const d = fps_const (c + d)"
-  by (simp add: fps_ext)
-
-lemma fps_const_sub [simp]: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
-  by (simp add: fps_ext)
-
-lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
-  by (simp add: fps_eq_iff fps_mult_nth sum.neutral)
-
-lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
-    Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
-  by (simp add: fps_ext)
-
-lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) =
-    Abs_fps (\<lambda>n. if n = 0 then f$0 + c else f$n)"
-  by (simp add: fps_ext)
-
-lemma fps_const_mult_left: "fps_const (c::'a::semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
-  unfolding fps_eq_iff fps_mult_nth
-  by (simp add: fps_const_def mult_delta_left sum.delta)
-
-lemma fps_const_mult_right: "f * fps_const (c::'a::semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
-  unfolding fps_eq_iff fps_mult_nth
-  by (simp add: fps_const_def mult_delta_right sum.delta')
-
-lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
-  by (simp add: fps_mult_nth mult_delta_left sum.delta)
-
-lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
-  by (simp add: fps_mult_nth mult_delta_right sum.delta')
-
-
-subsection \<open>Formal power series form an integral domain\<close>
-
-instance fps :: (ring) ring ..
-
-instance fps :: (ring_1) ring_1
-  by (intro_classes, auto simp add: distrib_right)
-
-instance fps :: (comm_ring_1) comm_ring_1
-  by (intro_classes, auto simp add: distrib_right)
-
-instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
-proof
-  fix a b :: "'a fps"
-  assume "a \<noteq> 0" and "b \<noteq> 0"
-  then obtain i j where i: "a $ i \<noteq> 0" "\<forall>k<i. a $ k = 0" and j: "b $ j \<noteq> 0" "\<forall>k<j. b $ k =0"
-    unfolding fps_nonzero_nth_minimal
-    by blast+
-  have "(a * b) $ (i + j) = (\<Sum>k=0..i+j. a $ k * b $ (i + j - k))"
-    by (rule fps_mult_nth)
-  also have "\<dots> = (a $ i * b $ (i + j - i)) + (\<Sum>k\<in>{0..i+j} - {i}. a $ k * b $ (i + j - k))"
-    by (rule sum.remove) simp_all
-  also have "(\<Sum>k\<in>{0..i+j}-{i}. a $ k * b $ (i + j - k)) = 0"
-  proof (rule sum.neutral [rule_format])
-    fix k assume "k \<in> {0..i+j} - {i}"
-    then have "k < i \<or> i+j-k < j"
-      by auto
-    then show "a $ k * b $ (i + j - k) = 0"
-      using i j by auto
-  qed
-  also have "a $ i * b $ (i + j - i) + 0 = a $ i * b $ j"
-    by simp
-  also have "a $ i * b $ j \<noteq> 0"
-    using i j by simp
-  finally have "(a*b) $ (i+j) \<noteq> 0" .
-  then show "a * b \<noteq> 0"
-    unfolding fps_nonzero_nth by blast
-qed
-
-instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
-
-instance fps :: (idom) idom ..
-
-lemma numeral_fps_const: "numeral k = fps_const (numeral k)"
-  by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1
-    fps_const_add [symmetric])
-
-lemma neg_numeral_fps_const:
-  "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)"
-  by (simp add: numeral_fps_const)
-
-lemma fps_numeral_nth: "numeral n $ i = (if i = 0 then numeral n else 0)"
-  by (simp add: numeral_fps_const)
-
-lemma fps_numeral_nth_0 [simp]: "numeral n $ 0 = numeral n"
-  by (simp add: numeral_fps_const)
-
-lemma fps_of_nat: "fps_const (of_nat c) = of_nat c"
-  by (induction c) (simp_all add: fps_const_add [symmetric] del: fps_const_add)
-
-lemma numeral_neq_fps_zero [simp]: "(numeral f :: 'a :: field_char_0 fps) \<noteq> 0"
-proof
-  assume "numeral f = (0 :: 'a fps)"
-  from arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] show False by simp
-qed 
-
-
-subsection \<open>The eXtractor series X\<close>
-
-lemma minus_one_power_iff: "(- (1::'a::comm_ring_1)) ^ n = (if even n then 1 else - 1)"
-  by (induct n) auto
-
-definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
-
-lemma X_mult_nth [simp]:
-  "(X * (f :: 'a::semiring_1 fps)) $n = (if n = 0 then 0 else f $ (n - 1))"
-proof (cases "n = 0")
-  case False
-  have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
-    by (simp add: fps_mult_nth)
-  also have "\<dots> = f $ (n - 1)"
-    using False by (simp add: X_def mult_delta_left sum.delta)
-  finally show ?thesis
-    using False by simp
-next
-  case True
-  then show ?thesis
-    by (simp add: fps_mult_nth X_def)
-qed
-
-lemma X_mult_right_nth[simp]:
-  "((a::'a::semiring_1 fps) * X) $ n = (if n = 0 then 0 else a $ (n - 1))"
-proof -
-  have "(a * X) $ n = (\<Sum>i = 0..n. a $ i * (if n - i = Suc 0 then 1 else 0))"
-    by (simp add: fps_times_def X_def)
-  also have "\<dots> = (\<Sum>i = 0..n. if i = n - 1 then if n = 0 then 0 else a $ i else 0)"
-    by (intro sum.cong) auto
-  also have "\<dots> = (if n = 0 then 0 else a $ (n - 1))" by (simp add: sum.delta)
-  finally show ?thesis .
-qed
-
-lemma fps_mult_X_commute: "X * (a :: 'a :: semiring_1 fps) = a * X" 
-  by (simp add: fps_eq_iff)
-
-lemma X_power_iff: "X^k = Abs_fps (\<lambda>n. if n = k then 1::'a::comm_ring_1 else 0)"
-proof (induct k)
-  case 0
-  then show ?case by (simp add: X_def fps_eq_iff)
-next
-  case (Suc k)
-  have "(X^Suc k) $ m = (if m = Suc k then 1::'a else 0)" for m
-  proof -
-    have "(X^Suc k) $ m = (if m = 0 then 0 else (X^k) $ (m - 1))"
-      by (simp del: One_nat_def)
-    then show ?thesis
-      using Suc.hyps by (auto cong del: if_weak_cong)
-  qed
-  then show ?case
-    by (simp add: fps_eq_iff)
-qed
-
-lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)"
-  by (simp add: X_def)
-
-lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else 0::'a::comm_ring_1)"
-  by (simp add: X_power_iff)
-
-lemma X_power_mult_nth: "(X^k * (f :: 'a::comm_ring_1 fps)) $n = (if n < k then 0 else f $ (n - k))"
-  apply (induct k arbitrary: n)
-  apply simp
-  unfolding power_Suc mult.assoc
-  apply (case_tac n)
-  apply auto
-  done
-
-lemma X_power_mult_right_nth:
-    "((f :: 'a::comm_ring_1 fps) * X^k) $n = (if n < k then 0 else f $ (n - k))"
-  by (metis X_power_mult_nth mult.commute)
-
-
-lemma X_neq_fps_const [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> fps_const c"
-proof
-  assume "(X::'a fps) = fps_const (c::'a)"
-  hence "X$1 = (fps_const (c::'a))$1" by (simp only:)
-  thus False by auto
-qed
-
-lemma X_neq_zero [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 0"
-  by (simp only: fps_const_0_eq_0[symmetric] X_neq_fps_const) simp
-
-lemma X_neq_one [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 1"
-  by (simp only: fps_const_1_eq_1[symmetric] X_neq_fps_const) simp
-
-lemma X_neq_numeral [simp]: "(X :: 'a :: {semiring_1,zero_neq_one} fps) \<noteq> numeral c"
-  by (simp only: numeral_fps_const X_neq_fps_const) simp
-
-lemma X_pow_eq_X_pow_iff [simp]:
-  "(X :: ('a :: {comm_ring_1}) fps) ^ m = X ^ n \<longleftrightarrow> m = n"
-proof
-  assume "(X :: 'a fps) ^ m = X ^ n"
-  hence "(X :: 'a fps) ^ m $ m = X ^ n $ m" by (simp only:)
-  thus "m = n" by (simp split: if_split_asm)
-qed simp_all
-
-
-subsection \<open>Subdegrees\<close>
-
-definition subdegree :: "('a::zero) fps \<Rightarrow> nat" where
-  "subdegree f = (if f = 0 then 0 else LEAST n. f$n \<noteq> 0)"
-
-lemma subdegreeI:
-  assumes "f $ d \<noteq> 0" and "\<And>i. i < d \<Longrightarrow> f $ i = 0"
-  shows   "subdegree f = d"
-proof-
-  from assms(1) have "f \<noteq> 0" by auto
-  moreover from assms(1) have "(LEAST i. f $ i \<noteq> 0) = d"
-  proof (rule Least_equality)
-    fix e assume "f $ e \<noteq> 0"
-    with assms(2) have "\<not>(e < d)" by blast
-    thus "e \<ge> d" by simp
-  qed
-  ultimately show ?thesis unfolding subdegree_def by simp
-qed
-
-lemma nth_subdegree_nonzero [simp,intro]: "f \<noteq> 0 \<Longrightarrow> f $ subdegree f \<noteq> 0"
-proof-
-  assume "f \<noteq> 0"
-  hence "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
-  also from \<open>f \<noteq> 0\<close> have "\<exists>n. f$n \<noteq> 0" using fps_nonzero_nth by blast
-  from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \<noteq> 0) \<noteq> 0" .
-  finally show ?thesis .
-qed
-
-lemma nth_less_subdegree_zero [dest]: "n < subdegree f \<Longrightarrow> f $ n = 0"
-proof (cases "f = 0")
-  assume "f \<noteq> 0" and less: "n < subdegree f"
-  note less
-  also from \<open>f \<noteq> 0\<close> have "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
-  finally show "f $ n = 0" using not_less_Least by blast
-qed simp_all
-
-lemma subdegree_geI:
-  assumes "f \<noteq> 0" "\<And>i. i < n \<Longrightarrow> f$i = 0"
-  shows   "subdegree f \<ge> n"
-proof (rule ccontr)
-  assume "\<not>(subdegree f \<ge> n)"
-  with assms(2) have "f $ subdegree f = 0" by simp
-  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
-  ultimately show False by contradiction
-qed
-
-lemma subdegree_greaterI:
-  assumes "f \<noteq> 0" "\<And>i. i \<le> n \<Longrightarrow> f$i = 0"
-  shows   "subdegree f > n"
-proof (rule ccontr)
-  assume "\<not>(subdegree f > n)"
-  with assms(2) have "f $ subdegree f = 0" by simp
-  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
-  ultimately show False by contradiction
-qed
-
-lemma subdegree_leI:
-  "f $ n \<noteq> 0 \<Longrightarrow> subdegree f \<le> n"
-  by (rule leI) auto
-
-
-lemma subdegree_0 [simp]: "subdegree 0 = 0"
-  by (simp add: subdegree_def)
-
-lemma subdegree_1 [simp]: "subdegree (1 :: ('a :: zero_neq_one) fps) = 0"
-  by (auto intro!: subdegreeI)
-
-lemma subdegree_X [simp]: "subdegree (X :: ('a :: zero_neq_one) fps) = 1"
-  by (auto intro!: subdegreeI simp: X_def)
-
-lemma subdegree_fps_const [simp]: "subdegree (fps_const c) = 0"
-  by (cases "c = 0") (auto intro!: subdegreeI)
-
-lemma subdegree_numeral [simp]: "subdegree (numeral n) = 0"
-  by (simp add: numeral_fps_const)
-
-lemma subdegree_eq_0_iff: "subdegree f = 0 \<longleftrightarrow> f = 0 \<or> f $ 0 \<noteq> 0"
-proof (cases "f = 0")
-  assume "f \<noteq> 0"
-  thus ?thesis
-    using nth_subdegree_nonzero[OF \<open>f \<noteq> 0\<close>] by (fastforce intro!: subdegreeI)
-qed simp_all
-
-lemma subdegree_eq_0 [simp]: "f $ 0 \<noteq> 0 \<Longrightarrow> subdegree f = 0"
-  by (simp add: subdegree_eq_0_iff)
-
-lemma nth_subdegree_mult [simp]:
-  fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps"
-  shows "(f * g) $ (subdegree f + subdegree g) = f $ subdegree f * g $ subdegree g"
-proof-
-  let ?n = "subdegree f + subdegree g"
-  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))"
-    by (simp add: fps_mult_nth)
-  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
-  proof (intro sum.cong)
-    fix x assume x: "x \<in> {0..?n}"
-    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
-    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
-      by (elim disjE conjE) auto
-  qed auto
-  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
-  finally show ?thesis .
-qed
-
-lemma subdegree_mult [simp]:
-  assumes "f \<noteq> 0" "g \<noteq> 0"
-  shows "subdegree ((f :: ('a :: {ring_no_zero_divisors}) fps) * g) = subdegree f + subdegree g"
-proof (rule subdegreeI)
-  let ?n = "subdegree f + subdegree g"
-  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth)
-  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
-  proof (intro sum.cong)
-    fix x assume x: "x \<in> {0..?n}"
-    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
-    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
-      by (elim disjE conjE) auto
-  qed auto
-  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
-  also from assms have "... \<noteq> 0" by auto
-  finally show "(f * g) $ (subdegree f + subdegree g) \<noteq> 0" .
-next
-  fix m assume m: "m < subdegree f + subdegree g"
-  have "(f * g) $ m = (\<Sum>i=0..m. f$i * g$(m-i))" by (simp add: fps_mult_nth)
-  also have "... = (\<Sum>i=0..m. 0)"
-  proof (rule sum.cong)
-    fix i assume "i \<in> {0..m}"
-    with m have "i < subdegree f \<or> m - i < subdegree g" by auto
-    thus "f$i * g$(m-i) = 0" by (elim disjE) auto
-  qed auto
-  finally show "(f * g) $ m = 0" by simp
-qed
-
-lemma subdegree_power [simp]:
-  "subdegree ((f :: ('a :: ring_1_no_zero_divisors) fps) ^ n) = n * subdegree f"
-  by (cases "f = 0"; induction n) simp_all
-
-lemma subdegree_uminus [simp]:
-  "subdegree (-(f::('a::group_add) fps)) = subdegree f"
-  by (simp add: subdegree_def)
-
-lemma subdegree_minus_commute [simp]:
-  "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)"
-proof -
-  have "f - g = -(g - f)" by simp
-  also have "subdegree ... = subdegree (g - f)" by (simp only: subdegree_uminus)
-  finally show ?thesis .
-qed
-
-lemma subdegree_add_ge:
-  assumes "f \<noteq> -(g :: ('a :: {group_add}) fps)"
-  shows   "subdegree (f + g) \<ge> min (subdegree f) (subdegree g)"
-proof (rule subdegree_geI)
-  from assms show "f + g \<noteq> 0" by (subst (asm) eq_neg_iff_add_eq_0)
-next
-  fix i assume "i < min (subdegree f) (subdegree g)"
-  hence "f $ i = 0" and "g $ i = 0" by auto
-  thus "(f + g) $ i = 0" by force
-qed
-
-lemma subdegree_add_eq1:
-  assumes "f \<noteq> 0"
-  assumes "subdegree f < subdegree (g :: ('a :: {group_add}) fps)"
-  shows   "subdegree (f + g) = subdegree f"
-proof (rule antisym[OF subdegree_leI])
-  from assms show "subdegree (f + g) \<ge> subdegree f"
-    by (intro order.trans[OF min.boundedI subdegree_add_ge]) auto
-  from assms have "f $ subdegree f \<noteq> 0" "g $ subdegree f = 0" by auto
-  thus "(f + g) $ subdegree f \<noteq> 0" by simp
-qed
-
-lemma subdegree_add_eq2:
-  assumes "g \<noteq> 0"
-  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
-  shows   "subdegree (f + g) = subdegree g"
-  using subdegree_add_eq1[OF assms] by (simp add: add.commute)
-
-lemma subdegree_diff_eq1:
-  assumes "f \<noteq> 0"
-  assumes "subdegree f < subdegree (g :: ('a :: {ab_group_add}) fps)"
-  shows   "subdegree (f - g) = subdegree f"
-  using subdegree_add_eq1[of f "-g"] assms by (simp add: add.commute)
-
-lemma subdegree_diff_eq2:
-  assumes "g \<noteq> 0"
-  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
-  shows   "subdegree (f - g) = subdegree g"
-  using subdegree_add_eq2[of "-g" f] assms by (simp add: add.commute)
-
-lemma subdegree_diff_ge [simp]:
-  assumes "f \<noteq> (g :: ('a :: {group_add}) fps)"
-  shows   "subdegree (f - g) \<ge> min (subdegree f) (subdegree g)"
-  using assms subdegree_add_ge[of f "-g"] by simp
-
-
-
-
-subsection \<open>Shifting and slicing\<close>
-
-definition fps_shift :: "nat \<Rightarrow> 'a fps \<Rightarrow> 'a fps" where
-  "fps_shift n f = Abs_fps (\<lambda>i. f $ (i + n))"
-
-lemma fps_shift_nth [simp]: "fps_shift n f $ i = f $ (i + n)"
-  by (simp add: fps_shift_def)
-
-lemma fps_shift_0 [simp]: "fps_shift 0 f = f"
-  by (intro fps_ext) (simp add: fps_shift_def)
-
-lemma fps_shift_zero [simp]: "fps_shift n 0 = 0"
-  by (intro fps_ext) (simp add: fps_shift_def)
-
-lemma fps_shift_one: "fps_shift n 1 = (if n = 0 then 1 else 0)"
-  by (intro fps_ext) (simp add: fps_shift_def)
-
-lemma fps_shift_fps_const: "fps_shift n (fps_const c) = (if n = 0 then fps_const c else 0)"
-  by (intro fps_ext) (simp add: fps_shift_def)
-
-lemma fps_shift_numeral: "fps_shift n (numeral c) = (if n = 0 then numeral c else 0)"
-  by (simp add: numeral_fps_const fps_shift_fps_const)
-
-lemma fps_shift_X_power [simp]:
-  "n \<le> m \<Longrightarrow> fps_shift n (X ^ m) = (X ^ (m - n) ::'a::comm_ring_1 fps)"
-  by (intro fps_ext) (auto simp: fps_shift_def )
-
-lemma fps_shift_times_X_power:
-  "n \<le> subdegree f \<Longrightarrow> fps_shift n f * X ^ n = (f :: 'a :: comm_ring_1 fps)"
-  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
-
-lemma fps_shift_times_X_power' [simp]:
-  "fps_shift n (f * X^n) = (f :: 'a :: comm_ring_1 fps)"
-  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
-
-lemma fps_shift_times_X_power'':
-  "m \<le> n \<Longrightarrow> fps_shift n (f * X^m) = fps_shift (n - m) (f :: 'a :: comm_ring_1 fps)"
-  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
-
-lemma fps_shift_subdegree [simp]:
-  "n \<le> subdegree f \<Longrightarrow> subdegree (fps_shift n f) = subdegree (f :: 'a :: comm_ring_1 fps) - n"
-  by (cases "f = 0") (force intro: nth_less_subdegree_zero subdegreeI)+
-
-lemma subdegree_decompose:
-  "f = fps_shift (subdegree f) f * X ^ subdegree (f :: ('a :: comm_ring_1) fps)"
-  by (rule fps_ext) (auto simp: X_power_mult_right_nth)
-
-lemma subdegree_decompose':
-  "n \<le> subdegree (f :: ('a :: comm_ring_1) fps) \<Longrightarrow> f = fps_shift n f * X^n"
-  by (rule fps_ext) (auto simp: X_power_mult_right_nth intro!: nth_less_subdegree_zero)
-
-lemma fps_shift_fps_shift:
-  "fps_shift (m + n) f = fps_shift m (fps_shift n f)"
-  by (rule fps_ext) (simp add: add_ac)
-
-lemma fps_shift_add:
-  "fps_shift n (f + g) = fps_shift n f + fps_shift n g"
-  by (simp add: fps_eq_iff)
-
-lemma fps_shift_mult:
-  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
-  shows   "fps_shift n (h*g) = h * fps_shift n g"
-proof -
-  from assms have "g = fps_shift n g * X^n" by (rule subdegree_decompose')
-  also have "h * ... = (h * fps_shift n g) * X^n" by simp
-  also have "fps_shift n ... = h * fps_shift n g" by simp
-  finally show ?thesis .
-qed
-
-lemma fps_shift_mult_right:
-  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
-  shows   "fps_shift n (g*h) = h * fps_shift n g"
-  by (subst mult.commute, subst fps_shift_mult) (simp_all add: assms)
-
-lemma nth_subdegree_zero_iff [simp]: "f $ subdegree f = 0 \<longleftrightarrow> f = 0"
-  by (cases "f = 0") auto
-
-lemma fps_shift_subdegree_zero_iff [simp]:
-  "fps_shift (subdegree f) f = 0 \<longleftrightarrow> f = 0"
-  by (subst (1) nth_subdegree_zero_iff[symmetric], cases "f = 0")
-     (simp_all del: nth_subdegree_zero_iff)
-
-
-definition "fps_cutoff n f = Abs_fps (\<lambda>i. if i < n then f$i else 0)"
-
-lemma fps_cutoff_nth [simp]: "fps_cutoff n f $ i = (if i < n then f$i else 0)"
-  unfolding fps_cutoff_def by simp
-
-lemma fps_cutoff_zero_iff: "fps_cutoff n f = 0 \<longleftrightarrow> (f = 0 \<or> n \<le> subdegree f)"
-proof
-  assume A: "fps_cutoff n f = 0"
-  thus "f = 0 \<or> n \<le> subdegree f"
-  proof (cases "f = 0")
-    assume "f \<noteq> 0"
-    with A have "n \<le> subdegree f"
-      by (intro subdegree_geI) (auto simp: fps_eq_iff split: if_split_asm)
-    thus ?thesis ..
-  qed simp
-qed (auto simp: fps_eq_iff intro: nth_less_subdegree_zero)
-
-lemma fps_cutoff_0 [simp]: "fps_cutoff 0 f = 0"
-  by (simp add: fps_eq_iff)
-
-lemma fps_cutoff_zero [simp]: "fps_cutoff n 0 = 0"
-  by (simp add: fps_eq_iff)
-
-lemma fps_cutoff_one: "fps_cutoff n 1 = (if n = 0 then 0 else 1)"
-  by (simp add: fps_eq_iff)
-
-lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)"
-  by (simp add: fps_eq_iff)
-
-lemma fps_cutoff_numeral: "fps_cutoff n (numeral c) = (if n = 0 then 0 else numeral c)"
-  by (simp add: numeral_fps_const fps_cutoff_fps_const)
-
-lemma fps_shift_cutoff:
-  "fps_shift n (f :: ('a :: comm_ring_1) fps) * X^n + fps_cutoff n f = f"
-  by (simp add: fps_eq_iff X_power_mult_right_nth)
-
-
-subsection \<open>Formal Power series form a metric space\<close>
-
-definition (in dist) "ball x r = {y. dist y x < r}"
-
-instantiation fps :: (comm_ring_1) dist
-begin
-
-definition
-  dist_fps_def: "dist (a :: 'a fps) b = (if a = b then 0 else inverse (2 ^ subdegree (a - b)))"
-
-lemma dist_fps_ge0: "dist (a :: 'a fps) b \<ge> 0"
-  by (simp add: dist_fps_def)
-
-lemma dist_fps_sym: "dist (a :: 'a fps) b = dist b a"
-  by (simp add: dist_fps_def)
-
-instance ..
-
-end
-
-instantiation fps :: (comm_ring_1) metric_space
-begin
-
-definition uniformity_fps_def [code del]:
-  "(uniformity :: ('a fps \<times> 'a fps) filter) = (INF e:{0 <..}. principal {(x, y). dist x y < e})"
-
-definition open_fps_def' [code del]:
-  "open (U :: 'a fps set) \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)"
-
-instance
-proof
-  show th: "dist a b = 0 \<longleftrightarrow> a = b" for a b :: "'a fps"
-    by (simp add: dist_fps_def split: if_split_asm)
-  then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp
-
-  fix a b c :: "'a fps"
-  consider "a = b" | "c = a \<or> c = b" | "a \<noteq> b" "a \<noteq> c" "b \<noteq> c" by blast
-  then show "dist a b \<le> dist a c + dist b c"
-  proof cases
-    case 1
-    then show ?thesis by (simp add: dist_fps_def)
-  next
-    case 2
-    then show ?thesis
-      by (cases "c = a") (simp_all add: th dist_fps_sym)
-  next
-    case neq: 3
-    have False if "dist a b > dist a c + dist b c"
-    proof -
-      let ?n = "subdegree (a - b)"
-      from neq have "dist a b > 0" "dist b c > 0" and "dist a c > 0" by (simp_all add: dist_fps_def)
-      with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all
-      with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)"
-        by (simp_all add: dist_fps_def field_simps)
-      hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0"
-        by (simp_all only: nth_less_subdegree_zero)
-      hence "(a - b) $ ?n = 0" by simp
-      moreover from neq have "(a - b) $ ?n \<noteq> 0" by (intro nth_subdegree_nonzero) simp_all
-      ultimately show False by contradiction
-    qed
-    thus ?thesis by (auto simp add: not_le[symmetric])
-  qed
-qed (rule open_fps_def' uniformity_fps_def)+
-
-end
-
-declare uniformity_Abort[where 'a="'a :: comm_ring_1 fps", code]
-
-lemma open_fps_def: "open (S :: 'a::comm_ring_1 fps set) = (\<forall>a \<in> S. \<exists>r. r >0 \<and> ball a r \<subseteq> S)"
-  unfolding open_dist ball_def subset_eq by simp
-
-text \<open>The infinite sums and justification of the notation in textbooks.\<close>
-
-lemma reals_power_lt_ex:
-  fixes x y :: real
-  assumes xp: "x > 0"
-    and y1: "y > 1"
-  shows "\<exists>k>0. (1/y)^k < x"
-proof -
-  have yp: "y > 0"
-    using y1 by simp
-  from reals_Archimedean2[of "max 0 (- log y x) + 1"]
-  obtain k :: nat where k: "real k > max 0 (- log y x) + 1"
-    by blast
-  from k have kp: "k > 0"
-    by simp
-  from k have "real k > - log y x"
-    by simp
-  then have "ln y * real k > - ln x"
-    unfolding log_def
-    using ln_gt_zero_iff[OF yp] y1
-    by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric])
-  then have "ln y * real k + ln x > 0"
-    by simp
-  then have "exp (real k * ln y + ln x) > exp 0"
-    by (simp add: ac_simps)
-  then have "y ^ k * x > 1"
-    unfolding exp_zero exp_add exp_real_of_nat_mult exp_ln [OF xp] exp_ln [OF yp]
-    by simp
-  then have "x > (1 / y)^k" using yp
-    by (simp add: field_simps)
-  then show ?thesis
-    using kp by blast
-qed
-
-lemma fps_sum_rep_nth: "(sum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
-    (if n \<le> m then a$n else 0::'a::comm_ring_1)"
-  apply (auto simp add: fps_sum_nth cond_value_iff cong del: if_weak_cong)
-  apply (simp add: sum.delta')
-  done
-
-lemma fps_notation: "(\<lambda>n. sum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) \<longlonglongrightarrow> a"
-  (is "?s \<longlonglongrightarrow> a")
-proof -
-  have "\<exists>n0. \<forall>n \<ge> n0. dist (?s n) a < r" if "r > 0" for r
-  proof -
-    obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0"
-      using reals_power_lt_ex[OF \<open>r > 0\<close>, of 2] by auto
-    show ?thesis
-    proof -
-      have "dist (?s n) a < r" if nn0: "n \<ge> n0" for n
-      proof -
-        from that have thnn0: "(1/2)^n \<le> (1/2 :: real)^n0"
-          by (simp add: divide_simps)
-        show ?thesis
-        proof (cases "?s n = a")
-          case True
-          then show ?thesis
-            unfolding dist_eq_0_iff[of "?s n" a, symmetric]
-            using \<open>r > 0\<close> by (simp del: dist_eq_0_iff)
-        next
-          case False
-          from False have dth: "dist (?s n) a = (1/2)^subdegree (?s n - a)"
-            by (simp add: dist_fps_def field_simps)
-          from False have kn: "subdegree (?s n - a) > n"
-            by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth)
-          then have "dist (?s n) a < (1/2)^n"
-            by (simp add: field_simps dist_fps_def)
-          also have "\<dots> \<le> (1/2)^n0"
-            using nn0 by (simp add: divide_simps)
-          also have "\<dots> < r"
-            using n0 by simp
-          finally show ?thesis .
-        qed
-      qed
-      then show ?thesis by blast
-    qed
-  qed
-  then show ?thesis
-    unfolding lim_sequentially by blast
-qed
-
-
-subsection \<open>Inverses of formal power series\<close>
-
-declare sum.cong[fundef_cong]
-
-instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse
-begin
-
-fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
-where
-  "natfun_inverse f 0 = inverse (f$0)"
-| "natfun_inverse f n = - inverse (f$0) * sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
-
-definition fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))"
-
-definition fps_divide_def:
-  "f div g = (if g = 0 then 0 else
-     let n = subdegree g; h = fps_shift n g
-     in  fps_shift n (f * inverse h))"
-
-instance ..
-
-end
-
-lemma fps_inverse_zero [simp]:
-  "inverse (0 :: 'a::{comm_monoid_add,inverse,times,uminus} fps) = 0"
-  by (simp add: fps_ext fps_inverse_def)
-
-lemma fps_inverse_one [simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1"
-  apply (auto simp add: expand_fps_eq fps_inverse_def)
-  apply (case_tac n)
-  apply auto
-  done
-
-lemma inverse_mult_eq_1 [intro]:
-  assumes f0: "f$0 \<noteq> (0::'a::field)"
-  shows "inverse f * f = 1"
-proof -
-  have c: "inverse f * f = f * inverse f"
-    by (simp add: mult.commute)
-  from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n"
-    by (simp add: fps_inverse_def)
-  from f0 have th0: "(inverse f * f) $ 0 = 1"
-    by (simp add: fps_mult_nth fps_inverse_def)
-  have "(inverse f * f)$n = 0" if np: "n > 0" for n
-  proof -
-    from np have eq: "{0..n} = {0} \<union> {1 .. n}"
-      by auto
-    have d: "{0} \<inter> {1 .. n} = {}"
-      by auto
-    from f0 np have th0: "- (inverse f $ n) =
-      (sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
-      by (cases n) (simp_all add: divide_inverse fps_inverse_def)
-    from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
-    have th1: "sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n"
-      by (simp add: field_simps)
-    have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
-      unfolding fps_mult_nth ifn ..
-    also have "\<dots> = f$0 * natfun_inverse f n + (\<Sum>i = 1..n. f$i * natfun_inverse f (n-i))"
-      by (simp add: eq)
-    also have "\<dots> = 0"
-      unfolding th1 ifn by simp
-    finally show ?thesis unfolding c .
-  qed
-  with th0 show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-lemma fps_inverse_0_iff[simp]: "(inverse f) $ 0 = (0::'a::division_ring) \<longleftrightarrow> f $ 0 = 0"
-  by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero)
-
-lemma fps_inverse_nth_0 [simp]: "inverse f $ 0 = inverse (f $ 0 :: 'a :: division_ring)"
-  by (simp add: fps_inverse_def)
-
-lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::division_ring) fps) \<longleftrightarrow> f $ 0 = 0"
-proof
-  assume A: "inverse f = 0"
-  have "0 = inverse f $ 0" by (subst A) simp
-  thus "f $ 0 = 0" by simp
-qed (simp add: fps_inverse_def)
-
-lemma fps_inverse_idempotent[intro, simp]:
-  assumes f0: "f$0 \<noteq> (0::'a::field)"
-  shows "inverse (inverse f) = f"
-proof -
-  from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
-  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
-  have "inverse f * f = inverse f * inverse (inverse f)"
-    by (simp add: ac_simps)
-  then show ?thesis
-    using f0 unfolding mult_cancel_left by simp
-qed
-
-lemma fps_inverse_unique:
-  assumes fg: "(f :: 'a :: field fps) * g = 1"
-  shows   "inverse f = g"
-proof -
-  have f0: "f $ 0 \<noteq> 0"
-  proof
-    assume "f $ 0 = 0"
-    hence "0 = (f * g) $ 0" by simp
-    also from fg have "(f * g) $ 0 = 1" by simp
-    finally show False by simp
-  qed
-  from inverse_mult_eq_1[OF this] fg
-  have th0: "inverse f * f = g * f"
-    by (simp add: ac_simps)
-  then show ?thesis
-    using f0
-    unfolding mult_cancel_right
-    by (auto simp add: expand_fps_eq)
-qed
-
-lemma fps_inverse_eq_0: "f$0 = 0 \<Longrightarrow> inverse (f :: 'a :: division_ring fps) = 0"
-  by simp
-  
-lemma sum_zero_lemma:
-  fixes n::nat
-  assumes "0 < n"
-  shows "(\<Sum>i = 0..n. if n = i then 1 else if n - i = 1 then - 1 else 0) = (0::'a::field)"
-proof -
-  let ?f = "\<lambda>i. if n = i then 1 else if n - i = 1 then - 1 else 0"
-  let ?g = "\<lambda>i. if i = n then 1 else if i = n - 1 then - 1 else 0"
-  let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
-  have th1: "sum ?f {0..n} = sum ?g {0..n}"
-    by (rule sum.cong) auto
-  have th2: "sum ?g {0..n - 1} = sum ?h {0..n - 1}"
-    apply (rule sum.cong)
-    using assms
-    apply auto
-    done
-  have eq: "{0 .. n} = {0.. n - 1} \<union> {n}"
-    by auto
-  from assms have d: "{0.. n - 1} \<inter> {n} = {}"
-    by auto
-  have f: "finite {0.. n - 1}" "finite {n}"
-    by auto
-  show ?thesis
-    unfolding th1
-    apply (simp add: sum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
-    unfolding th2
-    apply (simp add: sum.delta)
-    done
-qed
-
-lemma fps_inverse_mult: "inverse (f * g :: 'a::field fps) = inverse f * inverse g"
-proof (cases "f$0 = 0 \<or> g$0 = 0")
-  assume "\<not>(f$0 = 0 \<or> g$0 = 0)"
-  hence [simp]: "f$0 \<noteq> 0" "g$0 \<noteq> 0" by simp_all
-  show ?thesis
-  proof (rule fps_inverse_unique)
-    have "f * g * (inverse f * inverse g) = (inverse f * f) * (inverse g * g)" by simp
-    also have "... = 1" by (subst (1 2) inverse_mult_eq_1) simp_all
-    finally show "f * g * (inverse f * inverse g) = 1" .
-  qed
-next
-  assume A: "f$0 = 0 \<or> g$0 = 0"
-  hence "inverse (f * g) = 0" by simp
-  also from A have "... = inverse f * inverse g" by auto
-  finally show "inverse (f * g) = inverse f * inverse g" .
-qed
-
-
-lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) =
-    Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
-  apply (rule fps_inverse_unique)
-  apply (simp_all add: fps_eq_iff fps_mult_nth sum_zero_lemma)
-  done
-
-lemma subdegree_inverse [simp]: "subdegree (inverse (f::'a::field fps)) = 0"
-proof (cases "f$0 = 0")
-  assume nz: "f$0 \<noteq> 0"
-  hence "subdegree (inverse f) + subdegree f = subdegree (inverse f * f)"
-    by (subst subdegree_mult) auto
-  also from nz have "subdegree f = 0" by (simp add: subdegree_eq_0_iff)
-  also from nz have "inverse f * f = 1" by (rule inverse_mult_eq_1)
-  finally show "subdegree (inverse f) = 0" by simp
-qed (simp_all add: fps_inverse_def)
-
-lemma fps_is_unit_iff [simp]: "(f :: 'a :: field fps) dvd 1 \<longleftrightarrow> f $ 0 \<noteq> 0"
-proof
-  assume "f dvd 1"
-  then obtain g where "1 = f * g" by (elim dvdE)
-  from this[symmetric] have "(f*g) $ 0 = 1" by simp
-  thus "f $ 0 \<noteq> 0" by auto
-next
-  assume A: "f $ 0 \<noteq> 0"
-  thus "f dvd 1" by (simp add: inverse_mult_eq_1[OF A, symmetric])
-qed
-
-lemma subdegree_eq_0' [simp]: "(f :: 'a :: field fps) dvd 1 \<Longrightarrow> subdegree f = 0"
-  by simp
-
-lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \<noteq> 0 \<Longrightarrow> f dvd g"
-  by (rule dvd_trans, subst fps_is_unit_iff) simp_all
-
-instantiation fps :: (field) normalization_semidom
-begin
-
-definition fps_unit_factor_def [simp]:
-  "unit_factor f = fps_shift (subdegree f) f"
-
-definition fps_normalize_def [simp]:
-  "normalize f = (if f = 0 then 0 else X ^ subdegree f)"
-
-instance proof
-  fix f :: "'a fps"
-  show "unit_factor f * normalize f = f"
-    by (simp add: fps_shift_times_X_power)
-next
-  fix f g :: "'a fps"
-  show "unit_factor (f * g) = unit_factor f * unit_factor g"
-  proof (cases "f = 0 \<or> g = 0")
-    assume "\<not>(f = 0 \<or> g = 0)"
-    thus "unit_factor (f * g) = unit_factor f * unit_factor g"
-    unfolding fps_unit_factor_def
-      by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
-  qed auto
-next
-  fix f g :: "'a fps"
-  assume "g \<noteq> 0"
-  then have "f * (fps_shift (subdegree g) g * inverse (fps_shift (subdegree g) g)) = f"
-    by (metis add_cancel_right_left fps_shift_nth inverse_mult_eq_1 mult.commute mult_cancel_left2 nth_subdegree_nonzero)
-  then have "fps_shift (subdegree g) (g * (f * inverse (fps_shift (subdegree g) g))) = f"
-    by (simp add: fps_shift_mult_right mult.commute)
-  with \<open>g \<noteq> 0\<close> show "f * g / g = f"
-    by (simp add: fps_divide_def Let_def ac_simps)
-qed (auto simp add: fps_divide_def Let_def)
-
-end
-
-instantiation fps :: (field) ring_div
-begin
-
-definition fps_mod_def:
-  "f mod g = (if g = 0 then f else
-     let n = subdegree g; h = fps_shift n g
-     in  fps_cutoff n (f * inverse h) * h)"
-
-lemma fps_mod_eq_zero:
-  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree g"
-  shows   "f mod g = 0"
-  using assms by (cases "f = 0") (auto simp: fps_cutoff_zero_iff fps_mod_def Let_def)
-
-lemma fps_times_divide_eq:
-  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree (g :: 'a fps)"
-  shows   "f div g * g = f"
-proof (cases "f = 0")
-  assume nz: "f \<noteq> 0"
-  define n where "n = subdegree g"
-  define h where "h = fps_shift n g"
-  from assms have [simp]: "h $ 0 \<noteq> 0" unfolding h_def by (simp add: n_def)
-
-  from assms nz have "f div g * g = fps_shift n (f * inverse h) * g"
-    by (simp add: fps_divide_def Let_def h_def n_def)
-  also have "... = fps_shift n (f * inverse h) * X^n * h" unfolding h_def n_def
-    by (subst subdegree_decompose[of g]) simp
-  also have "fps_shift n (f * inverse h) * X^n = f * inverse h"
-    by (rule fps_shift_times_X_power) (simp_all add: nz assms n_def)
-  also have "... * h = f * (inverse h * h)" by simp
-  also have "inverse h * h = 1" by (rule inverse_mult_eq_1) simp
-  finally show ?thesis by simp
-qed (simp_all add: fps_divide_def Let_def)
-
-lemma
-  assumes "g$0 \<noteq> 0"
-  shows   fps_divide_unit: "f div g = f * inverse g" and fps_mod_unit [simp]: "f mod g = 0"
-proof -
-  from assms have [simp]: "subdegree g = 0" by (simp add: subdegree_eq_0_iff)
-  from assms show "f div g = f * inverse g"
-    by (auto simp: fps_divide_def Let_def subdegree_eq_0_iff)
-  from assms show "f mod g = 0" by (intro fps_mod_eq_zero) auto
-qed
-
-context
-begin
-private lemma fps_divide_cancel_aux1:
-  assumes "h$0 \<noteq> (0 :: 'a :: field)"
-  shows   "(h * f) div (h * g) = f div g"
-proof (cases "g = 0")
-  assume "g \<noteq> 0"
-  from assms have "h \<noteq> 0" by auto
-  note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
-  from assms have [simp]: "subdegree h = 0" by (simp add: subdegree_eq_0_iff)
-
-  have "(h * f) div (h * g) =
-          fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
-    by (simp add: fps_divide_def Let_def)
-  also have "h * f * inverse (fps_shift (subdegree g) (h*g)) =
-               (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
-    by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
-  also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
-  finally show "(h * f) div (h * g) = f div g" by (simp_all add: fps_divide_def Let_def)
-qed (simp_all add: fps_divide_def)
-
-private lemma fps_divide_cancel_aux2:
-  "(f * X^m) div (g * X^m) = f div (g :: 'a :: field fps)"
-proof (cases "g = 0")
-  assume [simp]: "g \<noteq> 0"
-  have "(f * X^m) div (g * X^m) =
-          fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*X^m))*X^m)"
-    by (simp add: fps_divide_def Let_def algebra_simps)
-  also have "... = f div g"
-    by (simp add: fps_shift_times_X_power'' fps_divide_def Let_def)
-  finally show ?thesis .
-qed (simp_all add: fps_divide_def)
-
-instance proof
-  fix f g :: "'a fps"
-  define n where "n = subdegree g"
-  define h where "h = fps_shift n g"
-
-  show "f div g * g + f mod g = f"
-  proof (cases "g = 0 \<or> f = 0")
-    assume "\<not>(g = 0 \<or> f = 0)"
-    hence nz [simp]: "f \<noteq> 0" "g \<noteq> 0" by simp_all
-    show ?thesis
-    proof (rule disjE[OF le_less_linear])
-      assume "subdegree f \<ge> subdegree g"
-      with nz show ?thesis by (simp add: fps_mod_eq_zero fps_times_divide_eq)
-    next
-      assume "subdegree f < subdegree g"
-      have g_decomp: "g = h * X^n" unfolding h_def n_def by (rule subdegree_decompose)
-      have "f div g * g + f mod g =
-              fps_shift n (f * inverse h) * g + fps_cutoff n (f * inverse h) * h"
-        by (simp add: fps_mod_def fps_divide_def Let_def n_def h_def)
-      also have "... = h * (fps_shift n (f * inverse h) * X^n + fps_cutoff n (f * inverse h))"
-        by (subst g_decomp) (simp add: algebra_simps)
-      also have "... = f * (inverse h * h)"
-        by (subst fps_shift_cutoff) simp
-      also have "inverse h * h = 1" by (rule inverse_mult_eq_1) (simp add: h_def n_def)
-      finally show ?thesis by simp
-    qed
-  qed (auto simp: fps_mod_def fps_divide_def Let_def)
-next
-
-  fix f g h :: "'a fps"
-  assume "h \<noteq> 0"
-  show "(h * f) div (h * g) = f div g"
-  proof -
-    define m where "m = subdegree h"
-    define h' where "h' = fps_shift m h"
-    have h_decomp: "h = h' * X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
-    from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
-    have "(h * f) div (h * g) = (h' * f * X^m) div (h' * g * X^m)"
-      by (simp add: h_decomp algebra_simps)
-    also have "... = f div g" by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
-    finally show ?thesis .
-  qed
-
-next
-  fix f g h :: "'a fps"
-  assume [simp]: "h \<noteq> 0"
-  define n h' where dfs: "n = subdegree h" "h' = fps_shift n h"
-  have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
-    by (simp add: fps_divide_def Let_def dfs[symmetric] algebra_simps fps_shift_add)
-  also have "h * inverse h' = (inverse h' * h') * X^n"
-    by (subst subdegree_decompose) (simp_all add: dfs)
-  also have "... = X^n" by (subst inverse_mult_eq_1) (simp_all add: dfs)
-  also have "fps_shift n (g * X^n) = g" by simp
-  also have "fps_shift n (f * inverse h') = f div h"
-    by (simp add: fps_divide_def Let_def dfs)
-  finally show "(f + g * h) div h = g + f div h" by simp
-qed
-
-end
-end
-
-lemma subdegree_mod:
-  assumes "f \<noteq> 0" "subdegree f < subdegree g"
-  shows   "subdegree (f mod g) = subdegree f"
-proof (cases "f div g * g = 0")
-  assume "f div g * g \<noteq> 0"
-  hence [simp]: "f div g \<noteq> 0" "g \<noteq> 0" by auto
-  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
-  also from assms have "subdegree ... = subdegree f"
-    by (intro subdegree_diff_eq1) simp_all
-  finally show ?thesis .
-next
-  assume zero: "f div g * g = 0"
-  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
-  also note zero
-  finally show ?thesis by simp
-qed
-
-lemma fps_divide_nth_0 [simp]: "g $ 0 \<noteq> 0 \<Longrightarrow> (f div g) $ 0 = f $ 0 / (g $ 0 :: _ :: field)"
-  by (simp add: fps_divide_unit divide_inverse)
-
-
-lemma dvd_imp_subdegree_le:
-  "(f :: 'a :: idom fps) dvd g \<Longrightarrow> g \<noteq> 0 \<Longrightarrow> subdegree f \<le> subdegree g"
-  by (auto elim: dvdE)
-
-lemma fps_dvd_iff:
-  assumes "(f :: 'a :: field fps) \<noteq> 0" "g \<noteq> 0"
-  shows   "f dvd g \<longleftrightarrow> subdegree f \<le> subdegree g"
-proof
-  assume "subdegree f \<le> subdegree g"
-  with assms have "g mod f = 0"
-    by (simp add: fps_mod_def Let_def fps_cutoff_zero_iff)
-  thus "f dvd g" by (simp add: dvd_eq_mod_eq_0)
-qed (simp add: assms dvd_imp_subdegree_le)
-
-lemma fps_shift_altdef:
-  "fps_shift n f = (f :: 'a :: field fps) div X^n"
-  by (simp add: fps_divide_def)
-  
-lemma fps_div_X_power_nth: "((f :: 'a :: field fps) div X^n) $ k = f $ (k + n)"
-  by (simp add: fps_shift_altdef [symmetric])
-
-lemma fps_div_X_nth: "((f :: 'a :: field fps) div X) $ k = f $ Suc k"
-  using fps_div_X_power_nth[of f 1] by simp
-
-lemma fps_const_inverse: "inverse (fps_const (a::'a::field)) = fps_const (inverse a)"
-  by (cases "a \<noteq> 0", rule fps_inverse_unique) (auto simp: fps_eq_iff)
-
-lemma fps_const_divide: "fps_const (x :: _ :: field) / fps_const y = fps_const (x / y)"
-  by (cases "y = 0") (simp_all add: fps_divide_unit fps_const_inverse divide_inverse)
-
-lemma inverse_fps_numeral:
-  "inverse (numeral n :: ('a :: field_char_0) fps) = fps_const (inverse (numeral n))"
-  by (intro fps_inverse_unique fps_ext) (simp_all add: fps_numeral_nth)
-
-lemma fps_numeral_divide_divide:
-  "x / numeral b / numeral c = (x / numeral (b * c) :: 'a :: field fps)"
-  by (cases "numeral b = (0::'a)"; cases "numeral c = (0::'a)")
-      (simp_all add: fps_divide_unit fps_inverse_mult [symmetric] numeral_fps_const numeral_mult 
-                del: numeral_mult [symmetric])
-
-lemma fps_numeral_mult_divide:
-  "numeral b * x / numeral c = (numeral b / numeral c * x :: 'a :: field fps)"
-  by (cases "numeral c = (0::'a)") (simp_all add: fps_divide_unit numeral_fps_const)
-
-lemmas fps_numeral_simps = 
-  fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const
-
-
-subsection \<open>Formal power series form a Euclidean ring\<close>
-
-instantiation fps :: (field) euclidean_ring_cancel
-begin
-
-definition fps_euclidean_size_def:
-  "euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)"
-
-instance proof
-  fix f g :: "'a fps" assume [simp]: "g \<noteq> 0"
-  show "euclidean_size f \<le> euclidean_size (f * g)"
-    by (cases "f = 0") (auto simp: fps_euclidean_size_def)
-  show "euclidean_size (f mod g) < euclidean_size g"
-    apply (cases "f = 0", simp add: fps_euclidean_size_def)
-    apply (rule disjE[OF le_less_linear[of "subdegree g" "subdegree f"]])
-    apply (simp_all add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod)
-    done
-qed (simp_all add: fps_euclidean_size_def)
-
-end
-
-instantiation fps :: (field) euclidean_ring_gcd
-begin
-definition fps_gcd_def: "(gcd :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.gcd"
-definition fps_lcm_def: "(lcm :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.lcm"
-definition fps_Gcd_def: "(Gcd :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Gcd"
-definition fps_Lcm_def: "(Lcm :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Lcm"
-instance by standard (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def)
-end
-
-lemma fps_gcd:
-  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
-  shows   "gcd f g = X ^ min (subdegree f) (subdegree g)"
-proof -
-  let ?m = "min (subdegree f) (subdegree g)"
-  show "gcd f g = X ^ ?m"
-  proof (rule sym, rule gcdI)
-    fix d assume "d dvd f" "d dvd g"
-    thus "d dvd X ^ ?m" by (cases "d = 0") (auto simp: fps_dvd_iff)
-  qed (simp_all add: fps_dvd_iff)
-qed
-
-lemma fps_gcd_altdef: "gcd (f :: 'a :: field fps) g =
-  (if f = 0 \<and> g = 0 then 0 else
-   if f = 0 then X ^ subdegree g else
-   if g = 0 then X ^ subdegree f else
-     X ^ min (subdegree f) (subdegree g))"
-  by (simp add: fps_gcd)
-
-lemma fps_lcm:
-  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
-  shows   "lcm f g = X ^ max (subdegree f) (subdegree g)"
-proof -
-  let ?m = "max (subdegree f) (subdegree g)"
-  show "lcm f g = X ^ ?m"
-  proof (rule sym, rule lcmI)
-    fix d assume "f dvd d" "g dvd d"
-    thus "X ^ ?m dvd d" by (cases "d = 0") (auto simp: fps_dvd_iff)
-  qed (simp_all add: fps_dvd_iff)
-qed
-
-lemma fps_lcm_altdef: "lcm (f :: 'a :: field fps) g =
-  (if f = 0 \<or> g = 0 then 0 else X ^ max (subdegree f) (subdegree g))"
-  by (simp add: fps_lcm)
-
-lemma fps_Gcd:
-  assumes "A - {0} \<noteq> {}"
-  shows   "Gcd A = X ^ (INF f:A-{0}. subdegree f)"
-proof (rule sym, rule GcdI)
-  fix f assume "f \<in> A"
-  thus "X ^ (INF f:A - {0}. subdegree f) dvd f"
-    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cINF_lower)
-next
-  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> d dvd f"
-  from assms obtain f where "f \<in> A - {0}" by auto
-  with d[of f] have [simp]: "d \<noteq> 0" by auto
-  from d assms have "subdegree d \<le> (INF f:A-{0}. subdegree f)"
-    by (intro cINF_greatest) (auto simp: fps_dvd_iff[symmetric])
-  with d assms show "d dvd X ^ (INF f:A-{0}. subdegree f)" by (simp add: fps_dvd_iff)
-qed simp_all
-
-lemma fps_Gcd_altdef: "Gcd (A :: 'a :: field fps set) =
-  (if A \<subseteq> {0} then 0 else X ^ (INF f:A-{0}. subdegree f))"
-  using fps_Gcd by auto
-
-lemma fps_Lcm:
-  assumes "A \<noteq> {}" "0 \<notin> A" "bdd_above (subdegree`A)"
-  shows   "Lcm A = X ^ (SUP f:A. subdegree f)"
-proof (rule sym, rule LcmI)
-  fix f assume "f \<in> A"
-  moreover from assms(3) have "bdd_above (subdegree ` A)" by auto
-  ultimately show "f dvd X ^ (SUP f:A. subdegree f)" using assms(2)
-    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cSUP_upper)
-next
-  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> f dvd d"
-  from assms obtain f where f: "f \<in> A" "f \<noteq> 0" by auto
-  show "X ^ (SUP f:A. subdegree f) dvd d"
-  proof (cases "d = 0")
-    assume "d \<noteq> 0"
-    moreover from d have "\<And>f. f \<in> A \<Longrightarrow> f \<noteq> 0 \<Longrightarrow> f dvd d" by blast
-    ultimately have "subdegree d \<ge> (SUP f:A. subdegree f)" using assms
-      by (intro cSUP_least) (auto simp: fps_dvd_iff)
-    with \<open>d \<noteq> 0\<close> show ?thesis by (simp add: fps_dvd_iff)
-  qed simp_all
-qed simp_all
-
-lemma fps_Lcm_altdef:
-  "Lcm (A :: 'a :: field fps set) =
-     (if 0 \<in> A \<or> \<not>bdd_above (subdegree`A) then 0 else
-      if A = {} then 1 else X ^ (SUP f:A. subdegree f))"
-proof (cases "bdd_above (subdegree`A)")
-  assume unbounded: "\<not>bdd_above (subdegree`A)"
-  have "Lcm A = 0"
-  proof (rule ccontr)
-    assume "Lcm A \<noteq> 0"
-    from unbounded obtain f where f: "f \<in> A" "subdegree (Lcm A) < subdegree f"
-      unfolding bdd_above_def by (auto simp: not_le)
-    moreover from f and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
-      by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all
-    ultimately show False by simp
-  qed
-  with unbounded show ?thesis by simp
-qed (simp_all add: fps_Lcm Lcm_eq_0_I)
-
-
-
-subsection \<open>Formal Derivatives, and the MacLaurin theorem around 0\<close>
-
-definition "fps_deriv f = Abs_fps (\<lambda>n. of_nat (n + 1) * f $ (n + 1))"
-
-lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n + 1)"
-  by (simp add: fps_deriv_def)
-
-lemma fps_0th_higher_deriv: 
-  "(fps_deriv ^^ n) f $ 0 = (fact n * f $ n :: 'a :: {comm_ring_1, semiring_char_0})"
-  by (induction n arbitrary: f) (simp_all del: funpow.simps add: funpow_Suc_right algebra_simps)
-
-lemma fps_deriv_linear[simp]:
-  "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
-    fps_const a * fps_deriv f + fps_const b * fps_deriv g"
-  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
-
-lemma fps_deriv_mult[simp]:
-  fixes f :: "'a::comm_ring_1 fps"
-  shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
-proof -
-  let ?D = "fps_deriv"
-  have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" for n
-  proof -
-    let ?Zn = "{0 ..n}"
-    let ?Zn1 = "{0 .. n + 1}"
-    let ?g = "\<lambda>i. of_nat (i+1) * g $ (i+1) * f $ (n - i) +
-        of_nat (i+1)* f $ (i+1) * g $ (n - i)"
-    let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
-        of_nat i* f $ i * g $ ((n + 1) - i)"
-    have s0: "sum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
-      sum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
-       by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
-    have s1: "sum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
-      sum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
-       by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
-    have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
-      by (simp only: mult.commute)
-    also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
-      by (simp add: fps_mult_nth sum.distrib[symmetric])
-    also have "\<dots> = sum ?h {0..n+1}"
-      by (rule sum.reindex_bij_witness_not_neutral
-            [where S'="{}" and T'="{0}" and j="Suc" and i="\<lambda>i. i - 1"]) auto
-    also have "\<dots> = (fps_deriv (f * g)) $ n"
-      apply (simp only: fps_deriv_nth fps_mult_nth sum.distrib)
-      unfolding s0 s1
-      unfolding sum.distrib[symmetric] sum_distrib_left
-      apply (rule sum.cong)
-      apply (auto simp add: of_nat_diff field_simps)
-      done
-    finally show ?thesis .
-  qed
-  then show ?thesis
-    unfolding fps_eq_iff by auto
-qed
-
-lemma fps_deriv_X[simp]: "fps_deriv X = 1"
-  by (simp add: fps_deriv_def X_def fps_eq_iff)
-
-lemma fps_deriv_neg[simp]:
-  "fps_deriv (- (f:: 'a::comm_ring_1 fps)) = - (fps_deriv f)"
-  by (simp add: fps_eq_iff fps_deriv_def)
-
-lemma fps_deriv_add[simp]:
-  "fps_deriv ((f:: 'a::comm_ring_1 fps) + g) = fps_deriv f + fps_deriv g"
-  using fps_deriv_linear[of 1 f 1 g] by simp
-
-lemma fps_deriv_sub[simp]:
-  "fps_deriv ((f:: 'a::comm_ring_1 fps) - g) = fps_deriv f - fps_deriv g"
-  using fps_deriv_add [of f "- g"] by simp
-
-lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
-  by (simp add: fps_ext fps_deriv_def fps_const_def)
-
-lemma fps_deriv_of_nat [simp]: "fps_deriv (of_nat n) = 0"
-  by (simp add: fps_of_nat [symmetric])
-
-lemma fps_deriv_numeral [simp]: "fps_deriv (numeral n) = 0"
-  by (simp add: numeral_fps_const)    
-
-lemma fps_deriv_mult_const_left[simp]:
-  "fps_deriv (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_deriv f"
-  by simp
-
-lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
-  by (simp add: fps_deriv_def fps_eq_iff)
-
-lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
-  by (simp add: fps_deriv_def fps_eq_iff )
-
-lemma fps_deriv_mult_const_right[simp]:
-  "fps_deriv (f * fps_const (c::'a::comm_ring_1)) = fps_deriv f * fps_const c"
-  by simp
-
-lemma fps_deriv_sum:
-  "fps_deriv (sum f S) = sum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
-proof (cases "finite S")
-  case False
-  then show ?thesis by simp
-next
-  case True
-  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
-qed
-
-lemma fps_deriv_eq_0_iff [simp]:
-  "fps_deriv f = 0 \<longleftrightarrow> f = fps_const (f$0 :: 'a::{idom,semiring_char_0})"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  show ?lhs if ?rhs
-  proof -
-    from that have "fps_deriv f = fps_deriv (fps_const (f$0))"
-      by simp
-    then show ?thesis
-      by simp
-  qed
-  show ?rhs if ?lhs
-  proof -
-    from that have "\<forall>n. (fps_deriv f)$n = 0"
-      by simp
-    then have "\<forall>n. f$(n+1) = 0"
-      by (simp del: of_nat_Suc of_nat_add One_nat_def)
-    then show ?thesis
-      apply (clarsimp simp add: fps_eq_iff fps_const_def)
-      apply (erule_tac x="n - 1" in allE)
-      apply simp
-      done
-  qed
-qed
-
-lemma fps_deriv_eq_iff:
-  fixes f :: "'a::{idom,semiring_char_0} fps"
-  shows "fps_deriv f = fps_deriv g \<longleftrightarrow> (f = fps_const(f$0 - g$0) + g)"
-proof -
-  have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0"
-    by simp
-  also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f - g) $ 0)"
-    unfolding fps_deriv_eq_0_iff ..
-  finally show ?thesis
-    by (simp add: field_simps)
-qed
-
-lemma fps_deriv_eq_iff_ex:
-  "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>c::'a::{idom,semiring_char_0}. f = fps_const c + g)"
-  by (auto simp: fps_deriv_eq_iff)
-
-
-fun fps_nth_deriv :: "nat \<Rightarrow> 'a::semiring_1 fps \<Rightarrow> 'a fps"
-where
-  "fps_nth_deriv 0 f = f"
-| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
-
-lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)"
-  by (induct n arbitrary: f) auto
-
-lemma fps_nth_deriv_linear[simp]:
-  "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
-    fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
-  by (induct n arbitrary: f g) (auto simp add: fps_nth_deriv_commute)
-
-lemma fps_nth_deriv_neg[simp]:
-  "fps_nth_deriv n (- (f :: 'a::comm_ring_1 fps)) = - (fps_nth_deriv n f)"
-  by (induct n arbitrary: f) simp_all
-
-lemma fps_nth_deriv_add[simp]:
-  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g"
-  using fps_nth_deriv_linear[of n 1 f 1 g] by simp
-
-lemma fps_nth_deriv_sub[simp]:
-  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
-  using fps_nth_deriv_add [of n f "- g"] by simp
-
-lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
-  by (induct n) simp_all
-
-lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)"
-  by (induct n) simp_all
-
-lemma fps_nth_deriv_const[simp]:
-  "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
-  by (cases n) simp_all
-
-lemma fps_nth_deriv_mult_const_left[simp]:
-  "fps_nth_deriv n (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_nth_deriv n f"
-  using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
-
-lemma fps_nth_deriv_mult_const_right[simp]:
-  "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
-  using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult.commute)
-
-lemma fps_nth_deriv_sum:
-  "fps_nth_deriv n (sum f S) = sum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
-proof (cases "finite S")
-  case True
-  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
-next
-  case False
-  then show ?thesis by simp
-qed
-
-lemma fps_deriv_maclauren_0:
-  "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k"
-  by (induct k arbitrary: f) (auto simp add: field_simps)
-
-
-subsection \<open>Powers\<close>
-
-lemma fps_power_zeroth_eq_one: "a$0 =1 \<Longrightarrow> a^n $ 0 = (1::'a::semiring_1)"
-  by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
-
-lemma fps_power_first_eq: "(a :: 'a::comm_ring_1 fps) $ 0 =1 \<Longrightarrow> a^n $ 1 = of_nat n * a$1"
-proof (induct n)
-  case 0
-  then show ?case by simp
-next
-  case (Suc n)
-  show ?case unfolding power_Suc fps_mult_nth
-    using Suc.hyps[OF \<open>a$0 = 1\<close>] \<open>a$0 = 1\<close> fps_power_zeroth_eq_one[OF \<open>a$0=1\<close>]
-    by (simp add: field_simps)
-qed
-
-lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
-  by (induct n) (auto simp add: fps_mult_nth)
-
-lemma startsby_zero_power:"a $0 = (0::'a::comm_ring_1) \<Longrightarrow> n > 0 \<Longrightarrow> a^n $0 = 0"
-  by (induct n) (auto simp add: fps_mult_nth)
-
-lemma startsby_power:"a $0 = (v::'a::comm_ring_1) \<Longrightarrow> a^n $0 = v^n"
-  by (induct n) (auto simp add: fps_mult_nth)
-
-lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::idom) \<longleftrightarrow> n \<noteq> 0 \<and> a$0 = 0"
-  apply (rule iffI)
-  apply (induct n)
-  apply (auto simp add: fps_mult_nth)
-  apply (rule startsby_zero_power, simp_all)
-  done
-
-lemma startsby_zero_power_prefix:
-  assumes a0: "a $ 0 = (0::'a::idom)"
-  shows "\<forall>n < k. a ^ k $ n = 0"
-  using a0
-proof (induct k rule: nat_less_induct)
-  fix k
-  assume H: "\<forall>m<k. a $0 =  0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $ 0 = 0"
-  show "\<forall>m<k. a ^ k $ m = 0"
-  proof (cases k)
-    case 0
-    then show ?thesis by simp
-  next
-    case (Suc l)
-    have "a^k $ m = 0" if mk: "m < k" for m
-    proof (cases "m = 0")
-      case True
-      then show ?thesis
-        using startsby_zero_power[of a k] Suc a0 by simp
-    next
-      case False
-      have "a ^k $ m = (a^l * a) $m"
-        by (simp add: Suc mult.commute)
-      also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
-        by (simp add: fps_mult_nth)
-      also have "\<dots> = 0"
-        apply (rule sum.neutral)
-        apply auto
-        apply (case_tac "x = m")
-        using a0 apply simp
-        apply (rule H[rule_format])
-        using a0 Suc mk apply auto
-        done
-      finally show ?thesis .
-    qed
-    then show ?thesis by blast
-  qed
-qed
-
-lemma startsby_zero_sum_depends:
-  assumes a0: "a $0 = (0::'a::idom)"
-    and kn: "n \<ge> k"
-  shows "sum (\<lambda>i. (a ^ i)$k) {0 .. n} = sum (\<lambda>i. (a ^ i)$k) {0 .. k}"
-  apply (rule sum.mono_neutral_right)
-  using kn
-  apply auto
-  apply (rule startsby_zero_power_prefix[rule_format, OF a0])
-  apply arith
-  done
-
-lemma startsby_zero_power_nth_same:
-  assumes a0: "a$0 = (0::'a::idom)"
-  shows "a^n $ n = (a$1) ^ n"
-proof (induct n)
-  case 0
-  then show ?case by simp
-next
-  case (Suc n)
-  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)"
-    by (simp add: field_simps)
-  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
-    by (simp add: fps_mult_nth)
-  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
-    apply (rule sum.mono_neutral_right)
-    apply simp
-    apply clarsimp
-    apply clarsimp
-    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
-    apply arith
-    done
-  also have "\<dots> = a^n $ n * a$1"
-    using a0 by simp
-  finally show ?case
-    using Suc.hyps by simp
-qed
-
-lemma fps_inverse_power:
-  fixes a :: "'a::field fps"
-  shows "inverse (a^n) = inverse a ^ n"
-  by (induction n) (simp_all add: fps_inverse_mult)
-
-lemma fps_deriv_power:
-  "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a::comm_ring_1) * fps_deriv a * a ^ (n - 1)"
-  apply (induct n)
-  apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add)
-  apply (case_tac n)
-  apply (auto simp add: field_simps)
-  done
-
-lemma fps_inverse_deriv:
-  fixes a :: "'a::field fps"
-  assumes a0: "a$0 \<noteq> 0"
-  shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
-proof -
-  from inverse_mult_eq_1[OF a0]
-  have "fps_deriv (inverse a * a) = 0" by simp
-  then have "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0"
-    by simp
-  then have "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0"
-    by simp
-  with inverse_mult_eq_1[OF a0]
-  have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) = 0"
-    unfolding power2_eq_square
-    apply (simp add: field_simps)
-    apply (simp add: mult.assoc[symmetric])
-    done
-  then have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * (inverse a)\<^sup>2 =
-      0 - fps_deriv a * (inverse a)\<^sup>2"
-    by simp
-  then show "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
-    by (simp add: field_simps)
-qed
-
-lemma fps_inverse_deriv':
-  fixes a :: "'a::field fps"
-  assumes a0: "a $ 0 \<noteq> 0"
-  shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2"
-  using fps_inverse_deriv[OF a0] a0
-  by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult)
-
-lemma inverse_mult_eq_1':
-  assumes f0: "f$0 \<noteq> (0::'a::field)"
-  shows "f * inverse f = 1"
-  by (metis mult.commute inverse_mult_eq_1 f0)
-
-lemma fps_inverse_minus [simp]: "inverse (-f) = -inverse (f :: 'a :: field fps)"
-  by (cases "f$0 = 0") (auto intro: fps_inverse_unique simp: inverse_mult_eq_1' fps_inverse_eq_0)
-  
-lemma divide_fps_const [simp]: "f / fps_const (c :: 'a :: field) = fps_const (inverse c) * f"
-  by (cases "c = 0") (simp_all add: fps_divide_unit fps_const_inverse)
-
-(* FIXME: The last part of this proof should go through by simp once we have a proper
-   theorem collection for simplifying division on rings *)
-lemma fps_divide_deriv:
-  assumes "b dvd (a :: 'a :: field fps)"
-  shows   "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b^2"
-proof -
-  have eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b div c" for a b c :: "'a :: field fps"
-    by (drule sym) (simp add: mult.assoc)
-  from assms have "a = a / b * b" by simp
-  also have "fps_deriv (a / b * b) = fps_deriv (a / b) * b + a / b * fps_deriv b" by simp
-  finally have "fps_deriv (a / b) * b^2 = fps_deriv a * b - a * fps_deriv b" using assms
-    by (simp add: power2_eq_square algebra_simps)
-  thus ?thesis by (cases "b = 0") (auto simp: eq_divide_imp)
-qed
-
-lemma fps_inverse_gp': "inverse (Abs_fps (\<lambda>n. 1::'a::field)) = 1 - X"
-  by (simp add: fps_inverse_gp fps_eq_iff X_def)
-
-lemma fps_one_over_one_minus_X_squared:
-  "inverse ((1 - X)^2 :: 'a :: field fps) = Abs_fps (\<lambda>n. of_nat (n+1))"
-proof -
-  have "inverse ((1 - X)^2 :: 'a fps) = fps_deriv (inverse (1 - X))"
-    by (subst fps_inverse_deriv) (simp_all add: fps_inverse_power)
-  also have "inverse (1 - X :: 'a fps) = Abs_fps (\<lambda>_. 1)"
-    by (subst fps_inverse_gp' [symmetric]) simp
-  also have "fps_deriv \<dots> = Abs_fps (\<lambda>n. of_nat (n + 1))"
-    by (simp add: fps_deriv_def)
-  finally show ?thesis .
-qed
-
-lemma fps_nth_deriv_X[simp]: "fps_nth_deriv n X = (if n = 0 then X else if n=1 then 1 else 0)"
-  by (cases n) simp_all
-
-lemma fps_inverse_X_plus1: "inverse (1 + X) = Abs_fps (\<lambda>n. (- (1::'a::field)) ^ n)"
-  (is "_ = ?r")
-proof -
-  have eq: "(1 + X) * ?r = 1"
-    unfolding minus_one_power_iff
-    by (auto simp add: field_simps fps_eq_iff)
-  show ?thesis
-    by (auto simp add: eq intro: fps_inverse_unique)
-qed
-
-
-subsection \<open>Integration\<close>
-
-definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
-  where "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
-
-lemma fps_deriv_fps_integral: "fps_deriv (fps_integral a a0) = a"
-  unfolding fps_integral_def fps_deriv_def
-  by (simp add: fps_eq_iff del: of_nat_Suc)
-
-lemma fps_integral_linear:
-  "fps_integral (fps_const a * f + fps_const b * g) (a*a0 + b*b0) =
-    fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0"
-  (is "?l = ?r")
-proof -
-  have "fps_deriv ?l = fps_deriv ?r"
-    by (simp add: fps_deriv_fps_integral)
-  moreover have "?l$0 = ?r$0"
-    by (simp add: fps_integral_def)
-  ultimately show ?thesis
-    unfolding fps_deriv_eq_iff by auto
-qed
-
-
-subsection \<open>Composition of FPSs\<close>
-
-definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps"  (infixl "oo" 55)
-  where "a oo b = Abs_fps (\<lambda>n. sum (\<lambda>i. a$i * (b^i$n)) {0..n})"
-
-lemma fps_compose_nth: "(a oo b)$n = sum (\<lambda>i. a$i * (b^i$n)) {0..n}"
-  by (simp add: fps_compose_def)
-
-lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0"
-  by (simp add: fps_compose_nth)
-
-lemma fps_compose_X[simp]: "a oo X = (a :: 'a::comm_ring_1 fps)"
-  by (simp add: fps_ext fps_compose_def mult_delta_right sum.delta')
-
-lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
-  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
-
-lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k"
-  unfolding numeral_fps_const by simp
-
-lemma neg_numeral_compose[simp]: "(- numeral k :: 'a::comm_ring_1 fps) oo b = - numeral k"
-  unfolding neg_numeral_fps_const by simp
-
-lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: 'a::comm_ring_1 fps)"
-  by (simp add: fps_eq_iff fps_compose_def mult_delta_left sum.delta not_le)
-
-
-subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
-
-subsubsection \<open>Rule 1\<close>
-  (* {a_{n+k}}_0^infty Corresponds to (f - sum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
-
-lemma fps_power_mult_eq_shift:
-  "X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
-    Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * X^i) {0 .. k}"
-  (is "?lhs = ?rhs")
-proof -
-  have "?lhs $ n = ?rhs $ n" for n :: nat
-  proof -
-    have "?lhs $ n = (if n < Suc k then 0 else a n)"
-      unfolding X_power_mult_nth by auto
-    also have "\<dots> = ?rhs $ n"
-    proof (induct k)
-      case 0
-      then show ?case
-        by (simp add: fps_sum_nth)
-    next
-      case (Suc k)
-      have "(Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
-        (Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} -
-          fps_const (a (Suc k)) * X^ Suc k) $ n"
-        by (simp add: field_simps)
-      also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
-        using Suc.hyps[symmetric] unfolding fps_sub_nth by simp
-      also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
-        unfolding X_power_mult_right_nth
-        apply (auto simp add: not_less fps_const_def)
-        apply (rule cong[of a a, OF refl])
-        apply arith
-        done
-      finally show ?case
-        by simp
-    qed
-    finally show ?thesis .
-  qed
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-
-subsubsection \<open>Rule 2\<close>
-
-  (* We can not reach the form of Wilf, but still near to it using rewrite rules*)
-  (* If f reprents {a_n} and P is a polynomial, then
-        P(xD) f represents {P(n) a_n}*)
-
-definition "XD = op * X \<circ> fps_deriv"
-
-lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: 'a::comm_ring_1 fps)"
-  by (simp add: XD_def field_simps)
-
-lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
-  by (simp add: XD_def field_simps)
-
-lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) =
-    fps_const c * XD a + fps_const d * XD (b :: 'a::comm_ring_1 fps)"
-  by simp
-
-lemma XDN_linear:
-  "(XD ^^ n) (fps_const c * a + fps_const d * b) =
-    fps_const c * (XD ^^ n) a + fps_const d * (XD ^^ n) (b :: 'a::comm_ring_1 fps)"
-  by (induct n) simp_all
-
-lemma fps_mult_X_deriv_shift: "X* fps_deriv a = Abs_fps (\<lambda>n. of_nat n* a$n)"
-  by (simp add: fps_eq_iff)
-
-lemma fps_mult_XD_shift:
-  "(XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
-  by (induct k arbitrary: a) (simp_all add: XD_def fps_eq_iff field_simps del: One_nat_def)
-
-
-subsubsection \<open>Rule 3\<close>
-
-text \<open>Rule 3 is trivial and is given by \<open>fps_times_def\<close>.\<close>
-
-
-subsubsection \<open>Rule 5 --- summation and "division" by (1 - X)\<close>
-
-lemma fps_divide_X_minus1_sum_lemma:
-  "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
-proof -
-  let ?sa = "Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
-  have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)"
-    by simp
-  have "a$n = ((1 - X) * ?sa) $ n" for n
-  proof (cases "n = 0")
-    case True
-    then show ?thesis
-      by (simp add: fps_mult_nth)
-  next
-    case False
-    then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1} \<union> {2..n} = {1..n}"
-      "{0..n - 1} \<union> {n} = {0..n}"
-      by (auto simp: set_eq_iff)
-    have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}" "{0..n - 1} \<inter> {n} = {}"
-      using False by simp_all
-    have f: "finite {0}" "finite {1}" "finite {2 .. n}"
-      "finite {0 .. n - 1}" "finite {n}" by simp_all
-    have "((1 - X) * ?sa) $ n = sum (\<lambda>i. (1 - X)$ i * ?sa $ (n - i)) {0 .. n}"
-      by (simp add: fps_mult_nth)
-    also have "\<dots> = a$n"
-      unfolding th0
-      unfolding sum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
-      unfolding sum.union_disjoint[OF f(2) f(3) d(2)]
-      apply (simp)
-      unfolding sum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
-      apply simp
-      done
-    finally show ?thesis
-      by simp
-  qed
-  then show ?thesis
-    unfolding fps_eq_iff by blast
-qed
-
-lemma fps_divide_X_minus1_sum:
-  "a /((1::'a::field fps) - X) = Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
-proof -
-  let ?X = "1 - (X::'a fps)"
-  have th0: "?X $ 0 \<noteq> 0"
-    by simp
-  have "a /?X = ?X *  Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) * inverse ?X"
-    using fps_divide_X_minus1_sum_lemma[of a, symmetric] th0
-    by (simp add: fps_divide_def mult.assoc)
-  also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) "
-    by (simp add: ac_simps)
-  finally show ?thesis
-    by (simp add: inverse_mult_eq_1[OF th0])
-qed
-
-
-subsubsection \<open>Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
-  finite product of FPS, also the relvant instance of powers of a FPS\<close>
-
-definition "natpermute n k = {l :: nat list. length l = k \<and> sum_list l = n}"
-
-lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
-  apply (auto simp add: natpermute_def)
-  apply (case_tac x)
-  apply auto
-  done
-
-lemma append_natpermute_less_eq:
-  assumes "xs @ ys \<in> natpermute n k"
-  shows "sum_list xs \<le> n"
-    and "sum_list ys \<le> n"
-proof -
-  from assms have "sum_list (xs @ ys) = n"
-    by (simp add: natpermute_def)
-  then have "sum_list xs + sum_list ys = n"
-    by simp
-  then show "sum_list xs \<le> n" and "sum_list ys \<le> n"
-    by simp_all
-qed
-
-lemma natpermute_split:
-  assumes "h \<le> k"
-  shows "natpermute n k =
-    (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})"
-  (is "?L = ?R" is "_ = (\<Union>m \<in>{0..n}. ?S m)")
-proof
-  show "?R \<subseteq> ?L"
-  proof
-    fix l
-    assume l: "l \<in> ?R"
-    from l obtain m xs ys where h: "m \<in> {0..n}"
-      and xs: "xs \<in> natpermute m h"
-      and ys: "ys \<in> natpermute (n - m) (k - h)"
-      and leq: "l = xs@ys" by blast
-    from xs have xs': "sum_list xs = m"
-      by (simp add: natpermute_def)
-    from ys have ys': "sum_list ys = n - m"
-      by (simp add: natpermute_def)
-    show "l \<in> ?L" using leq xs ys h
-      apply (clarsimp simp add: natpermute_def)
-      unfolding xs' ys'
-      using assms xs ys
-      unfolding natpermute_def
-      apply simp
-      done
-  qed
-  show "?L \<subseteq> ?R"
-  proof
-    fix l
-    assume l: "l \<in> natpermute n k"
-    let ?xs = "take h l"
-    let ?ys = "drop h l"
-    let ?m = "sum_list ?xs"
-    from l have ls: "sum_list (?xs @ ?ys) = n"
-      by (simp add: natpermute_def)
-    have xs: "?xs \<in> natpermute ?m h" using l assms
-      by (simp add: natpermute_def)
-    have l_take_drop: "sum_list l = sum_list (take h l @ drop h l)"
-      by simp
-    then have ys: "?ys \<in> natpermute (n - ?m) (k - h)"
-      using l assms ls by (auto simp add: natpermute_def simp del: append_take_drop_id)
-    from ls have m: "?m \<in> {0..n}"
-      by (simp add: l_take_drop del: append_take_drop_id)
-    from xs ys ls show "l \<in> ?R"
-      apply auto
-      apply (rule bexI [where x = "?m"])
-      apply (rule exI [where x = "?xs"])
-      apply (rule exI [where x = "?ys"])
-      using ls l
-      apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id)
-      apply simp
-      done
-  qed
-qed
-
-lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
-  by (auto simp add: natpermute_def)
-
-lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})"
-  apply (auto simp add: set_replicate_conv_if natpermute_def)
-  apply (rule nth_equalityI)
-  apply simp_all
-  done
-
-lemma natpermute_finite: "finite (natpermute n k)"
-proof (induct k arbitrary: n)
-  case 0
-  then show ?case
-    apply (subst natpermute_split[of 0 0, simplified])
-    apply (simp add: natpermute_0)
-    done
-next
-  case (Suc k)
-  then show ?case unfolding natpermute_split [of k "Suc k", simplified]
-    apply -
-    apply (rule finite_UN_I)
-    apply simp
-    unfolding One_nat_def[symmetric] natlist_trivial_1
-    apply simp
-    done
-qed
-
-lemma natpermute_contain_maximal:
-  "{xs \<in> natpermute n (k + 1). n \<in> set xs} = (\<Union>i\<in>{0 .. k}. {(replicate (k + 1) 0) [i:=n]})"
-  (is "?A = ?B")
-proof
-  show "?A \<subseteq> ?B"
-  proof
-    fix xs
-    assume "xs \<in> ?A"
-    then have H: "xs \<in> natpermute n (k + 1)" and n: "n \<in> set xs"
-      by blast+
-    then obtain i where i: "i \<in> {0.. k}" "xs!i = n"
-      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
-    have eqs: "({0..k} - {i}) \<union> {i} = {0..k}"
-      using i by auto
-    have f: "finite({0..k} - {i})" "finite {i}"
-      by auto
-    have d: "({0..k} - {i}) \<inter> {i} = {}"
-      using i by auto
-    from H have "n = sum (nth xs) {0..k}"
-      apply (simp add: natpermute_def)
-      apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth)
-      done
-    also have "\<dots> = n + sum (nth xs) ({0..k} - {i})"
-      unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp
-    finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
-      by auto
-    from H have xsl: "length xs = k+1"
-      by (simp add: natpermute_def)
-    from i have i': "i < length (replicate (k+1) 0)"   "i < k+1"
-      unfolding length_replicate by presburger+
-    have "xs = replicate (k+1) 0 [i := n]"
-      apply (rule nth_equalityI)
-      unfolding xsl length_list_update length_replicate
-      apply simp
-      apply clarify
-      unfolding nth_list_update[OF i'(1)]
-      using i zxs
-      apply (case_tac "ia = i")
-      apply (auto simp del: replicate.simps)
-      done
-    then show "xs \<in> ?B" using i by blast
-  qed
-  show "?B \<subseteq> ?A"
-  proof
-    fix xs
-    assume "xs \<in> ?B"
-    then obtain i where i: "i \<in> {0..k}" and xs: "xs = replicate (k + 1) 0 [i:=n]"
-      by auto
-    have nxs: "n \<in> set xs"
-      unfolding xs
-      apply (rule set_update_memI)
-      using i apply simp
-      done
-    have xsl: "length xs = k + 1"
-      by (simp only: xs length_replicate length_list_update)
-    have "sum_list xs = sum (nth xs) {0..<k+1}"
-      unfolding sum_list_sum_nth xsl ..
-    also have "\<dots> = sum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
-      by (rule sum.cong) (simp_all add: xs del: replicate.simps)
-    also have "\<dots> = n" using i by (simp add: sum.delta)
-    finally have "xs \<in> natpermute n (k + 1)"
-      using xsl unfolding natpermute_def mem_Collect_eq by blast
-    then show "xs \<in> ?A"
-      using nxs by blast
-  qed
-qed
-
-text \<open>The general form.\<close>
-lemma fps_prod_nth:
-  fixes m :: nat
-    and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
-  shows "(prod a {0 .. m}) $ n =
-    sum (\<lambda>v. prod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
-  (is "?P m n")
-proof (induct m arbitrary: n rule: nat_less_induct)
-  fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
-  show "?P m n"
-  proof (cases m)
-    case 0
-    then show ?thesis
-      apply simp
-      unfolding natlist_trivial_1[where n = n, unfolded One_nat_def]
-      apply simp
-      done
-  next
-    case (Suc k)
-    then have km: "k < m" by arith
-    have u0: "{0 .. k} \<union> {m} = {0..m}"
-      using Suc by (simp add: set_eq_iff) presburger
-    have f0: "finite {0 .. k}" "finite {m}" by auto
-    have d0: "{0 .. k} \<inter> {m} = {}" using Suc by auto
-    have "(prod a {0 .. m}) $ n = (prod a {0 .. k} * a m) $ n"
-      unfolding prod.union_disjoint[OF f0 d0, unfolded u0] by simp
-    also have "\<dots> = (\<Sum>i = 0..n. (\<Sum>v\<in>natpermute i (k + 1). \<Prod>j\<in>{0..k}. a j $ v ! j) * a m $ (n - i))"
-      unfolding fps_mult_nth H[rule_format, OF km] ..
-    also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
-      apply (simp add: Suc)
-      unfolding natpermute_split[of m "m + 1", simplified, of n,
-        unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
-      apply (subst sum.UNION_disjoint)
-      apply simp
-      apply simp
-      unfolding image_Collect[symmetric]
-      apply clarsimp
-      apply (rule finite_imageI)
-      apply (rule natpermute_finite)
-      apply (clarsimp simp add: set_eq_iff)
-      apply auto
-      apply (rule sum.cong)
-      apply (rule refl)
-      unfolding sum_distrib_right
-      apply (rule sym)
-      apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in sum.reindex_cong)
-      apply (simp add: inj_on_def)
-      apply auto
-      unfolding prod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
-      apply (clarsimp simp add: natpermute_def nth_append)
-      done
-    finally show ?thesis .
-  qed
-qed
-
-text \<open>The special form for powers.\<close>
-lemma fps_power_nth_Suc:
-  fixes m :: nat
-    and a :: "'a::comm_ring_1 fps"
-  shows "(a ^ Suc m)$n = sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
-proof -
-  have th0: "a^Suc m = prod (\<lambda>i. a) {0..m}"
-    by (simp add: prod_constant)
-  show ?thesis unfolding th0 fps_prod_nth ..
-qed
-
-lemma fps_power_nth:
-  fixes m :: nat
-    and a :: "'a::comm_ring_1 fps"
-  shows "(a ^m)$n =
-    (if m=0 then 1$n else sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
-  by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc)
-
-lemma fps_nth_power_0:
-  fixes m :: nat
-    and a :: "'a::comm_ring_1 fps"
-  shows "(a ^m)$0 = (a$0) ^ m"
-proof (cases m)
-  case 0
-  then show ?thesis by simp
-next
-  case (Suc n)
-  then have c: "m = card {0..n}" by simp
-  have "(a ^m)$0 = prod (\<lambda>i. a$0) {0..n}"
-    by (simp add: Suc fps_power_nth del: replicate.simps power_Suc)
-  also have "\<dots> = (a$0) ^ m"
-   unfolding c by (rule prod_constant)
- finally show ?thesis .
-qed
-
-lemma natpermute_max_card:
-  assumes n0: "n \<noteq> 0"
-  shows "card {xs \<in> natpermute n (k + 1). n \<in> set xs} = k + 1"
-  unfolding natpermute_contain_maximal
-proof -
-  let ?A = "\<lambda>i. {replicate (k + 1) 0[i := n]}"
-  let ?K = "{0 ..k}"
-  have fK: "finite ?K"
-    by simp
-  have fAK: "\<forall>i\<in>?K. finite (?A i)"
-    by auto
-  have d: "\<forall>i\<in> ?K. \<forall>j\<in> ?K. i \<noteq> j \<longrightarrow>
-    {replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
-  proof clarify
-    fix i j
-    assume i: "i \<in> ?K" and j: "j \<in> ?K" and ij: "i \<noteq> j"
-    have False if eq: "replicate (k+1) 0 [i:=n] = replicate (k+1) 0 [j:= n]"
-    proof -
-      have "(replicate (k+1) 0 [i:=n] ! i) = n"
-        using i by (simp del: replicate.simps)
-      moreover
-      have "(replicate (k+1) 0 [j:=n] ! i) = 0"
-        using i ij by (simp del: replicate.simps)
-      ultimately show ?thesis
-        using eq n0 by (simp del: replicate.simps)
-    qed
-    then show "{replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
-      by auto
-  qed
-  from card_UN_disjoint[OF fK fAK d]
-  show "card (\<Union>i\<in>{0..k}. {replicate (k + 1) 0[i := n]}) = k + 1"
-    by simp
-qed
-
-lemma fps_power_Suc_nth:
-  fixes f :: "'a :: comm_ring_1 fps"
-  assumes k: "k > 0"
-  shows "(f ^ Suc m) $ k = 
-           of_nat (Suc m) * (f $ k * (f $ 0) ^ m) +
-           (\<Sum>v\<in>{v\<in>natpermute k (m+1). k \<notin> set v}. \<Prod>j = 0..m. f $ v ! j)"
-proof -
-  define A B 
-    where "A = {v\<in>natpermute k (m+1). k \<in> set v}" 
-      and  "B = {v\<in>natpermute k (m+1). k \<notin> set v}"
-  have [simp]: "finite A" "finite B" "A \<inter> B = {}" by (auto simp: A_def B_def natpermute_finite)
-
-  from natpermute_max_card[of k m] k have card_A: "card A = m + 1" by (simp add: A_def)
-  {
-    fix v assume v: "v \<in> A"
-    from v have [simp]: "length v = Suc m" by (simp add: A_def natpermute_def)
-    from v have "\<exists>j. j \<le> m \<and> v ! j = k" 
-      by (auto simp: set_conv_nth A_def natpermute_def less_Suc_eq_le)
-    then guess j by (elim exE conjE) note j = this
-    
-    from v have "k = sum_list v" by (simp add: A_def natpermute_def)
-    also have "\<dots> = (\<Sum>i=0..m. v ! i)"
-      by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum_op_ivl_Suc)
-    also from j have "{0..m} = insert j ({0..m}-{j})" by auto
-    also from j have "(\<Sum>i\<in>\<dots>. v ! i) = k + (\<Sum>i\<in>{0..m}-{j}. v ! i)"
-      by (subst sum.insert) simp_all
-    finally have "(\<Sum>i\<in>{0..m}-{j}. v ! i) = 0" by simp
-    hence zero: "v ! i = 0" if "i \<in> {0..m}-{j}" for i using that
-      by (subst (asm) sum_eq_0_iff) auto
-      
-    from j have "{0..m} = insert j ({0..m} - {j})" by auto
-    also from j have "(\<Prod>i\<in>\<dots>. f $ (v ! i)) = f $ k * (\<Prod>i\<in>{0..m} - {j}. f $ (v ! i))"
-      by (subst prod.insert) auto
-    also have "(\<Prod>i\<in>{0..m} - {j}. f $ (v ! i)) = (\<Prod>i\<in>{0..m} - {j}. f $ 0)"
-      by (intro prod.cong) (simp_all add: zero)
-    also from j have "\<dots> = (f $ 0) ^ m" by (subst prod_constant) simp_all
-    finally have "(\<Prod>j = 0..m. f $ (v ! j)) = f $ k * (f $ 0) ^ m" .
-  } note A = this
-  
-  have "(f ^ Suc m) $ k = (\<Sum>v\<in>natpermute k (m + 1). \<Prod>j = 0..m. f $ v ! j)"
-    by (rule fps_power_nth_Suc)
-  also have "natpermute k (m+1) = A \<union> B" unfolding A_def B_def by blast
-  also have "(\<Sum>v\<in>\<dots>. \<Prod>j = 0..m. f $ (v ! j)) = 
-               (\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) + (\<Sum>v\<in>B. \<Prod>j = 0..m. f $ (v ! j))"
-    by (intro sum.union_disjoint) simp_all   
-  also have "(\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)"
-    by (simp add: A card_A)
-  finally show ?thesis by (simp add: B_def)
-qed 
-  
-lemma fps_power_Suc_eqD:
-  fixes f g :: "'a :: {idom,semiring_char_0} fps"
-  assumes "f ^ Suc m = g ^ Suc m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0"
-  shows   "f = g"
-proof (rule fps_ext)
-  fix k :: nat
-  show "f $ k = g $ k"
-  proof (induction k rule: less_induct)
-    case (less k)
-    show ?case
-    proof (cases "k = 0")
-      case False
-      let ?h = "\<lambda>f. (\<Sum>v | v \<in> natpermute k (m + 1) \<and> k \<notin> set v. \<Prod>j = 0..m. f $ v ! j)"
-      from False fps_power_Suc_nth[of k f m] fps_power_Suc_nth[of k g m]
-        have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h f =
-                g $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h g" using assms 
-        by (simp add: mult_ac del: power_Suc of_nat_Suc)
-      also have "v ! i < k" if "v \<in> {v\<in>natpermute k (m+1). k \<notin> set v}" "i \<le> m" for v i
-        using that elem_le_sum_list_nat[of i v] unfolding natpermute_def
-        by (auto simp: set_conv_nth dest!: spec[of _ i])
-      hence "?h f = ?h g"
-        by (intro sum.cong refl prod.cong less lessI) (auto simp: natpermute_def)
-      finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)"
-        by simp
-      with assms show "f $ k = g $ k" 
-        by (subst (asm) mult_right_cancel) (auto simp del: of_nat_Suc)
-    qed (simp_all add: assms)
-  qed
-qed
-
-lemma fps_power_Suc_eqD':
-  fixes f g :: "'a :: {idom,semiring_char_0} fps"
-  assumes "f ^ Suc m = g ^ Suc m" "f $ subdegree f = g $ subdegree g"
-  shows   "f = g"
-proof (cases "f = 0")
-  case False
-  have "Suc m * subdegree f = subdegree (f ^ Suc m)"
-    by (rule subdegree_power [symmetric])
-  also have "f ^ Suc m = g ^ Suc m" by fact
-  also have "subdegree \<dots> = Suc m * subdegree g" by (rule subdegree_power)
-  finally have [simp]: "subdegree f = subdegree g"
-    by (subst (asm) Suc_mult_cancel1)
-  have "fps_shift (subdegree f) f * X ^ subdegree f = f"
-    by (rule subdegree_decompose [symmetric])
-  also have "\<dots> ^ Suc m = g ^ Suc m" by fact
-  also have "g = fps_shift (subdegree g) g * X ^ subdegree g"
-    by (rule subdegree_decompose)
-  also have "subdegree f = subdegree g" by fact
-  finally have "fps_shift (subdegree g) f ^ Suc m = fps_shift (subdegree g) g ^ Suc m"
-    by (simp add: algebra_simps power_mult_distrib del: power_Suc)
-  hence "fps_shift (subdegree g) f = fps_shift (subdegree g) g"
-    by (rule fps_power_Suc_eqD) (insert assms False, auto)
-  with subdegree_decompose[of f] subdegree_decompose[of g] show ?thesis by simp
-qed (insert assms, simp_all)
-
-lemma fps_power_eqD':
-  fixes f g :: "'a :: {idom,semiring_char_0} fps"
-  assumes "f ^ m = g ^ m" "f $ subdegree f = g $ subdegree g" "m > 0"
-  shows   "f = g"
-  using fps_power_Suc_eqD'[of f "m-1" g] assms by simp
-
-lemma fps_power_eqD:
-  fixes f g :: "'a :: {idom,semiring_char_0} fps"
-  assumes "f ^ m = g ^ m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0" "m > 0"
-  shows   "f = g"
-  by (rule fps_power_eqD'[of f m g]) (insert assms, simp_all)
-
-lemma fps_compose_inj_right:
-  assumes a0: "a$0 = (0::'a::idom)"
-    and a1: "a$1 \<noteq> 0"
-  shows "(b oo a = c oo a) \<longleftrightarrow> b = c"
-  (is "?lhs \<longleftrightarrow>?rhs")
-proof
-  show ?lhs if ?rhs using that by simp
-  show ?rhs if ?lhs
-  proof -
-    have "b$n = c$n" for n
-    proof (induct n rule: nat_less_induct)
-      fix n
-      assume H: "\<forall>m<n. b$m = c$m"
-      show "b$n = c$n"
-      proof (cases n)
-        case 0
-        from \<open>?lhs\<close> have "(b oo a)$n = (c oo a)$n"
-          by simp
-        then show ?thesis
-          using 0 by (simp add: fps_compose_nth)
-      next
-        case (Suc n1)
-        have f: "finite {0 .. n1}" "finite {n}" by simp_all
-        have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using Suc by auto
-        have d: "{0 .. n1} \<inter> {n} = {}" using Suc by auto
-        have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
-          apply (rule sum.cong)
-          using H Suc
-          apply auto
-          done
-        have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
-          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq
-          using startsby_zero_power_nth_same[OF a0]
-          by simp
-        have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
-          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq]
-          using startsby_zero_power_nth_same[OF a0]
-          by simp
-        from \<open>?lhs\<close>[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
-        show ?thesis by auto
-      qed
-    qed
-    then show ?rhs by (simp add: fps_eq_iff)
-  qed
-qed
-
-
-subsection \<open>Radicals\<close>
-
-declare prod.cong [fundef_cong]
-
-function radical :: "(nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a::field fps \<Rightarrow> nat \<Rightarrow> 'a"
-where
-  "radical r 0 a 0 = 1"
-| "radical r 0 a (Suc n) = 0"
-| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
-| "radical r (Suc k) a (Suc n) =
-    (a$ Suc n - sum (\<lambda>xs. prod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
-      {xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
-    (of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
-  by pat_completeness auto
-
-termination radical
-proof
-  let ?R = "measure (\<lambda>(r, k, a, n). n)"
-  {
-    show "wf ?R" by auto
-  next
-    fix r k a n xs i
-    assume xs: "xs \<in> {xs \<in> natpermute (Suc n) (Suc k). Suc n \<notin> set xs}" and i: "i \<in> {0..k}"
-    have False if c: "Suc n \<le> xs ! i"
-    proof -
-      from xs i have "xs !i \<noteq> Suc n"
-        by (auto simp add: in_set_conv_nth natpermute_def)
-      with c have c': "Suc n < xs!i" by arith
-      have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
-        by simp_all
-      have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
-        by auto
-      have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
-        using i by auto
-      from xs have "Suc n = sum_list xs"
-        by (simp add: natpermute_def)
-      also have "\<dots> = sum (nth xs) {0..<Suc k}" using xs
-        by (simp add: natpermute_def sum_list_sum_nth)
-      also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
-        unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
-        unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
-        by simp
-      finally show ?thesis using c' by simp
-    qed
-    then show "((r, Suc k, a, xs!i), r, Suc k, a, Suc n) \<in> ?R"
-      apply auto
-      apply (metis not_less)
-      done
-  next
-    fix r k a n
-    show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \<in> ?R" by simp
-  }
-qed
-
-definition "fps_radical r n a = Abs_fps (radical r n a)"
-
-lemma fps_radical0[simp]: "fps_radical r 0 a = 1"
-  apply (auto simp add: fps_eq_iff fps_radical_def)
-  apply (case_tac n)
-  apply auto
-  done
-
-lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n = 0 then 1 else r n (a$0))"
-  by (cases n) (simp_all add: fps_radical_def)
-
-lemma fps_radical_power_nth[simp]:
-  assumes r: "(r k (a$0)) ^ k = a$0"
-  shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)"
-proof (cases k)
-  case 0
-  then show ?thesis by simp
-next
-  case (Suc h)
-  have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
-    unfolding fps_power_nth Suc by simp
-  also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
-    apply (rule prod.cong)
-    apply simp
-    using Suc
-    apply (subgoal_tac "replicate k 0 ! x = 0")
-    apply (auto intro: nth_replicate simp del: replicate.simps)
-    done
-  also have "\<dots> = a$0"
-    using r Suc by (simp add: prod_constant)
-  finally show ?thesis
-    using Suc by simp
-qed
-
-lemma power_radical:
-  fixes a:: "'a::field_char_0 fps"
-  assumes a0: "a$0 \<noteq> 0"
-  shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \<longleftrightarrow> (fps_radical r (Suc k) a) ^ (Suc k) = a"
-    (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  let ?r = "fps_radical r (Suc k) a"
-  show ?rhs if r0: ?lhs
-  proof -
-    from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
-    have "?r ^ Suc k $ z = a$z" for z
-    proof (induct z rule: nat_less_induct)
-      fix n
-      assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
-      show "?r ^ Suc k $ n = a $n"
-      proof (cases n)
-        case 0
-        then show ?thesis
-          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp
-      next
-        case (Suc n1)
-        then have "n \<noteq> 0" by simp
-        let ?Pnk = "natpermute n (k + 1)"
-        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-        have f: "finite ?Pnkn" "finite ?Pnknn"
-          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-          by (metis natpermute_finite)+
-        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
-        proof (rule sum.cong)
-          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
-          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
-            fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
-          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-            unfolding natpermute_contain_maximal by auto
-          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
-              (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
-            apply (rule prod.cong, simp)
-            using i r0
-            apply (simp del: replicate.simps)
-            done
-          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
-            using i r0 by (simp add: prod_gen_delta)
-          finally show ?ths .
-        qed rule
-        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
-          by (simp add: natpermute_max_card[OF \<open>n \<noteq> 0\<close>, simplified])
-        also have "\<dots> = a$n - sum ?f ?Pnknn"
-          unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
-        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
-        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
-          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
-        also have "\<dots> = a$n" unfolding fn by simp
-        finally show ?thesis .
-      qed
-    qed
-    then show ?thesis using r0 by (simp add: fps_eq_iff)
-  qed
-  show ?lhs if ?rhs
-  proof -
-    from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0"
-      by simp
-    then show ?thesis
-      unfolding fps_power_nth_Suc
-      by (simp add: prod_constant del: replicate.simps)
-  qed
-qed
-
-(*
-lemma power_radical:
-  fixes a:: "'a::field_char_0 fps"
-  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \<noteq> 0"
-  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
-proof-
-  let ?r = "fps_radical r (Suc k) a"
-  from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
-  {fix z have "?r ^ Suc k $ z = a$z"
-    proof(induct z rule: nat_less_induct)
-      fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
-      {assume "n = 0" then have "?r ^ Suc k $ n = a $n"
-          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
-      moreover
-      {fix n1 assume n1: "n = Suc n1"
-        have fK: "finite {0..k}" by simp
-        have nz: "n \<noteq> 0" using n1 by arith
-        let ?Pnk = "natpermute n (k + 1)"
-        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-        have f: "finite ?Pnkn" "finite ?Pnknn"
-          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-          by (metis natpermute_finite)+
-        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
-        proof(rule sum.cong2)
-          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
-          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
-          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-            unfolding natpermute_contain_maximal by auto
-          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
-            apply (rule prod.cong, simp)
-            using i r0 by (simp del: replicate.simps)
-          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
-            unfolding prod_gen_delta[OF fK] using i r0 by simp
-          finally show ?ths .
-        qed
-        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
-          by (simp add: natpermute_max_card[OF nz, simplified])
-        also have "\<dots> = a$n - sum ?f ?Pnknn"
-          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
-        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
-        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
-          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
-        also have "\<dots> = a$n" unfolding fn by simp
-        finally have "?r ^ Suc k $ n = a $n" .}
-      ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
-  qed }
-  then show ?thesis by (simp add: fps_eq_iff)
-qed
-
-*)
-lemma eq_divide_imp':
-  fixes c :: "'a::field"
-  shows "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
-  by (simp add: field_simps)
-
-lemma radical_unique:
-  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
-    and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0"
-    and b0: "b$0 \<noteq> 0"
-  shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
-    (is "?lhs \<longleftrightarrow> ?rhs" is "_ \<longleftrightarrow> a = ?r")
-proof
-  show ?lhs if ?rhs
-    using that using power_radical[OF b0, of r k, unfolded r0] by simp
-  show ?rhs if ?lhs
-  proof -
-    have r00: "r (Suc k) (b$0) \<noteq> 0" using b0 r0 by auto
-    have ceq: "card {0..k} = Suc k" by simp
-    from a0 have a0r0: "a$0 = ?r$0" by simp
-    have "a $ n = ?r $ n" for n
-    proof (induct n rule: nat_less_induct)
-      fix n
-      assume h: "\<forall>m<n. a$m = ?r $m"
-      show "a$n = ?r $ n"
-      proof (cases n)
-        case 0
-        then show ?thesis using a0 by simp
-      next
-        case (Suc n1)
-        have fK: "finite {0..k}" by simp
-        have nz: "n \<noteq> 0" using Suc by simp
-        let ?Pnk = "natpermute n (Suc k)"
-        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
-        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
-        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
-        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-        have f: "finite ?Pnkn" "finite ?Pnknn"
-          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
-          by (metis natpermute_finite)+
-        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-        let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
-        have "sum ?g ?Pnkn = sum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
-        proof (rule sum.cong)
-          fix v
-          assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
-          let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
-          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
-            unfolding Suc_eq_plus1 natpermute_contain_maximal
-            by (auto simp del: replicate.simps)
-          have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
-            apply (rule prod.cong, simp)
-            using i a0
-            apply (simp del: replicate.simps)
-            done
-          also have "\<dots> = a $ n * (?r $ 0)^k"
-            using i by (simp add: prod_gen_delta)
-          finally show ?ths .
-        qed rule
-        then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
-          by (simp add: natpermute_max_card[OF nz, simplified])
-        have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn"
-        proof (rule sum.cong, rule refl, rule prod.cong, simp)
-          fix xs i
-          assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
-          have False if c: "n \<le> xs ! i"
-          proof -
-            from xs i have "xs ! i \<noteq> n"
-              by (auto simp add: in_set_conv_nth natpermute_def)
-            with c have c': "n < xs!i" by arith
-            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
-              by simp_all
-            have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
-              by auto
-            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
-              using i by auto
-            from xs have "n = sum_list xs"
-              by (simp add: natpermute_def)
-            also have "\<dots> = sum (nth xs) {0..<Suc k}"
-              using xs by (simp add: natpermute_def sum_list_sum_nth)
-            also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
-              unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
-              unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
-              by simp
-            finally show ?thesis using c' by simp
-          qed
-          then have thn: "xs!i < n" by presburger
-          from h[rule_format, OF thn] show "a$(xs !i) = ?r$(xs!i)" .
-        qed
-        have th00: "\<And>x::'a. of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
-          by (simp add: field_simps del: of_nat_Suc)
-        from \<open>?lhs\<close> have "b$n = a^Suc k $ n"
-          by (simp add: fps_eq_iff)
-        also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn"
-          unfolding fps_power_nth_Suc
-          using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
-            unfolded eq, of ?g] by simp
-        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn"
-          unfolding th0 th1 ..
-        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn"
-          by simp
-        then have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
-          apply -
-          apply (rule eq_divide_imp')
-          using r00
-          apply (simp del: of_nat_Suc)
-          apply (simp add: ac_simps)
-          done
-        then show ?thesis
-          apply (simp del: of_nat_Suc)
-          unfolding fps_radical_def Suc
-          apply (simp add: field_simps Suc th00 del: of_nat_Suc)
-          done
-      qed
-    qed
-    then show ?rhs by (simp add: fps_eq_iff)
-  qed
-qed
-
-
-lemma radical_power:
-  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
-    and a0: "(a$0 :: 'a::field_char_0) \<noteq> 0"
-  shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
-proof -
-  let ?ak = "a^ Suc k"
-  have ak0: "?ak $ 0 = (a$0) ^ Suc k"
-    by (simp add: fps_nth_power_0 del: power_Suc)
-  from r0 have th0: "r (Suc k) (a ^ Suc k $ 0) ^ Suc k = a ^ Suc k $ 0"
-    using ak0 by auto
-  from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0"
-    by auto
-  from ak0 a0 have ak00: "?ak $ 0 \<noteq>0 "
-    by auto
-  from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis
-    by metis
-qed
-
-lemma fps_deriv_radical:
-  fixes a :: "'a::field_char_0 fps"
-  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
-    and a0: "a$0 \<noteq> 0"
-  shows "fps_deriv (fps_radical r (Suc k) a) =
-    fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)"
-proof -
-  let ?r = "fps_radical r (Suc k) a"
-  let ?w = "(fps_const (of_nat (Suc k)) * ?r ^ k)"
-  from a0 r0 have r0': "r (Suc k) (a$0) \<noteq> 0"
-    by auto
-  from r0' have w0: "?w $ 0 \<noteq> 0"
-    by (simp del: of_nat_Suc)
-  note th0 = inverse_mult_eq_1[OF w0]
-  let ?iw = "inverse ?w"
-  from iffD1[OF power_radical[of a r], OF a0 r0]
-  have "fps_deriv (?r ^ Suc k) = fps_deriv a"
-    by simp
-  then have "fps_deriv ?r * ?w = fps_deriv a"
-    by (simp add: fps_deriv_power ac_simps del: power_Suc)
-  then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a"
-    by simp
-  with a0 r0 have "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w"
-    by (subst fps_divide_unit) (auto simp del: of_nat_Suc)
-  then show ?thesis unfolding th0 by simp
-qed
-
-lemma radical_mult_distrib:
-  fixes a :: "'a::field_char_0 fps"
-  assumes k: "k > 0"
-    and ra0: "r k (a $ 0) ^ k = a $ 0"
-    and rb0: "r k (b $ 0) ^ k = b $ 0"
-    and a0: "a $ 0 \<noteq> 0"
-    and b0: "b $ 0 \<noteq> 0"
-  shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \<longleftrightarrow>
-    fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b"
-    (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  show ?rhs if r0': ?lhs
-  proof -
-    from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0"
-      by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
-    show ?thesis
-    proof (cases k)
-      case 0
-      then show ?thesis using r0' by simp
-    next
-      case (Suc h)
-      let ?ra = "fps_radical r (Suc h) a"
-      let ?rb = "fps_radical r (Suc h) b"
-      have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
-        using r0' Suc by (simp add: fps_mult_nth)
-      have ab0: "(a*b) $ 0 \<noteq> 0"
-        using a0 b0 by (simp add: fps_mult_nth)
-      from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded Suc] th0 ab0, symmetric]
-        iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded Suc]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded Suc]] Suc r0'
-      show ?thesis
-        by (auto simp add: power_mult_distrib simp del: power_Suc)
-    qed
-  qed
-  show ?lhs if ?rhs
-  proof -
-    from that have "(fps_radical r k (a * b)) $ 0 = (fps_radical r k a * fps_radical r k b) $ 0"
-      by simp
-    then show ?thesis
-      using k by (simp add: fps_mult_nth)
-  qed
-qed
-
-(*
-lemma radical_mult_distrib:
-  fixes a:: "'a::field_char_0 fps"
-  assumes
-  ra0: "r k (a $ 0) ^ k = a $ 0"
-  and rb0: "r k (b $ 0) ^ k = b $ 0"
-  and r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)"
-  and a0: "a$0 \<noteq> 0"
-  and b0: "b$0 \<noteq> 0"
-  shows "fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)"
-proof-
-  from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
-    by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
-  {assume "k=0" then have ?thesis by simp}
-  moreover
-  {fix h assume k: "k = Suc h"
-  let ?ra = "fps_radical r (Suc h) a"
-  let ?rb = "fps_radical r (Suc h) b"
-  have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
-    using r0' k by (simp add: fps_mult_nth)
-  have ab0: "(a*b) $ 0 \<noteq> 0" using a0 b0 by (simp add: fps_mult_nth)
-  from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric]
-    power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k
-  have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)}
-ultimately show ?thesis by (cases k, auto)
-qed
-*)
-
-lemma fps_divide_1 [simp]: "(a :: 'a::field fps) / 1 = a"
-  by (fact div_by_1)
-
-lemma radical_divide:
-  fixes a :: "'a::field_char_0 fps"
-  assumes kp: "k > 0"
-    and ra0: "(r k (a $ 0)) ^ k = a $ 0"
-    and rb0: "(r k (b $ 0)) ^ k = b $ 0"
-    and a0: "a$0 \<noteq> 0"
-    and b0: "b$0 \<noteq> 0"
-  shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \<longleftrightarrow>
-    fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
-  (is "?lhs = ?rhs")
-proof
-  let ?r = "fps_radical r k"
-  from kp obtain h where k: "k = Suc h"
-    by (cases k) auto
-  have ra0': "r k (a$0) \<noteq> 0" using a0 ra0 k by auto
-  have rb0': "r k (b$0) \<noteq> 0" using b0 rb0 k by auto
-
-  show ?lhs if ?rhs
-  proof -
-    from that have "?r (a/b) $ 0 = (?r a / ?r b)$0"
-      by simp
-    then show ?thesis
-      using k a0 b0 rb0' by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
-  qed
-  show ?rhs if ?lhs
-  proof -
-    from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0"
-      by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def)
-    have th0: "r k ((a/b)$0) ^ k = (a/b)$0"
-      by (simp add: \<open>?lhs\<close> power_divide ra0 rb0)
-    from a0 b0 ra0' rb0' kp \<open>?lhs\<close>
-    have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0"
-      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
-    from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \<noteq> 0"
-      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero)
-    note tha[simp] = iffD1[OF power_radical[where r=r and k=h], OF a0 ra0[unfolded k], unfolded k[symmetric]]
-    note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]]
-    from b0 rb0' have th2: "(?r a / ?r b)^k = a/b"
-      by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric])
-
-    from iffD1[OF radical_unique[where r=r and a="?r a / ?r b" and b="a/b" and k=h], symmetric, unfolded k[symmetric], OF th0 th1 ab0' th2]
-    show ?thesis .
-  qed
-qed
-
-lemma radical_inverse:
-  fixes a :: "'a::field_char_0 fps"
-  assumes k: "k > 0"
-    and ra0: "r k (a $ 0) ^ k = a $ 0"
-    and r1: "(r k 1)^k = 1"
-    and a0: "a$0 \<noteq> 0"
-  shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \<longleftrightarrow>
-    fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a"
-  using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0
-  by (simp add: divide_inverse fps_divide_def)
-
-
-subsection \<open>Derivative of composition\<close>
-
-lemma fps_compose_deriv:
-  fixes a :: "'a::idom fps"
-  assumes b0: "b$0 = 0"
-  shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b"
-proof -
-  have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n
-  proof -
-    have "(fps_deriv (a oo b))$n = sum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
-      by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc)
-    also have "\<dots> = sum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
-      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
-    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
-      unfolding fps_mult_left_const_nth  by (simp add: field_simps)
-    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
-      unfolding fps_mult_nth ..
-    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
-      apply (rule sum.mono_neutral_right)
-      apply (auto simp add: mult_delta_left sum.delta not_le)
-      done
-    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
-      unfolding fps_deriv_nth
-      by (rule sum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
-    finally have th0: "(fps_deriv (a oo b))$n =
-      sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
-
-    have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
-      unfolding fps_mult_nth by (simp add: ac_simps)
-    also have "\<dots> = sum (\<lambda>i. sum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
-      unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc
-      apply (rule sum.cong)
-      apply (rule refl)
-      apply (rule sum.mono_neutral_left)
-      apply (simp_all add: subset_eq)
-      apply clarify
-      apply (subgoal_tac "b^i$x = 0")
-      apply simp
-      apply (rule startsby_zero_power_prefix[OF b0, rule_format])
-      apply simp
-      done
-    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
-      unfolding sum_distrib_left
-      apply (subst sum.commute)
-      apply (rule sum.cong, rule refl)+
-      apply simp
-      done
-    finally show ?thesis
-      unfolding th0 by simp
-  qed
-  then show ?thesis by (simp add: fps_eq_iff)
-qed
-
-lemma fps_mult_X_plus_1_nth:
-  "((1+X)*a) $n = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
-proof (cases n)
-  case 0
-  then show ?thesis
-    by (simp add: fps_mult_nth)
-next
-  case (Suc m)
-  have "((1 + X)*a) $ n = sum (\<lambda>i. (1 + X) $ i * a $ (n - i)) {0..n}"
-    by (simp add: fps_mult_nth)
-  also have "\<dots> = sum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
-    unfolding Suc by (rule sum.mono_neutral_right) auto
-  also have "\<dots> = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
-    by (simp add: Suc)
-  finally show ?thesis .
-qed
-
-
-subsection \<open>Finite FPS (i.e. polynomials) and X\<close>
-
-lemma fps_poly_sum_X:
-  assumes "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
-  shows "a = sum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
-proof -
-  have "a$i = ?r$i" for i
-    unfolding fps_sum_nth fps_mult_left_const_nth X_power_nth
-    by (simp add: mult_delta_right sum.delta' assms)
-  then show ?thesis
-    unfolding fps_eq_iff by blast
-qed
-
-
-subsection \<open>Compositional inverses\<close>
-
-fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
-where
-  "compinv a 0 = X$0"
-| "compinv a (Suc n) =
-    (X$ Suc n - sum (\<lambda>i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
-
-definition "fps_inv a = Abs_fps (compinv a)"
-
-lemma fps_inv:
-  assumes a0: "a$0 = 0"
-    and a1: "a$1 \<noteq> 0"
-  shows "fps_inv a oo a = X"
-proof -
-  let ?i = "fps_inv a oo a"
-  have "?i $n = X$n" for n
-  proof (induct n rule: nat_less_induct)
-    fix n
-    assume h: "\<forall>m<n. ?i$m = X$m"
-    show "?i $ n = X$n"
-    proof (cases n)
-      case 0
-      then show ?thesis using a0
-        by (simp add: fps_compose_nth fps_inv_def)
-    next
-      case (Suc n1)
-      have "?i $ n = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
-        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
-      also have "\<dots> = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
-        (X$ Suc n1 - sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
-        using a0 a1 Suc by (simp add: fps_inv_def)
-      also have "\<dots> = X$n" using Suc by simp
-      finally show ?thesis .
-    qed
-  qed
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-
-fun gcompinv :: "'a fps \<Rightarrow> 'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
-where
-  "gcompinv b a 0 = b$0"
-| "gcompinv b a (Suc n) =
-    (b$ Suc n - sum (\<lambda>i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
-
-definition "fps_ginv b a = Abs_fps (gcompinv b a)"
-
-lemma fps_ginv:
-  assumes a0: "a$0 = 0"
-    and a1: "a$1 \<noteq> 0"
-  shows "fps_ginv b a oo a = b"
-proof -
-  let ?i = "fps_ginv b a oo a"
-  have "?i $n = b$n" for n
-  proof (induct n rule: nat_less_induct)
-    fix n
-    assume h: "\<forall>m<n. ?i$m = b$m"
-    show "?i $ n = b$n"
-    proof (cases n)
-      case 0
-      then show ?thesis using a0
-        by (simp add: fps_compose_nth fps_ginv_def)
-    next
-      case (Suc n1)
-      have "?i $ n = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
-        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
-      also have "\<dots> = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
-        (b$ Suc n1 - sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
-        using a0 a1 Suc by (simp add: fps_ginv_def)
-      also have "\<dots> = b$n" using Suc by simp
-      finally show ?thesis .
-    qed
-  qed
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-lemma fps_inv_ginv: "fps_inv = fps_ginv X"
-  apply (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def)
-  apply (induct_tac n rule: nat_less_induct)
-  apply auto
-  apply (case_tac na)
-  apply simp
-  apply simp
-  done
-
-lemma fps_compose_1[simp]: "1 oo a = 1"
-  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
-
-lemma fps_compose_0[simp]: "0 oo a = 0"
-  by (simp add: fps_eq_iff fps_compose_nth)
-
-lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)"
-  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral)
-
-lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib)
-
-lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\<lambda>i. f i oo a) S"
-proof (cases "finite S")
-  case True
-  show ?thesis
-  proof (rule finite_induct[OF True])
-    show "sum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
-      by simp
-  next
-    fix x F
-    assume fF: "finite F"
-      and xF: "x \<notin> F"
-      and h: "sum f F oo a = sum (\<lambda>i. f i oo a) F"
-    show "sum f (insert x F) oo a  = sum (\<lambda>i. f i oo a) (insert x F)"
-      using fF xF h by (simp add: fps_compose_add_distrib)
-  qed
-next
-  case False
-  then show ?thesis by simp
-qed
-
-lemma convolution_eq:
-  "sum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
-    sum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
-  by (rule sum.reindex_bij_witness[where i=fst and j="\<lambda>i. (i, n - i)"]) auto
-
-lemma product_composition_lemma:
-  assumes c0: "c$0 = (0::'a::idom)"
-    and d0: "d$0 = 0"
-  shows "((a oo c) * (b oo d))$n =
-    sum (\<lambda>(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}"  (is "?l = ?r")
-proof -
-  let ?S = "{(k::nat, m::nat). k + m \<le> n}"
-  have s: "?S \<subseteq> {0..n} \<times> {0..n}" by (auto simp add: subset_eq)
-  have f: "finite {(k::nat, m::nat). k + m \<le> n}"
-    apply (rule finite_subset[OF s])
-    apply auto
-    done
-  have "?r =  sum (\<lambda>i. sum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
-    apply (simp add: fps_mult_nth sum_distrib_left)
-    apply (subst sum.commute)
-    apply (rule sum.cong)
-    apply (auto simp add: field_simps)
-    done
-  also have "\<dots> = ?l"
-    apply (simp add: fps_mult_nth fps_compose_nth sum_product)
-    apply (rule sum.cong)
-    apply (rule refl)
-    apply (simp add: sum.cartesian_product mult.assoc)
-    apply (rule sum.mono_neutral_right[OF f])
-    apply (simp add: subset_eq)
-    apply presburger
-    apply clarsimp
-    apply (rule ccontr)
-    apply (clarsimp simp add: not_le)
-    apply (case_tac "x < aa")
-    apply simp
-    apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0])
-    apply blast
-    apply simp
-    apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0])
-    apply blast
-    done
-  finally show ?thesis by simp
-qed
-
-lemma product_composition_lemma':
-  assumes c0: "c$0 = (0::'a::idom)"
-    and d0: "d$0 = 0"
-  shows "((a oo c) * (b oo d))$n =
-    sum (\<lambda>k. sum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}"  (is "?l = ?r")
-  unfolding product_composition_lemma[OF c0 d0]
-  unfolding sum.cartesian_product
-  apply (rule sum.mono_neutral_left)
-  apply simp
-  apply (clarsimp simp add: subset_eq)
-  apply clarsimp
-  apply (rule ccontr)
-  apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
-  apply simp
-  unfolding fps_mult_nth
-  apply (rule sum.neutral)
-  apply (clarsimp simp add: not_le)
-  apply (case_tac "x < aa")
-  apply (rule startsby_zero_power_prefix[OF c0, rule_format])
-  apply simp
-  apply (subgoal_tac "n - x < ba")
-  apply (frule_tac k = "ba" in startsby_zero_power_prefix[OF d0, rule_format])
-  apply simp
-  apply arith
-  done
-
-
-lemma sum_pair_less_iff:
-  "sum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
-    sum (\<lambda>s. sum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
-  (is "?l = ?r")
-proof -
-  let ?KM = "{(k,m). k + m \<le> n}"
-  let ?f = "\<lambda>s. UNION {(0::nat)..s} (\<lambda>i. {(i,s - i)})"
-  have th0: "?KM = UNION {0..n} ?f"
-    by auto
-  show "?l = ?r "
-    unfolding th0
-    apply (subst sum.UNION_disjoint)
-    apply auto
-    apply (subst sum.UNION_disjoint)
-    apply auto
-    done
-qed
-
-lemma fps_compose_mult_distrib_lemma:
-  assumes c0: "c$0 = (0::'a::idom)"
-  shows "((a oo c) * (b oo c))$n = sum (\<lambda>s. sum (\<lambda>i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
-  unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
-  unfolding sum_pair_less_iff[where a = "\<lambda>k. a$k" and b="\<lambda>m. b$m" and c="\<lambda>s. (c ^ s)$n" and n = n] ..
-
-lemma fps_compose_mult_distrib:
-  assumes c0: "c $ 0 = (0::'a::idom)"
-  shows "(a * b) oo c = (a oo c) * (b oo c)"
-  apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
-  apply (simp add: fps_compose_nth fps_mult_nth sum_distrib_right)
-  done
-
-lemma fps_compose_prod_distrib:
-  assumes c0: "c$0 = (0::'a::idom)"
-  shows "prod a S oo c = prod (\<lambda>k. a k oo c) S"
-  apply (cases "finite S")
-  apply simp_all
-  apply (induct S rule: finite_induct)
-  apply simp
-  apply (simp add: fps_compose_mult_distrib[OF c0])
-  done
-
-lemma fps_compose_divide:
-  assumes [simp]: "g dvd f" "h $ 0 = 0"
-  shows   "fps_compose f h = fps_compose (f / g :: 'a :: field fps) h * fps_compose g h"
-proof -
-  have "f = (f / g) * g" by simp
-  also have "fps_compose \<dots> h = fps_compose (f / g) h * fps_compose g h"
-    by (subst fps_compose_mult_distrib) simp_all
-  finally show ?thesis .
-qed
-
-lemma fps_compose_divide_distrib:
-  assumes "g dvd f" "h $ 0 = 0" "fps_compose g h \<noteq> 0"
-  shows   "fps_compose (f / g :: 'a :: field fps) h = fps_compose f h / fps_compose g h"
-  using fps_compose_divide[OF assms(1,2)] assms(3) by simp
-
-lemma fps_compose_power:
-  assumes c0: "c$0 = (0::'a::idom)"
-  shows "(a oo c)^n = a^n oo c"
-proof (cases n)
-  case 0
-  then show ?thesis by simp
-next
-  case (Suc m)
-  have th0: "a^n = prod (\<lambda>k. a) {0..m}" "(a oo c) ^ n = prod (\<lambda>k. a oo c) {0..m}"
-    by (simp_all add: prod_constant Suc)
-  then show ?thesis
-    by (simp add: fps_compose_prod_distrib[OF c0])
-qed
-
-lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric])
-    
-lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
-  using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
-
-lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
-  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
-
-lemma fps_inverse_compose:
-  assumes b0: "(b$0 :: 'a::field) = 0"
-    and a0: "a$0 \<noteq> 0"
-  shows "inverse a oo b = inverse (a oo b)"
-proof -
-  let ?ia = "inverse a"
-  let ?ab = "a oo b"
-  let ?iab = "inverse ?ab"
-
-  from a0 have ia0: "?ia $ 0 \<noteq> 0" by simp
-  from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
-  have "(?ia oo b) *  (a oo b) = 1"
-    unfolding fps_compose_mult_distrib[OF b0, symmetric]
-    unfolding inverse_mult_eq_1[OF a0]
-    fps_compose_1 ..
-
-  then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
-  then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
-  then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp
-qed
-
-lemma fps_divide_compose:
-  assumes c0: "(c$0 :: 'a::field) = 0"
-    and b0: "b$0 \<noteq> 0"
-  shows "(a/b) oo c = (a oo c) / (b oo c)"
-    using b0 c0 by (simp add: fps_divide_unit fps_inverse_compose fps_compose_mult_distrib)
-
-lemma gp:
-  assumes a0: "a$0 = (0::'a::field)"
-  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)"
-    (is "?one oo a = _")
-proof -
-  have o0: "?one $ 0 \<noteq> 0" by simp
-  have th0: "(1 - X) $ 0 \<noteq> (0::'a)" by simp
-  from fps_inverse_gp[where ?'a = 'a]
-  have "inverse ?one = 1 - X" by (simp add: fps_eq_iff)
-  then have "inverse (inverse ?one) = inverse (1 - X)" by simp
-  then have th: "?one = 1/(1 - X)" unfolding fps_inverse_idempotent[OF o0]
-    by (simp add: fps_divide_def)
-  show ?thesis
-    unfolding th
-    unfolding fps_divide_compose[OF a0 th0]
-    fps_compose_1 fps_compose_sub_distrib X_fps_compose_startby0[OF a0] ..
-qed
-
-lemma fps_const_power [simp]: "fps_const (c::'a::ring_1) ^ n = fps_const (c^n)"
-  by (induct n) auto
-
-lemma fps_compose_radical:
-  assumes b0: "b$0 = (0::'a::field_char_0)"
-    and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
-    and a0: "a$0 \<noteq> 0"
-  shows "fps_radical r (Suc k)  a oo b = fps_radical r (Suc k) (a oo b)"
-proof -
-  let ?r = "fps_radical r (Suc k)"
-  let ?ab = "a oo b"
-  have ab0: "?ab $ 0 = a$0"
-    by (simp add: fps_compose_def)
-  from ab0 a0 ra0 have rab0: "?ab $ 0 \<noteq> 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0"
-    by simp_all
-  have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0"
-    by (simp add: ab0 fps_compose_def)
-  have th0: "(?r a oo b) ^ (Suc k) = a  oo b"
-    unfolding fps_compose_power[OF b0]
-    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  ..
-  from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0]
-  show ?thesis  .
-qed
-
-lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
-  by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc)
-
-lemma fps_const_mult_apply_right:
-  "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
-  by (auto simp add: fps_const_mult_apply_left mult.commute)
-
-lemma fps_compose_assoc:
-  assumes c0: "c$0 = (0::'a::idom)"
-    and b0: "b$0 = 0"
-  shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
-proof -
-  have "?l$n = ?r$n" for n
-  proof -
-    have "?l$n = (sum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
-      by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
-        sum_distrib_left mult.assoc fps_sum_nth)
-    also have "\<dots> = ((sum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
-      by (simp add: fps_compose_sum_distrib)
-    also have "\<dots> = ?r$n"
-      apply (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc)
-      apply (rule sum.cong)
-      apply (rule refl)
-      apply (rule sum.mono_neutral_right)
-      apply (auto simp add: not_le)
-      apply (erule startsby_zero_power_prefix[OF b0, rule_format])
-      done
-    finally show ?thesis .
-  qed
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-
-lemma fps_X_power_compose:
-  assumes a0: "a$0=0"
-  shows "X^k oo a = (a::'a::idom fps)^k"
-  (is "?l = ?r")
-proof (cases k)
-  case 0
-  then show ?thesis by simp
-next
-  case (Suc h)
-  have "?l $ n = ?r $n" for n
-  proof -
-    consider "k > n" | "k \<le> n" by arith
-    then show ?thesis
-    proof cases
-      case 1
-      then show ?thesis
-        using a0 startsby_zero_power_prefix[OF a0] Suc
-        by (simp add: fps_compose_nth del: power_Suc)
-    next
-      case 2
-      then show ?thesis
-        by (simp add: fps_compose_nth mult_delta_left sum.delta)
-    qed
-  qed
-  then show ?thesis
-    unfolding fps_eq_iff by blast
-qed
-
-lemma fps_inv_right:
-  assumes a0: "a$0 = 0"
-    and a1: "a$1 \<noteq> 0"
-  shows "a oo fps_inv a = X"
-proof -
-  let ?ia = "fps_inv a"
-  let ?iaa = "a oo fps_inv a"
-  have th0: "?ia $ 0 = 0"
-    by (simp add: fps_inv_def)
-  have th1: "?iaa $ 0 = 0"
-    using a0 a1 by (simp add: fps_inv_def fps_compose_nth)
-  have th2: "X$0 = 0"
-    by simp
-  from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo X"
-    by simp
-  then have "(a oo fps_inv a) oo a = X oo a"
-    by (simp add: fps_compose_assoc[OF a0 th0] X_fps_compose_startby0[OF a0])
-  with fps_compose_inj_right[OF a0 a1] show ?thesis
-    by simp
-qed
-
-lemma fps_inv_deriv:
-  assumes a0: "a$0 = (0::'a::field)"
-    and a1: "a$1 \<noteq> 0"
-  shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)"
-proof -
-  let ?ia = "fps_inv a"
-  let ?d = "fps_deriv a oo ?ia"
-  let ?dia = "fps_deriv ?ia"
-  have ia0: "?ia$0 = 0"
-    by (simp add: fps_inv_def)
-  have th0: "?d$0 \<noteq> 0"
-    using a1 by (simp add: fps_compose_nth)
-  from fps_inv_right[OF a0 a1] have "?d * ?dia = 1"
-    by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] )
-  then have "inverse ?d * ?d * ?dia = inverse ?d * 1"
-    by simp
-  with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d"
-    by simp
-qed
-
-lemma fps_inv_idempotent:
-  assumes a0: "a$0 = 0"
-    and a1: "a$1 \<noteq> 0"
-  shows "fps_inv (fps_inv a) = a"
-proof -
-  let ?r = "fps_inv"
-  have ra0: "?r a $ 0 = 0"
-    by (simp add: fps_inv_def)
-  from a1 have ra1: "?r a $ 1 \<noteq> 0"
-    by (simp add: fps_inv_def field_simps)
-  have X0: "X$0 = 0"
-    by simp
-  from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
-  then have "?r (?r a) oo ?r a oo a = X oo a"
-    by simp
-  then have "?r (?r a) oo (?r a oo a) = a"
-    unfolding X_fps_compose_startby0[OF a0]
-    unfolding fps_compose_assoc[OF a0 ra0, symmetric] .
-  then show ?thesis
-    unfolding fps_inv[OF a0 a1] by simp
-qed
-
-lemma fps_ginv_ginv:
-  assumes a0: "a$0 = 0"
-    and a1: "a$1 \<noteq> 0"
-    and c0: "c$0 = 0"
-    and  c1: "c$1 \<noteq> 0"
-  shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c"
-proof -
-  let ?r = "fps_ginv"
-  from c0 have rca0: "?r c a $0 = 0"
-    by (simp add: fps_ginv_def)
-  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0"
-    by (simp add: fps_ginv_def field_simps)
-  from fps_ginv[OF rca0 rca1]
-  have "?r b (?r c a) oo ?r c a = b" .
-  then have "?r b (?r c a) oo ?r c a oo a = b oo a"
-    by simp
-  then have "?r b (?r c a) oo (?r c a oo a) = b oo a"
-    apply (subst fps_compose_assoc)
-    using a0 c0
-    apply (auto simp add: fps_ginv_def)
-    done
-  then have "?r b (?r c a) oo c = b oo a"
-    unfolding fps_ginv[OF a0 a1] .
-  then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c"
-    by simp
-  then have "?r b (?r c a) oo (c oo fps_inv c) = b oo a oo fps_inv c"
-    apply (subst fps_compose_assoc)
-    using a0 c0
-    apply (auto simp add: fps_inv_def)
-    done
-  then show ?thesis
-    unfolding fps_inv_right[OF c0 c1] by simp
-qed
-
-lemma fps_ginv_deriv:
-  assumes a0:"a$0 = (0::'a::field)"
-    and a1: "a$1 \<noteq> 0"
-  shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv X a"
-proof -
-  let ?ia = "fps_ginv b a"
-  let ?iXa = "fps_ginv X a"
-  let ?d = "fps_deriv"
-  let ?dia = "?d ?ia"
-  have iXa0: "?iXa $ 0 = 0"
-    by (simp add: fps_ginv_def)
-  have da0: "?d a $ 0 \<noteq> 0"
-    using a1 by simp
-  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b"
-    by simp
-  then have "(?d ?ia oo a) * ?d a = ?d b"
-    unfolding fps_compose_deriv[OF a0] .
-  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)"
-    by simp
-  with a1 have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a"
-    by (simp add: fps_divide_unit)
-  then have "(?d ?ia oo a) oo ?iXa =  (?d b / ?d a) oo ?iXa"
-    unfolding inverse_mult_eq_1[OF da0] by simp
-  then have "?d ?ia oo (a oo ?iXa) =  (?d b / ?d a) oo ?iXa"
-    unfolding fps_compose_assoc[OF iXa0 a0] .
-  then show ?thesis unfolding fps_inv_ginv[symmetric]
-    unfolding fps_inv_right[OF a0 a1] by simp
-qed
-
-lemma fps_compose_linear:
-  "fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * X) = Abs_fps (\<lambda>n. c^n * f $ n)"
-  by (simp add: fps_eq_iff fps_compose_def power_mult_distrib
-                if_distrib sum.delta' cong: if_cong)
-              
-lemma fps_compose_uminus': 
-  "fps_compose f (-X :: 'a :: comm_ring_1 fps) = Abs_fps (\<lambda>n. (-1)^n * f $ n)"
-  using fps_compose_linear[of f "-1"] 
-  by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp
-
-subsection \<open>Elementary series\<close>
-
-subsubsection \<open>Exponential series\<close>
-
-definition "fps_exp x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
-
-lemma fps_exp_deriv[simp]: "fps_deriv (fps_exp a) = fps_const (a::'a::field_char_0) * fps_exp a" 
-  (is "?l = ?r")
-proof -
-  have "?l$n = ?r $ n" for n
-    apply (auto simp add: fps_exp_def field_simps power_Suc[symmetric]
-      simp del: fact_Suc of_nat_Suc power_Suc)
-    apply (simp add: field_simps)
-    done
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-lemma fps_exp_unique_ODE:
-  "fps_deriv a = fps_const c * a \<longleftrightarrow> a = fps_const (a$0) * fps_exp (c::'a::field_char_0)"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  show ?rhs if ?lhs
-  proof -
-    from that have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
-      by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
-    have th': "a$n = a$0 * c ^ n/ (fact n)" for n
-    proof (induct n)
-      case 0
-      then show ?case by simp
-    next
-      case Suc
-      then show ?case
-        unfolding th
-        using fact_gt_zero
-        apply (simp add: field_simps del: of_nat_Suc fact_Suc)
-        apply simp
-        done
-    qed
-    show ?thesis
-      by (auto simp add: fps_eq_iff fps_const_mult_left fps_exp_def intro: th')
-  qed
-  show ?lhs if ?rhs
-    using that by (metis fps_exp_deriv fps_deriv_mult_const_left mult.left_commute)
-qed
-
-lemma fps_exp_add_mult: "fps_exp (a + b) = fps_exp (a::'a::field_char_0) * fps_exp b" (is "?l = ?r")
-proof -
-  have "fps_deriv ?r = fps_const (a + b) * ?r"
-    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
-  then have "?r = ?l"
-    by (simp only: fps_exp_unique_ODE) (simp add: fps_mult_nth fps_exp_def)
-  then show ?thesis ..
-qed
-
-lemma fps_exp_nth[simp]: "fps_exp a $ n = a^n / of_nat (fact n)"
-  by (simp add: fps_exp_def)
-
-lemma fps_exp_0[simp]: "fps_exp (0::'a::field) = 1"
-  by (simp add: fps_eq_iff power_0_left)
-
-lemma fps_exp_neg: "fps_exp (- a) = inverse (fps_exp (a::'a::field_char_0))"
-proof -
-  from fps_exp_add_mult[of a "- a"] have th0: "fps_exp a * fps_exp (- a) = 1" by simp
-  from fps_inverse_unique[OF th0] show ?thesis by simp
-qed
-
-lemma fps_exp_nth_deriv[simp]: 
-  "fps_nth_deriv n (fps_exp (a::'a::field_char_0)) = (fps_const a)^n * (fps_exp a)"
-  by (induct n) auto
-
-lemma X_compose_fps_exp[simp]: "X oo fps_exp (a::'a::field) = fps_exp a - 1"
-  by (simp add: fps_eq_iff X_fps_compose)
-
-lemma fps_inv_fps_exp_compose:
-  assumes a: "a \<noteq> 0"
-  shows "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = X"
-    and "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = X"
-proof -
-  let ?b = "fps_exp a - 1"
-  have b0: "?b $ 0 = 0"
-    by simp
-  have b1: "?b $ 1 \<noteq> 0"
-    by (simp add: a)
-  from fps_inv[OF b0 b1] show "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = X" .
-  from fps_inv_right[OF b0 b1] show "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = X" .
-qed
-
-lemma fps_exp_power_mult: "(fps_exp (c::'a::field_char_0))^n = fps_exp (of_nat n * c)"
-  by (induct n) (auto simp add: field_simps fps_exp_add_mult)
-
-lemma radical_fps_exp:
-  assumes r: "r (Suc k) 1 = 1"
-  shows "fps_radical r (Suc k) (fps_exp (c::'a::field_char_0)) = fps_exp (c / of_nat (Suc k))"
-proof -
-  let ?ck = "(c / of_nat (Suc k))"
-  let ?r = "fps_radical r (Suc k)"
-  have eq0[simp]: "?ck * of_nat (Suc k) = c" "of_nat (Suc k) * ?ck = c"
-    by (simp_all del: of_nat_Suc)
-  have th0: "fps_exp ?ck ^ (Suc k) = fps_exp c" unfolding fps_exp_power_mult eq0 ..
-  have th: "r (Suc k) (fps_exp c $0) ^ Suc k = fps_exp c $ 0"
-    "r (Suc k) (fps_exp c $ 0) = fps_exp ?ck $ 0" "fps_exp c $ 0 \<noteq> 0" using r by simp_all
-  from th0 radical_unique[where r=r and k=k, OF th] show ?thesis
-    by auto
-qed
-
-lemma fps_exp_compose_linear [simp]: 
-  "fps_exp (d::'a::field_char_0) oo (fps_const c * X) = fps_exp (c * d)"
-  by (simp add: fps_compose_linear fps_exp_def fps_eq_iff power_mult_distrib)
-  
-lemma fps_fps_exp_compose_minus [simp]: 
-  "fps_compose (fps_exp c) (-X) = fps_exp (-c :: 'a :: field_char_0)"
-  using fps_exp_compose_linear[of c "-1 :: 'a"] 
-  unfolding fps_const_neg [symmetric] fps_const_1_eq_1 by simp
-
-lemma fps_exp_eq_iff [simp]: "fps_exp c = fps_exp d \<longleftrightarrow> c = (d :: 'a :: field_char_0)"
-proof
-  assume "fps_exp c = fps_exp d"
-  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] show "c = d" by simp
-qed simp_all
-
-lemma fps_exp_eq_fps_const_iff [simp]: 
-  "fps_exp (c :: 'a :: field_char_0) = fps_const c' \<longleftrightarrow> c = 0 \<and> c' = 1"
-proof
-  assume "c = 0 \<and> c' = 1"
-  thus "fps_exp c = fps_const c'" by (auto simp: fps_eq_iff)
-next
-  assume "fps_exp c = fps_const c'"
-  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] 
-    show "c = 0 \<and> c' = 1" by simp_all
-qed
-
-lemma fps_exp_neq_0 [simp]: "\<not>fps_exp (c :: 'a :: field_char_0) = 0"
-  unfolding fps_const_0_eq_0 [symmetric] fps_exp_eq_fps_const_iff by simp  
-
-lemma fps_exp_eq_1_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = 1 \<longleftrightarrow> c = 0"
-  unfolding fps_const_1_eq_1 [symmetric] fps_exp_eq_fps_const_iff by simp
-    
-lemma fps_exp_neq_numeral_iff [simp]: 
-  "fps_exp (c :: 'a :: field_char_0) = numeral n \<longleftrightarrow> c = 0 \<and> n = Num.One"
-  unfolding numeral_fps_const fps_exp_eq_fps_const_iff by simp
-
-
-subsubsection \<open>Logarithmic series\<close>
-
-lemma Abs_fps_if_0:
-  "Abs_fps (\<lambda>n. if n = 0 then (v::'a::ring_1) else f n) =
-    fps_const v + X * Abs_fps (\<lambda>n. f (Suc n))"
-  by (auto simp add: fps_eq_iff)
-
-definition fps_ln :: "'a::field_char_0 \<Rightarrow> 'a fps"
-  where "fps_ln c = fps_const (1/c) * Abs_fps (\<lambda>n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)"
-
-lemma fps_ln_deriv: "fps_deriv (fps_ln c) = fps_const (1/c) * inverse (1 + X)"
-  unfolding fps_inverse_X_plus1
-  by (simp add: fps_ln_def fps_eq_iff del: of_nat_Suc)
-
-lemma fps_ln_nth: "fps_ln c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
-  by (simp add: fps_ln_def field_simps)
-
-lemma fps_ln_0 [simp]: "fps_ln c $ 0 = 0" by (simp add: fps_ln_def)
-
-lemma fps_ln_fps_exp_inv:
-  fixes a :: "'a::field_char_0"
-  assumes a: "a \<noteq> 0"
-  shows "fps_ln a = fps_inv (fps_exp a - 1)"  (is "?l = ?r")
-proof -
-  let ?b = "fps_exp a - 1"
-  have b0: "?b $ 0 = 0" by simp
-  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
-  have "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) =
-    (fps_const a * (fps_exp a - 1) + fps_const a) oo fps_inv (fps_exp a - 1)"
-    by (simp add: field_simps)
-  also have "\<dots> = fps_const a * (X + 1)"
-    apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
-    apply (simp add: field_simps)
-    done
-  finally have eq: "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_const a * (X + 1)" .
-  from fps_inv_deriv[OF b0 b1, unfolded eq]
-  have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
-    using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
-  then have "fps_deriv ?l = fps_deriv ?r"
-    by (simp add: fps_ln_deriv add.commute fps_divide_def divide_inverse)
-  then show ?thesis unfolding fps_deriv_eq_iff
-    by (simp add: fps_ln_nth fps_inv_def)
-qed
-
-lemma fps_ln_mult_add:
-  assumes c0: "c\<noteq>0"
-    and d0: "d\<noteq>0"
-  shows "fps_ln c + fps_ln d = fps_const (c+d) * fps_ln (c*d)"
-  (is "?r = ?l")
-proof-
-  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
-  have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + X)"
-    by (simp add: fps_ln_deriv fps_const_add[symmetric] algebra_simps del: fps_const_add)
-  also have "\<dots> = fps_deriv ?l"
-    apply (simp add: fps_ln_deriv)
-    apply (simp add: fps_eq_iff eq)
-    done
-  finally show ?thesis
-    unfolding fps_deriv_eq_iff by simp
-qed
-
-lemma X_dvd_fps_ln [simp]: "X dvd fps_ln c"
-proof -
-  have "fps_ln c = X * Abs_fps (\<lambda>n. (-1) ^ n / (of_nat (Suc n) * c))"
-    by (intro fps_ext) (auto simp: fps_ln_def of_nat_diff)
-  thus ?thesis by simp
-qed
-
-
-subsubsection \<open>Binomial series\<close>
-
-definition "fps_binomial a = Abs_fps (\<lambda>n. a gchoose n)"
-
-lemma fps_binomial_nth[simp]: "fps_binomial a $ n = a gchoose n"
-  by (simp add: fps_binomial_def)
-
-lemma fps_binomial_ODE_unique:
-  fixes c :: "'a::field_char_0"
-  shows "fps_deriv a = (fps_const c * a) / (1 + X) \<longleftrightarrow> a = fps_const (a$0) * fps_binomial c"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  let ?da = "fps_deriv a"
-  let ?x1 = "(1 + X):: 'a fps"
-  let ?l = "?x1 * ?da"
-  let ?r = "fps_const c * a"
-
-  have eq: "?l = ?r \<longleftrightarrow> ?lhs"
-  proof -
-    have x10: "?x1 $ 0 \<noteq> 0" by simp
-    have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
-    also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
-      apply (simp only: fps_divide_def  mult.assoc[symmetric] inverse_mult_eq_1[OF x10])
-      apply (simp add: field_simps)
-      done
-    finally show ?thesis .
-  qed
-
-  show ?rhs if ?lhs
-  proof -
-    from eq that have h: "?l = ?r" ..
-    have th0: "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" for n
-    proof -
-      from h have "?l $ n = ?r $ n" by simp
-      then show ?thesis
-        apply (simp add: field_simps del: of_nat_Suc)
-        apply (cases n)
-        apply (simp_all add: field_simps del: of_nat_Suc)
-        done
-    qed
-    have th1: "a $ n = (c gchoose n) * a $ 0" for n
-    proof (induct n)
-      case 0
-      then show ?case by simp
-    next
-      case (Suc m)
-      then show ?case
-        unfolding th0
-        apply (simp add: field_simps del: of_nat_Suc)
-        unfolding mult.assoc[symmetric] gbinomial_mult_1
-        apply (simp add: field_simps)
-        done
-    qed
-    show ?thesis
-      apply (simp add: fps_eq_iff)
-      apply (subst th1)
-      apply (simp add: field_simps)
-      done
-  qed
-
-  show ?lhs if ?rhs
-  proof -
-    have th00: "x * (a $ 0 * y) = a $ 0 * (x * y)" for x y
-      by (simp add: mult.commute)
-    have "?l = ?r"
-      apply (subst \<open>?rhs\<close>)
-      apply (subst (2) \<open>?rhs\<close>)
-      apply (clarsimp simp add: fps_eq_iff field_simps)
-      unfolding mult.assoc[symmetric] th00 gbinomial_mult_1
-      apply (simp add: field_simps gbinomial_mult_1)
-      done
-    with eq show ?thesis ..
-  qed
-qed
-
-lemma fps_binomial_ODE_unique':
-  "(fps_deriv a = fps_const c * a / (1 + X) \<and> a $ 0 = 1) \<longleftrightarrow> (a = fps_binomial c)"
-  by (subst fps_binomial_ODE_unique) auto
-
-lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + X)"
-proof -
-  let ?a = "fps_binomial c"
-  have th0: "?a = fps_const (?a$0) * ?a" by (simp)
-  from iffD2[OF fps_binomial_ODE_unique, OF th0] show ?thesis .
-qed
-
-lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r")
-proof -
-  let ?P = "?r - ?l"
-  let ?b = "fps_binomial"
-  let ?db = "\<lambda>x. fps_deriv (?b x)"
-  have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
-  also have "\<dots> = inverse (1 + X) *
-      (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
-    unfolding fps_binomial_deriv
-    by (simp add: fps_divide_def field_simps)
-  also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
-    by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add)
-  finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
-    by (simp add: fps_divide_def)
-  have "?P = fps_const (?P$0) * ?b (c + d)"
-    unfolding fps_binomial_ODE_unique[symmetric]
-    using th0 by simp
-  then have "?P = 0" by (simp add: fps_mult_nth)
-  then show ?thesis by simp
-qed
-
-lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + X)"
-  (is "?l = inverse ?r")
-proof-
-  have th: "?r$0 \<noteq> 0" by simp
-  have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)"
-    by (simp add: fps_inverse_deriv[OF th] fps_divide_def
-      power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg)
-  have eq: "inverse ?r $ 0 = 1"
-    by (simp add: fps_inverse_def)
-  from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq
-  show ?thesis by (simp add: fps_inverse_def)
-qed
-
-lemma fps_binomial_of_nat: "fps_binomial (of_nat n) = (1 + X :: 'a :: field_char_0 fps) ^ n"
-proof (cases "n = 0")
-  case [simp]: True
-  have "fps_deriv ((1 + X) ^ n :: 'a fps) = 0" by simp
-  also have "\<dots> = fps_const (of_nat n) * (1 + X) ^ n / (1 + X)" by (simp add: fps_binomial_def)
-  finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) simp_all
-next
-  case False
-  have "fps_deriv ((1 + X) ^ n :: 'a fps) = fps_const (of_nat n) * (1 + X) ^ (n - 1)"
-    by (simp add: fps_deriv_power)
-  also have "(1 + X :: 'a fps) $ 0 \<noteq> 0" by simp
-  hence "(1 + X :: 'a fps) \<noteq> 0" by (intro notI) (simp only: , simp)
-  with False have "(1 + X :: 'a fps) ^ (n - 1) = (1 + X) ^ n / (1 + X)"
-    by (cases n) (simp_all )
-  also have "fps_const (of_nat n :: 'a) * ((1 + X) ^ n / (1 + X)) =
-               fps_const (of_nat n) * (1 + X) ^ n / (1 + X)"
-    by (simp add: unit_div_mult_swap)
-  finally show ?thesis
-    by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) (simp_all add: fps_power_nth)
-qed
-
-lemma fps_binomial_0 [simp]: "fps_binomial 0 = 1"
-  using fps_binomial_of_nat[of 0] by simp
-  
-lemma fps_binomial_power: "fps_binomial a ^ n = fps_binomial (of_nat n * a)"
-  by (induction n) (simp_all add: fps_binomial_add_mult ring_distribs)
-
-lemma fps_binomial_1: "fps_binomial 1 = 1 + X"
-  using fps_binomial_of_nat[of 1] by simp
-
-lemma fps_binomial_minus_of_nat:
-  "fps_binomial (- of_nat n) = inverse ((1 + X :: 'a :: field_char_0 fps) ^ n)"
-  by (rule sym, rule fps_inverse_unique)
-     (simp add: fps_binomial_of_nat [symmetric] fps_binomial_add_mult [symmetric])
-
-lemma one_minus_const_X_power:
-  "c \<noteq> 0 \<Longrightarrow> (1 - fps_const c * X) ^ n =
-     fps_compose (fps_binomial (of_nat n)) (-fps_const c * X)"
-  by (subst fps_binomial_of_nat)
-     (simp add: fps_compose_power [symmetric] fps_compose_add_distrib fps_const_neg [symmetric] 
-           del: fps_const_neg)
-
-lemma one_minus_X_const_neg_power:
-  "inverse ((1 - fps_const c * X) ^ n) = 
-       fps_compose (fps_binomial (-of_nat n)) (-fps_const c * X)"
-proof (cases "c = 0")
-  case False
-  thus ?thesis
-  by (subst fps_binomial_minus_of_nat)
-     (simp add: fps_compose_power [symmetric] fps_inverse_compose fps_compose_add_distrib
-                fps_const_neg [symmetric] del: fps_const_neg)
-qed simp
-
-lemma X_plus_const_power:
-  "c \<noteq> 0 \<Longrightarrow> (X + fps_const c) ^ n =
-     fps_const (c^n) * fps_compose (fps_binomial (of_nat n)) (fps_const (inverse c) * X)"
-  by (subst fps_binomial_of_nat)
-     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
-                fps_const_power [symmetric] power_mult_distrib [symmetric] 
-                algebra_simps inverse_mult_eq_1' del: fps_const_power)
-
-lemma X_plus_const_neg_power:
-  "c \<noteq> 0 \<Longrightarrow> inverse ((X + fps_const c) ^ n) =
-     fps_const (inverse c^n) * fps_compose (fps_binomial (-of_nat n)) (fps_const (inverse c) * X)"
-  by (subst fps_binomial_minus_of_nat)
-     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
-                fps_const_power [symmetric] power_mult_distrib [symmetric] fps_inverse_compose 
-                algebra_simps fps_const_inverse [symmetric] fps_inverse_mult [symmetric]
-                fps_inverse_power [symmetric] inverse_mult_eq_1'
-           del: fps_const_power)
-
-
-lemma one_minus_const_X_neg_power':
-  "n > 0 \<Longrightarrow> inverse ((1 - fps_const (c :: 'a :: field_char_0) * X) ^ n) =
-       Abs_fps (\<lambda>k. of_nat ((n + k - 1) choose k) * c^k)"
-  apply (rule fps_ext)
-  apply (subst one_minus_X_const_neg_power, subst fps_const_neg, subst fps_compose_linear)
-  apply (simp add: power_mult_distrib [symmetric] mult.assoc [symmetric] 
-                   gbinomial_minus binomial_gbinomial of_nat_diff)
-  done
-
-text \<open>Vandermonde's Identity as a consequence.\<close>
-lemma gbinomial_Vandermonde:
-  "sum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
-proof -
-  let ?ba = "fps_binomial a"
-  let ?bb = "fps_binomial b"
-  let ?bab = "fps_binomial (a + b)"
-  from fps_binomial_add_mult[of a b] have "?bab $ n = (?ba * ?bb)$n" by simp
-  then show ?thesis by (simp add: fps_mult_nth)
-qed
-
-lemma binomial_Vandermonde:
-  "sum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
-  using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n]
-  by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric]
-                 of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
-
-lemma binomial_Vandermonde_same: "sum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
-  using binomial_Vandermonde[of n n n, symmetric]
-  unfolding mult_2
-  apply (simp add: power2_eq_square)
-  apply (rule sum.cong)
-  apply (auto intro:  binomial_symmetric)
-  done
-
-lemma Vandermonde_pochhammer_lemma:
-  fixes a :: "'a::field_char_0"
-  assumes b: "\<forall>j\<in>{0 ..<n}. b \<noteq> of_nat j"
-  shows "sum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
-      (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} =
-    pochhammer (- (a + b)) n / pochhammer (- b) n"
-  (is "?l = ?r")
-proof -
-  let ?m1 = "\<lambda>m. (- 1 :: 'a) ^ m"
-  let ?f = "\<lambda>m. of_nat (fact m)"
-  let ?p = "\<lambda>(x::'a). pochhammer (- x)"
-  from b have bn0: "?p b n \<noteq> 0"
-    unfolding pochhammer_eq_0_iff by simp
-  have th00:
-    "b gchoose (n - k) =
-        (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
-      (is ?gchoose)
-    "pochhammer (1 + b - of_nat n) k \<noteq> 0"
-      (is ?pochhammer)
-    if kn: "k \<in> {0..n}" for k
-  proof -
-    from kn have "k \<le> n" by simp
-    have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0"
-    proof
-      assume "pochhammer (1 + b - of_nat n) n = 0"
-      then have c: "pochhammer (b - of_nat n + 1) n = 0"
-        by (simp add: algebra_simps)
-      then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j"
-        unfolding pochhammer_eq_0_iff by blast
-      from j have "b = of_nat n - of_nat j - of_nat 1"
-        by (simp add: algebra_simps)
-      then have "b = of_nat (n - j - 1)"
-        using j kn by (simp add: of_nat_diff)
-      with b show False using j by auto
-    qed
-
-    from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \<noteq> 0"
-      by (rule pochhammer_neq_0_mono)
-
-    consider "k = 0 \<or> n = 0" | "k \<noteq> 0" "n \<noteq> 0"
-      by blast
-    then have "b gchoose (n - k) =
-      (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
-    proof cases
-      case 1
-      then show ?thesis
-        using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer)
-    next
-      case neq: 2
-      then obtain m where m: "n = Suc m"
-        by (cases n) auto
-      from neq(1) obtain h where h: "k = Suc h"
-        by (cases k) auto
-      show ?thesis
-      proof (cases "k = n")
-        case True
-        then show ?thesis
-          using pochhammer_minus'[where k=k and b=b]
-          apply (simp add: pochhammer_same)
-          using bn0
-          apply (simp add: field_simps power_add[symmetric])
-          done
-      next
-        case False
-        with kn have kn': "k < n"
-          by simp
-        have m1nk: "?m1 n = prod (\<lambda>i. - 1) {..m}" "?m1 k = prod (\<lambda>i. - 1) {0..h}"
-          by (simp_all add: prod_constant m h)
-        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
-          using bn0 kn
-          unfolding pochhammer_eq_0_iff
-          apply auto
-          apply (erule_tac x= "n - ka - 1" in allE)
-          apply (auto simp add: algebra_simps of_nat_diff)
-          done
-        have eq1: "prod (\<lambda>k. (1::'a) + of_nat m - of_nat k) {..h} =
-          prod of_nat {Suc (m - h) .. Suc m}"
-          using kn' h m
-          by (intro prod.reindex_bij_witness[where i="\<lambda>k. Suc m - k" and j="\<lambda>k. Suc m - k"])
-             (auto simp: of_nat_diff)
-        have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
-          apply (simp add: pochhammer_minus field_simps)
-          using \<open>k \<le> n\<close> apply (simp add: fact_split [of k n])
-          apply (simp add: pochhammer_prod)
-          using prod.atLeast_lessThan_shift_bounds [where ?'a = 'a, of "\<lambda>i. 1 + of_nat i" 0 "n - k" k]
-          apply (auto simp add: of_nat_diff field_simps)
-          done
-        have th20: "?m1 n * ?p b n = prod (\<lambda>i. b - of_nat i) {0..m}"
-          apply (simp add: pochhammer_minus field_simps m)
-          apply (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift)
-          done
-        have th21:"pochhammer (b - of_nat n + 1) k = prod (\<lambda>i. b - of_nat i) {n - k .. n - 1}"
-          using kn apply (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift)
-          using prod.atLeast_atMost_shift_0 [of "m - h" m, where ?'a = 'a]
-          apply (auto simp add: of_nat_diff field_simps)
-          done
-        have "?m1 n * ?p b n =
-          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1} * pochhammer (b - of_nat n + 1) k"
-          using kn' m h unfolding th20 th21 apply simp
-          apply (subst prod.union_disjoint [symmetric])
-          apply auto
-          apply (rule prod.cong)
-          apply auto
-          done
-        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k =
-          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1}"
-          using nz' by (simp add: field_simps)
-        have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) =
-          ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
-          using bnz0
-          by (simp add: field_simps)
-        also have "\<dots> = b gchoose (n - k)"
-          unfolding th1 th2
-          using kn' m h
-          apply (simp add: field_simps gbinomial_mult_fact)
-          apply (rule prod.cong)
-          apply auto
-          done
-        finally show ?thesis by simp
-      qed
-    qed
-    then show ?gchoose and ?pochhammer
-      apply (cases "n = 0")
-      using nz'
-      apply auto
-      done
-  qed
-  have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))"
-    unfolding gbinomial_pochhammer
-    using bn0 by (auto simp add: field_simps)
-  also have "\<dots> = ?l"
-    unfolding gbinomial_Vandermonde[symmetric]
-    apply (simp add: th00)
-    unfolding gbinomial_pochhammer
-    using bn0
-    apply (simp add: sum_distrib_right sum_distrib_left field_simps)
-    done
-  finally show ?thesis by simp
-qed
-
-lemma Vandermonde_pochhammer:
-  fixes a :: "'a::field_char_0"
-  assumes c: "\<forall>i \<in> {0..< n}. c \<noteq> - of_nat i"
-  shows "sum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
-    (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
-proof -
-  let ?a = "- a"
-  let ?b = "c + of_nat n - 1"
-  have h: "\<forall> j \<in>{0..< n}. ?b \<noteq> of_nat j"
-    using c
-    apply (auto simp add: algebra_simps of_nat_diff)
-    apply (erule_tac x = "n - j - 1" in ballE)
-    apply (auto simp add: of_nat_diff algebra_simps)
-    done
-  have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n"
-    unfolding pochhammer_minus
-    by (simp add: algebra_simps)
-  have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n"
-    unfolding pochhammer_minus
-    by simp
-  have nz: "pochhammer c n \<noteq> 0" using c
-    by (simp add: pochhammer_eq_0_iff)
-  from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
-  show ?thesis
-    using nz by (simp add: field_simps sum_distrib_left)
-qed
-
-
-subsubsection \<open>Formal trigonometric functions\<close>
-
-definition "fps_sin (c::'a::field_char_0) =
-  Abs_fps (\<lambda>n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))"
-
-definition "fps_cos (c::'a::field_char_0) =
-  Abs_fps (\<lambda>n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)"
-
-lemma fps_sin_deriv:
-  "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
-  (is "?lhs = ?rhs")
-proof (rule fps_ext)
-  fix n :: nat
-  show "?lhs $ n = ?rhs $ n"
-  proof (cases "even n")
-    case True
-    have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
-    also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
-      using True by (simp add: fps_sin_def)
-    also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
-      unfolding fact_Suc of_nat_mult
-      by (simp add: field_simps del: of_nat_add of_nat_Suc)
-    also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
-      by (simp add: field_simps del: of_nat_add of_nat_Suc)
-    finally show ?thesis
-      using True by (simp add: fps_cos_def field_simps)
-  next
-    case False
-    then show ?thesis
-      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
-  qed
-qed
-
-lemma fps_cos_deriv: "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)"
-  (is "?lhs = ?rhs")
-proof (rule fps_ext)
-  have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n
-    by simp
-  show "?lhs $ n = ?rhs $ n" for n
-  proof (cases "even n")
-    case False
-    then have n0: "n \<noteq> 0" by presburger
-    from False have th1: "Suc ((n - 1) div 2) = Suc n div 2"
-      by (cases n) simp_all
-    have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
-    also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
-      using False by (simp add: fps_cos_def)
-    also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
-      unfolding fact_Suc of_nat_mult
-      by (simp add: field_simps del: of_nat_add of_nat_Suc)
-    also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
-      by (simp add: field_simps del: of_nat_add of_nat_Suc)
-    also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
-      unfolding th0 unfolding th1 by simp
-    finally show ?thesis
-      using False by (simp add: fps_sin_def field_simps)
-  next
-    case True
-    then show ?thesis
-      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
-  qed
-qed
-
-lemma fps_sin_cos_sum_of_squares: "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1"
-  (is "?lhs = _")
-proof -
-  have "fps_deriv ?lhs = 0"
-    apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv)
-    apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
-    done
-  then have "?lhs = fps_const (?lhs $ 0)"
-    unfolding fps_deriv_eq_0_iff .
-  also have "\<dots> = 1"
-    by (auto simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def)
-  finally show ?thesis .
-qed
-
-lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0"
-  unfolding fps_sin_def by simp
-
-lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c"
-  unfolding fps_sin_def by simp
-
-lemma fps_sin_nth_add_2:
-    "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
-  unfolding fps_sin_def
-  apply (cases n)
-  apply simp
-  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
-  apply simp
-  done
-
-lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1"
-  unfolding fps_cos_def by simp
-
-lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0"
-  unfolding fps_cos_def by simp
-
-lemma fps_cos_nth_add_2:
-  "fps_cos c $ (n + 2) = - (c * c * fps_cos c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
-  unfolding fps_cos_def
-  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
-  apply simp
-  done
-
-lemma nat_induct2: "P 0 \<Longrightarrow> P 1 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (n + 2)) \<Longrightarrow> P (n::nat)"
-  unfolding One_nat_def numeral_2_eq_2
-  apply (induct n rule: nat_less_induct)
-  apply (case_tac n)
-  apply simp
-  apply (rename_tac m)
-  apply (case_tac m)
-  apply simp
-  apply (rename_tac k)
-  apply (case_tac k)
-  apply simp_all
-  done
-
-lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2"
-  by simp
-
-lemma eq_fps_sin:
-  assumes 0: "a $ 0 = 0"
-    and 1: "a $ 1 = c"
-    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
-  shows "a = fps_sin c"
-  apply (rule fps_ext)
-  apply (induct_tac n rule: nat_induct2)
-  apply (simp add: 0)
-  apply (simp add: 1 del: One_nat_def)
-  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
-  apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2
-              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
-  apply (subst minus_divide_left)
-  apply (subst nonzero_eq_divide_eq)
-  apply (simp del: of_nat_add of_nat_Suc)
-  apply (simp only: ac_simps)
-  done
-
-lemma eq_fps_cos:
-  assumes 0: "a $ 0 = 1"
-    and 1: "a $ 1 = 0"
-    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
-  shows "a = fps_cos c"
-  apply (rule fps_ext)
-  apply (induct_tac n rule: nat_induct2)
-  apply (simp add: 0)
-  apply (simp add: 1 del: One_nat_def)
-  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
-  apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2
-              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
-  apply (subst minus_divide_left)
-  apply (subst nonzero_eq_divide_eq)
-  apply (simp del: of_nat_add of_nat_Suc)
-  apply (simp only: ac_simps)
-  done
-
-lemma mult_nth_0 [simp]: "(a * b) $ 0 = a $ 0 * b $ 0"
-  by (simp add: fps_mult_nth)
-
-lemma mult_nth_1 [simp]: "(a * b) $ 1 = a $ 0 * b $ 1 + a $ 1 * b $ 0"
-  by (simp add: fps_mult_nth)
-
-lemma fps_sin_add: "fps_sin (a + b) = fps_sin a * fps_cos b + fps_cos a * fps_sin b"
-  apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def)
-  apply (simp del: fps_const_neg fps_const_add fps_const_mult
-              add: fps_const_add [symmetric] fps_const_neg [symmetric]
-                   fps_sin_deriv fps_cos_deriv algebra_simps)
-  done
-
-lemma fps_cos_add: "fps_cos (a + b) = fps_cos a * fps_cos b - fps_sin a * fps_sin b"
-  apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def)
-  apply (simp del: fps_const_neg fps_const_add fps_const_mult
-              add: fps_const_add [symmetric] fps_const_neg [symmetric]
-                   fps_sin_deriv fps_cos_deriv algebra_simps)
-  done
-
-lemma fps_sin_even: "fps_sin (- c) = - fps_sin c"
-  by (auto simp add: fps_eq_iff fps_sin_def)
-
-lemma fps_cos_odd: "fps_cos (- c) = fps_cos c"
-  by (auto simp add: fps_eq_iff fps_cos_def)
-
-definition "fps_tan c = fps_sin c / fps_cos c"
-
-lemma fps_tan_deriv: "fps_deriv (fps_tan c) = fps_const c / (fps_cos c)\<^sup>2"
-proof -
-  have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
-  from this have "fps_cos c \<noteq> 0" by (intro notI) simp
-  hence "fps_deriv (fps_tan c) =
-           fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)"
-    by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps
-                  fps_sin_deriv fps_cos_deriv fps_const_neg[symmetric] div_mult_swap
-             del: fps_const_neg)
-  also note fps_sin_cos_sum_of_squares
-  finally show ?thesis by simp
-qed
-
-text \<open>Connection to @{const "fps_exp"} over the complex numbers --- Euler and de Moivre.\<close>
-
-lemma fps_exp_ii_sin_cos: "fps_exp (\<i> * c) = fps_cos c + fps_const \<i> * fps_sin c"
-  (is "?l = ?r")
-proof -
-  have "?l $ n = ?r $ n" for n
-  proof (cases "even n")
-    case True
-    then obtain m where m: "n = 2 * m" ..
-    show ?thesis
-      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"])
-  next
-    case False
-    then obtain m where m: "n = 2 * m + 1" ..
-    show ?thesis
-      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
-        power_mult power_minus [of "c ^ 2"])
-  qed
-  then show ?thesis
-    by (simp add: fps_eq_iff)
-qed
-
-lemma fps_exp_minus_ii_sin_cos: "fps_exp (- (\<i> * c)) = fps_cos c - fps_const \<i> * fps_sin c"
-  unfolding minus_mult_right fps_exp_ii_sin_cos by (simp add: fps_sin_even fps_cos_odd)
-
-lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
-  by (fact fps_const_sub)
-
-lemma fps_of_int: "fps_const (of_int c) = of_int c"
-  by (induction c) (simp_all add: fps_const_minus [symmetric] fps_of_nat fps_const_neg [symmetric] 
-                             del: fps_const_minus fps_const_neg)
-
-lemma fps_deriv_of_int [simp]: "fps_deriv (of_int n) = 0"
-  by (simp add: fps_of_int [symmetric])
-
-lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a::comm_ring_1)"
-  by (fact numeral_fps_const) (* FIXME: duplicate *)
-
-lemma fps_cos_fps_exp_ii: "fps_cos c = (fps_exp (\<i> * c) + fps_exp (- \<i> * c)) / fps_const 2"
-proof -
-  have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2"
-    by (simp add: numeral_fps_const)
-  show ?thesis
-    unfolding fps_exp_ii_sin_cos minus_mult_commute
-    by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_unit fps_const_inverse th)
-qed
-
-lemma fps_sin_fps_exp_ii: "fps_sin c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / fps_const (2*\<i>)"
-proof -
-  have th: "fps_const \<i> * fps_sin c + fps_const \<i> * fps_sin c = fps_sin c * fps_const (2 * \<i>)"
-    by (simp add: fps_eq_iff numeral_fps_const)
-  show ?thesis
-    unfolding fps_exp_ii_sin_cos minus_mult_commute
-    by (simp add: fps_sin_even fps_cos_odd fps_divide_unit fps_const_inverse th)
-qed
-
-lemma fps_tan_fps_exp_ii:
-  "fps_tan c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / 
-      (fps_const \<i> * (fps_exp (\<i> * c) + fps_exp (- \<i> * c)))"
-  unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii mult_minus_left fps_exp_neg
-  apply (simp add: fps_divide_unit fps_inverse_mult fps_const_mult[symmetric] fps_const_inverse del: fps_const_mult)
-  apply simp
-  done
-
-lemma fps_demoivre:
-  "(fps_cos a + fps_const \<i> * fps_sin a)^n =
-    fps_cos (of_nat n * a) + fps_const \<i> * fps_sin (of_nat n * a)"
-  unfolding fps_exp_ii_sin_cos[symmetric] fps_exp_power_mult
-  by (simp add: ac_simps)
-
-
-subsection \<open>Hypergeometric series\<close>
-
-definition "fps_hypergeo as bs (c::'a::{field_char_0,field}) =
-  Abs_fps (\<lambda>n. (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
-    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
-
-lemma fps_hypergeo_nth[simp]: "fps_hypergeo as bs c $ n =
-  (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
-    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
-  by (simp add: fps_hypergeo_def)
-
-lemma foldl_mult_start:
-  fixes v :: "'a::comm_ring_1"
-  shows "foldl (\<lambda>r x. r * f x) v as * x = foldl (\<lambda>r x. r * f x) (v * x) as "
-  by (induct as arbitrary: x v) (auto simp add: algebra_simps)
-
-lemma foldr_mult_foldl:
-  fixes v :: "'a::comm_ring_1"
-  shows "foldr (\<lambda>x r. r * f x) as v = foldl (\<lambda>r x. r * f x) v as"
-  by (induct as arbitrary: v) (auto simp add: foldl_mult_start)
-
-lemma fps_hypergeo_nth_alt:
-  "fps_hypergeo as bs c $ n = foldr (\<lambda>a r. r * pochhammer a n) as (c ^ n) /
-    foldr (\<lambda>b r. r * pochhammer b n) bs (of_nat (fact n))"
-  by (simp add: foldl_mult_start foldr_mult_foldl)
-
-lemma fps_hypergeo_fps_exp[simp]: "fps_hypergeo [] [] c = fps_exp c"
-  by (simp add: fps_eq_iff)
-
-lemma fps_hypergeo_1_0[simp]: "fps_hypergeo [1] [] c = 1/(1 - fps_const c * X)"
-proof -
-  let ?a = "(Abs_fps (\<lambda>n. 1)) oo (fps_const c * X)"
-  have th0: "(fps_const c * X) $ 0 = 0" by simp
-  show ?thesis unfolding gp[OF th0, symmetric]
-    by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
-      fps_compose_nth power_mult_distrib cond_value_iff sum.delta' cong del: if_weak_cong)
-qed
-
-lemma fps_hypergeo_B[simp]: "fps_hypergeo [-a] [] (- 1) = fps_binomial a"
-  by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps)
-
-lemma fps_hypergeo_0[simp]: "fps_hypergeo as bs c $ 0 = 1"
-  apply simp
-  apply (subgoal_tac "\<forall>as. foldl (\<lambda>(r::'a) (a::'a). r) 1 as = 1")
-  apply auto
-  apply (induct_tac as)
-  apply auto
-  done
-
-lemma foldl_prod_prod:
-  "foldl (\<lambda>(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (\<lambda>r x. r * g x) w as =
-    foldl (\<lambda>r x. r * f x * g x) (v * w) as"
-  by (induct as arbitrary: v w) (auto simp add: algebra_simps)
-
-
-lemma fps_hypergeo_rec:
-  "fps_hypergeo as bs c $ Suc n = ((foldl (\<lambda>r a. r* (a + of_nat n)) c as) /
-    (foldl (\<lambda>r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * fps_hypergeo as bs c $ n"
-  apply (simp del: of_nat_Suc of_nat_add fact_Suc)
-  apply (simp add: foldl_mult_start del: fact_Suc of_nat_Suc)
-  unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc
-  apply (simp add: algebra_simps)
-  done
-
-lemma XD_nth[simp]: "XD a $ n = (if n = 0 then 0 else of_nat n * a$n)"
-  by (simp add: XD_def)
-
-lemma XD_0th[simp]: "XD a $ 0 = 0"
-  by simp
-lemma XD_Suc[simp]:" XD a $ Suc n = of_nat (Suc n) * a $ Suc n"
-  by simp
-
-definition "XDp c a = XD a + fps_const c * a"
-
-lemma XDp_nth[simp]: "XDp c a $ n = (c + of_nat n) * a$n"
-  by (simp add: XDp_def algebra_simps)
-
-lemma XDp_commute: "XDp b \<circ> XDp (c::'a::comm_ring_1) = XDp c \<circ> XDp b"
-  by (auto simp add: XDp_def fun_eq_iff fps_eq_iff algebra_simps)
-
-lemma XDp0 [simp]: "XDp 0 = XD"
-  by (simp add: fun_eq_iff fps_eq_iff)
-
-lemma XDp_fps_integral [simp]: "XDp 0 (fps_integral a c) = X * a"
-  by (simp add: fps_eq_iff fps_integral_def)
-
-lemma fps_hypergeo_minus_nat:
-  "fps_hypergeo [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0,field}) $ k =
-    (if k \<le> n then
-      pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k))
-     else 0)"
-  "fps_hypergeo [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0,field}) $ k =
-    (if k \<le> m then
-      pochhammer (- of_nat m) k * c ^ k / (pochhammer (- of_nat (m + n)) k * of_nat (fact k))
-     else 0)"
-  by (auto simp add: pochhammer_eq_0_iff)
-
-lemma sum_eq_if: "sum f {(n::nat) .. m} = (if m < n then 0 else f n + sum f {n+1 .. m})"
-  apply simp
-  apply (subst sum.insert[symmetric])
-  apply (auto simp add: not_less sum_head_Suc)
-  done
-
-lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))"
-  by (cases n) (simp_all add: pochhammer_rec)
-
-lemma XDp_foldr_nth [simp]: "foldr (\<lambda>c r. XDp c \<circ> r) cs (\<lambda>c. XDp c a) c0 $ n =
-    foldr (\<lambda>c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n"
-  by (induct cs arbitrary: c0) (auto simp add: algebra_simps)
-
-lemma genric_XDp_foldr_nth:
-  assumes f: "\<forall>n c a. f c a $ n = (of_nat n + k c) * a$n"
-  shows "foldr (\<lambda>c r. f c \<circ> r) cs (\<lambda>c. g c a) c0 $ n =
-    foldr (\<lambda>c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
-  by (induct cs arbitrary: c0) (auto simp add: algebra_simps f)
-
-lemma dist_less_imp_nth_equal:
-  assumes "dist f g < inverse (2 ^ i)"
-    and"j \<le> i"
-  shows "f $ j = g $ j"
-proof (rule ccontr)
-  assume "f $ j \<noteq> g $ j"
-  hence "f \<noteq> g" by auto
-  with assms have "i < subdegree (f - g)"
-    by (simp add: if_split_asm dist_fps_def)
-  also have "\<dots> \<le> j"
-    using \<open>f $ j \<noteq> g $ j\<close> by (intro subdegree_leI) simp_all
-  finally show False using \<open>j \<le> i\<close> by simp
-qed
-
-lemma nth_equal_imp_dist_less:
-  assumes "\<And>j. j \<le> i \<Longrightarrow> f $ j = g $ j"
-  shows "dist f g < inverse (2 ^ i)"
-proof (cases "f = g")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  with assms have "dist f g = inverse (2 ^ subdegree (f - g))"
-    by (simp add: if_split_asm dist_fps_def)
-  moreover
-  from assms and False have "i < subdegree (f - g)"
-    by (intro subdegree_greaterI) simp_all
-  ultimately show ?thesis by simp
-qed
-
-lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \<longleftrightarrow> (\<forall>j \<le> i. f $ j = g $ j)"
-  using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast
-
-instance fps :: (comm_ring_1) complete_space
-proof
-  fix X :: "nat \<Rightarrow> 'a fps"
-  assume "Cauchy X"
-  obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j"
-  proof -
-    have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. X M $ j = X m $ j" for i
-    proof -
-      have "0 < inverse ((2::real)^i)" by simp
-      from metric_CauchyD[OF \<open>Cauchy X\<close> this] dist_less_imp_nth_equal
-      show ?thesis by blast
-    qed
-    then show ?thesis using that by metis
-  qed
-
-  show "convergent X"
-  proof (rule convergentI)
-    show "X \<longlonglongrightarrow> Abs_fps (\<lambda>i. X (M i) $ i)"
-      unfolding tendsto_iff
-    proof safe
-      fix e::real assume e: "0 < e"
-      have "(\<lambda>n. inverse (2 ^ n) :: real) \<longlonglongrightarrow> 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all
-      from this and e have "eventually (\<lambda>i. inverse (2 ^ i) < e) sequentially"
-        by (rule order_tendstoD)
-      then obtain i where "inverse (2 ^ i) < e"
-        by (auto simp: eventually_sequentially)
-      have "eventually (\<lambda>x. M i \<le> x) sequentially"
-        by (auto simp: eventually_sequentially)
-      then show "eventually (\<lambda>x. dist (X x) (Abs_fps (\<lambda>i. X (M i) $ i)) < e) sequentially"
-      proof eventually_elim
-        fix x
-        assume x: "M i \<le> x"
-        have "X (M i) $ j = X (M j) $ j" if "j \<le> i" for j
-          using M that by (metis nat_le_linear)
-        with x have "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < inverse (2 ^ i)"
-          using M by (force simp: dist_less_eq_nth_equal)
-        also note \<open>inverse (2 ^ i) < e\<close>
-        finally show "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < e" .
-      qed
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Library/Fraction_Field.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,476 +0,0 @@
-(*  Title:      HOL/Library/Fraction_Field.thy
-    Author:     Amine Chaieb, University of Cambridge
-*)
-
-section\<open>A formalization of the fraction field of any integral domain;
-         generalization of theory Rat from int to any integral domain\<close>
-
-theory Fraction_Field
-imports Main
-begin
-
-subsection \<open>General fractions construction\<close>
-
-subsubsection \<open>Construction of the type of fractions\<close>
-
-context idom begin
-
-definition fractrel :: "'a \<times> 'a \<Rightarrow> 'a * 'a \<Rightarrow> bool" where
-  "fractrel = (\<lambda>x y. snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"
-
-lemma fractrel_iff [simp]:
-  "fractrel x y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
-  by (simp add: fractrel_def)
-
-lemma symp_fractrel: "symp fractrel"
-  by (simp add: symp_def)
-
-lemma transp_fractrel: "transp fractrel"
-proof (rule transpI, unfold split_paired_all)
-  fix a b a' b' a'' b'' :: 'a
-  assume A: "fractrel (a, b) (a', b')"
-  assume B: "fractrel (a', b') (a'', b'')"
-  have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps)
-  also from A have "a * b' = a' * b" by auto
-  also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps)
-  also from B have "a' * b'' = a'' * b'" by auto
-  also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: ac_simps)
-  finally have "b' * (a * b'') = b' * (a'' * b)" .
-  moreover from B have "b' \<noteq> 0" by auto
-  ultimately have "a * b'' = a'' * b" by simp
-  with A B show "fractrel (a, b) (a'', b'')" by auto
-qed
-
-lemma part_equivp_fractrel: "part_equivp fractrel"
-using _ symp_fractrel transp_fractrel
-by(rule part_equivpI)(rule exI[where x="(0, 1)"]; simp)
-
-end
-
-quotient_type (overloaded) 'a fract = "'a :: idom \<times> 'a" / partial: "fractrel"
-by(rule part_equivp_fractrel)
-
-subsubsection \<open>Representation and basic operations\<close>
-
-lift_definition Fract :: "'a :: idom \<Rightarrow> 'a \<Rightarrow> 'a fract"
-  is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)"
-  by simp
-
-lemma Fract_cases [cases type: fract]:
-  obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0"
-by transfer simp
-
-lemma Fract_induct [case_names Fract, induct type: fract]:
-  "(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q"
-  by (cases q) simp
-
-lemma eq_fract:
-  shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
-    and "\<And>a. Fract a 0 = Fract 0 1"
-    and "\<And>a c. Fract 0 a = Fract 0 c"
-by(transfer; simp)+
-
-instantiation fract :: (idom) comm_ring_1
-begin
-
-lift_definition zero_fract :: "'a fract" is "(0, 1)" by simp
-
-lemma Zero_fract_def: "0 = Fract 0 1"
-by transfer simp
-
-lift_definition one_fract :: "'a fract" is "(1, 1)" by simp
-
-lemma One_fract_def: "1 = Fract 1 1"
-by transfer simp
-
-lift_definition plus_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
-  is "\<lambda>q r. (fst q * snd r + fst r * snd q, snd q * snd r)"
-by(auto simp add: algebra_simps)
-
-lemma add_fract [simp]:
-  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
-by transfer simp
-
-lift_definition uminus_fract :: "'a fract \<Rightarrow> 'a fract"
-  is "\<lambda>x. (- fst x, snd x)"
-by simp
-
-lemma minus_fract [simp]:
-  fixes a b :: "'a::idom"
-  shows "- Fract a b = Fract (- a) b"
-by transfer simp
-
-lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
-  by (cases "b = 0") (simp_all add: eq_fract)
-
-definition diff_fract_def: "q - r = q + - (r::'a fract)"
-
-lemma diff_fract [simp]:
-  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
-  by (simp add: diff_fract_def)
-
-lift_definition times_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
-  is "\<lambda>q r. (fst q * fst r, snd q * snd r)"
-by(simp add: algebra_simps)
-
-lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"
-by transfer simp
-
-lemma mult_fract_cancel:
-  "c \<noteq> 0 \<Longrightarrow> Fract (c * a) (c * b) = Fract a b"
-by transfer simp
-
-instance
-proof
-  fix q r s :: "'a fract"
-  show "(q * r) * s = q * (r * s)"
-    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
-  show "q * r = r * q"
-    by (cases q, cases r) (simp add: eq_fract algebra_simps)
-  show "1 * q = q"
-    by (cases q) (simp add: One_fract_def eq_fract)
-  show "(q + r) + s = q + (r + s)"
-    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
-  show "q + r = r + q"
-    by (cases q, cases r) (simp add: eq_fract algebra_simps)
-  show "0 + q = q"
-    by (cases q) (simp add: Zero_fract_def eq_fract)
-  show "- q + q = 0"
-    by (cases q) (simp add: Zero_fract_def eq_fract)
-  show "q - r = q + - r"
-    by (cases q, cases r) (simp add: eq_fract)
-  show "(q + r) * s = q * s + r * s"
-    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
-  show "(0::'a fract) \<noteq> 1"
-    by (simp add: Zero_fract_def One_fract_def eq_fract)
-qed
-
-end
-
-lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1"
-  by (induct k) (simp_all add: Zero_fract_def One_fract_def)
-
-lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
-  by (rule of_nat_fract [symmetric])
-
-lemma fract_collapse:
-  "Fract 0 k = 0"
-  "Fract 1 1 = 1"
-  "Fract k 0 = 0"
-by(transfer; simp)+
-
-lemma fract_expand:
-  "0 = Fract 0 1"
-  "1 = Fract 1 1"
-  by (simp_all add: fract_collapse)
-
-lemma Fract_cases_nonzero:
-  obtains (Fract) a b where "q = Fract a b" and "b \<noteq> 0" and "a \<noteq> 0"
-    | (0) "q = 0"
-proof (cases "q = 0")
-  case True
-  then show thesis using 0 by auto
-next
-  case False
-  then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
-  with False have "0 \<noteq> Fract a b" by simp
-  with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
-  with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto
-qed
-
-
-subsubsection \<open>The field of rational numbers\<close>
-
-context idom
-begin
-
-subclass ring_no_zero_divisors ..
-
-end
-
-instantiation fract :: (idom) field
-begin
-
-lift_definition inverse_fract :: "'a fract \<Rightarrow> 'a fract"
-  is "\<lambda>x. if fst x = 0 then (0, 1) else (snd x, fst x)"
-by(auto simp add: algebra_simps)
-
-lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"
-by transfer simp
-
-definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)"
-
-lemma divide_fract [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)"
-  by (simp add: divide_fract_def)
-
-instance
-proof
-  fix q :: "'a fract"
-  assume "q \<noteq> 0"
-  then show "inverse q * q = 1"
-    by (cases q rule: Fract_cases_nonzero)
-      (simp_all add: fract_expand eq_fract mult.commute)
-next
-  fix q r :: "'a fract"
-  show "q div r = q * inverse r" by (simp add: divide_fract_def)
-next
-  show "inverse 0 = (0:: 'a fract)"
-    by (simp add: fract_expand) (simp add: fract_collapse)
-qed
-
-end
-
-
-subsubsection \<open>The ordered field of fractions over an ordered idom\<close>
-
-instantiation fract :: (linordered_idom) linorder
-begin
-
-lemma less_eq_fract_respect:
-  fixes a b a' b' c d c' d' :: 'a
-  assumes neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-  assumes eq1: "a * b' = a' * b"
-  assumes eq2: "c * d' = c' * d"
-  shows "((a * d) * (b * d) \<le> (c * b) * (b * d)) \<longleftrightarrow> ((a' * d') * (b' * d') \<le> (c' * b') * (b' * d'))"
-proof -
-  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
-  {
-    fix a b c d x :: 'a
-    assume x: "x \<noteq> 0"
-    have "?le a b c d = ?le (a * x) (b * x) c d"
-    proof -
-      from x have "0 < x * x"
-        by (auto simp add: zero_less_mult_iff)
-      then have "?le a b c d =
-          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
-        by (simp add: mult_le_cancel_right)
-      also have "... = ?le (a * x) (b * x) c d"
-        by (simp add: ac_simps)
-      finally show ?thesis .
-    qed
-  } note le_factor = this
-
-  let ?D = "b * d" and ?D' = "b' * d'"
-  from neq have D: "?D \<noteq> 0" by simp
-  from neq have "?D' \<noteq> 0" by simp
-  then have "?le a b c d = ?le (a * ?D') (b * ?D') c d"
-    by (rule le_factor)
-  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
-    by (simp add: ac_simps)
-  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
-    by (simp only: eq1 eq2)
-  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
-    by (simp add: ac_simps)
-  also from D have "... = ?le a' b' c' d'"
-    by (rule le_factor [symmetric])
-  finally show "?le a b c d = ?le a' b' c' d'" .
-qed
-
-lift_definition less_eq_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> bool"
-  is "\<lambda>q r. (fst q * snd r) * (snd q * snd r) \<le> (fst r * snd q) * (snd q * snd r)"
-by (clarsimp simp add: less_eq_fract_respect)
-
-definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
-
-lemma le_fract [simp]:
-  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
-  by transfer simp
-
-lemma less_fract [simp]:
-  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
-  by (simp add: less_fract_def less_le_not_le ac_simps)
-
-instance
-proof
-  fix q r s :: "'a fract"
-  assume "q \<le> r" and "r \<le> s"
-  then show "q \<le> s"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: 'a
-    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
-    assume 1: "Fract a b \<le> Fract c d"
-    assume 2: "Fract c d \<le> Fract e f"
-    show "Fract a b \<le> Fract e f"
-    proof -
-      from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
-        by (auto simp add: zero_less_mult_iff linorder_neq_iff)
-      have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
-      proof -
-        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-          by simp
-        with ff show ?thesis by (simp add: mult_le_cancel_right)
-      qed
-      also have "... = (c * f) * (d * f) * (b * b)"
-        by (simp only: ac_simps)
-      also have "... \<le> (e * d) * (d * f) * (b * b)"
-      proof -
-        from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
-          by simp
-        with bb show ?thesis by (simp add: mult_le_cancel_right)
-      qed
-      finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
-        by (simp only: ac_simps)
-      with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
-        by (simp add: mult_le_cancel_right)
-      with neq show ?thesis by simp
-    qed
-  qed
-next
-  fix q r :: "'a fract"
-  assume "q \<le> r" and "r \<le> q"
-  then show "q = r"
-  proof (induct q, induct r)
-    fix a b c d :: 'a
-    assume neq: "b \<noteq> 0" "d \<noteq> 0"
-    assume 1: "Fract a b \<le> Fract c d"
-    assume 2: "Fract c d \<le> Fract a b"
-    show "Fract a b = Fract c d"
-    proof -
-      from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-        by simp
-      also have "... \<le> (a * d) * (b * d)"
-      proof -
-        from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
-          by simp
-        then show ?thesis by (simp only: ac_simps)
-      qed
-      finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
-      moreover from neq have "b * d \<noteq> 0" by simp
-      ultimately have "a * d = c * b" by simp
-      with neq show ?thesis by (simp add: eq_fract)
-    qed
-  qed
-next
-  fix q r :: "'a fract"
-  show "q \<le> q"
-    by (induct q) simp
-  show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
-    by (simp only: less_fract_def)
-  show "q \<le> r \<or> r \<le> q"
-    by (induct q, induct r)
-       (simp add: mult.commute, rule linorder_linear)
-qed
-
-end
-
-instantiation fract :: (linordered_idom) linordered_field
-begin
-
-definition abs_fract_def2:
-  "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
-
-definition sgn_fract_def:
-  "sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
-
-theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
-  unfolding abs_fract_def2 not_le [symmetric]
-  by transfer (auto simp add: zero_less_mult_iff le_less)
-
-instance proof
-  fix q r s :: "'a fract"
-  assume "q \<le> r"
-  then show "s + q \<le> s + r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: 'a
-    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
-    assume le: "Fract a b \<le> Fract c d"
-    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
-    proof -
-      let ?F = "f * f" from neq have F: "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-        by simp
-      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
-        by (simp add: mult_le_cancel_right)
-      with neq show ?thesis by (simp add: field_simps)
-    qed
-  qed
-next
-  fix q r s :: "'a fract"
-  assume "q < r" and "0 < s"
-  then show "s * q < s * r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: 'a
-    assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0"
-    assume le: "Fract a b < Fract c d"
-    assume gt: "0 < Fract e f"
-    show "Fract e f * Fract a b < Fract e f * Fract c d"
-    proof -
-      let ?E = "e * f" and ?F = "f * f"
-      from neq gt have "0 < ?E"
-        by (auto simp add: Zero_fract_def order_less_le eq_fract)
-      moreover from neq have "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
-        by simp
-      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
-        by (simp add: mult_less_cancel_right)
-      with neq show ?thesis
-        by (simp add: ac_simps)
-    qed
-  qed
-qed (fact sgn_fract_def abs_fract_def2)+
-
-end
-
-instantiation fract :: (linordered_idom) distrib_lattice
-begin
-
-definition inf_fract_def:
-  "(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
-
-definition sup_fract_def:
-  "(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
-
-instance
-  by standard (simp_all add: inf_fract_def sup_fract_def max_min_distrib2)
-  
-end
-
-lemma fract_induct_pos [case_names Fract]:
-  fixes P :: "'a::linordered_idom fract \<Rightarrow> bool"
-  assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
-  shows "P q"
-proof (cases q)
-  case (Fract a b)
-  {
-    fix a b :: 'a
-    assume b: "b < 0"
-    have "P (Fract a b)"
-    proof -
-      from b have "0 < - b" by simp
-      then have "P (Fract (- a) (- b))"
-        by (rule step)
-      then show "P (Fract a b)"
-        by (simp add: order_less_imp_not_eq [OF b])
-    qed
-  }
-  with Fract show "P q"
-    by (auto simp add: linorder_neq_iff step)
-qed
-
-lemma zero_less_Fract_iff: "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
-  by (auto simp add: Zero_fract_def zero_less_mult_iff)
-
-lemma Fract_less_zero_iff: "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
-  by (auto simp add: Zero_fract_def mult_less_0_iff)
-
-lemma zero_le_Fract_iff: "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
-  by (auto simp add: Zero_fract_def zero_le_mult_iff)
-
-lemma Fract_le_zero_iff: "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
-  by (auto simp add: Zero_fract_def mult_le_0_iff)
-
-lemma one_less_Fract_iff: "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
-  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
-
-lemma Fract_less_one_iff: "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
-  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
-
-lemma one_le_Fract_iff: "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
-  by (auto simp add: One_fract_def mult_le_cancel_right)
-
-lemma Fract_le_one_iff: "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
-  by (auto simp add: One_fract_def mult_le_cancel_right)
-
-end
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1207 +0,0 @@
-(* Author: Amine Chaieb, TU Muenchen *)
-
-section \<open>Fundamental Theorem of Algebra\<close>
-
-theory Fundamental_Theorem_Algebra
-imports Polynomial Complex_Main
-begin
-
-subsection \<open>More lemmas about module of complex numbers\<close>
-
-text \<open>The triangle inequality for cmod\<close>
-
-lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
-  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
-
-
-subsection \<open>Basic lemmas about polynomials\<close>
-
-lemma poly_bound_exists:
-  fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
-  shows "\<exists>m. m > 0 \<and> (\<forall>z. norm z \<le> r \<longrightarrow> norm (poly p z) \<le> m)"
-proof (induct p)
-  case 0
-  then show ?case by (rule exI[where x=1]) simp
-next
-  case (pCons c cs)
-  from pCons.hyps obtain m where m: "\<forall>z. norm z \<le> r \<longrightarrow> norm (poly cs z) \<le> m"
-    by blast
-  let ?k = " 1 + norm c + \<bar>r * m\<bar>"
-  have kp: "?k > 0"
-    using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
-  have "norm (poly (pCons c cs) z) \<le> ?k" if H: "norm z \<le> r" for z
-  proof -
-    from m H have th: "norm (poly cs z) \<le> m"
-      by blast
-    from H have rp: "r \<ge> 0"
-      using norm_ge_zero[of z] by arith
-    have "norm (poly (pCons c cs) z) \<le> norm c + norm (z * poly cs z)"
-      using norm_triangle_ineq[of c "z* poly cs z"] by simp
-    also have "\<dots> \<le> norm c + r * m"
-      using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]]
-      by (simp add: norm_mult)
-    also have "\<dots> \<le> ?k"
-      by simp
-    finally show ?thesis .
-  qed
-  with kp show ?case by blast
-qed
-
-
-text \<open>Offsetting the variable in a polynomial gives another of same degree\<close>
-
-definition offset_poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
-  where "offset_poly p h = fold_coeffs (\<lambda>a q. smult h q + pCons a q) p 0"
-
-lemma offset_poly_0: "offset_poly 0 h = 0"
-  by (simp add: offset_poly_def)
-
-lemma offset_poly_pCons:
-  "offset_poly (pCons a p) h =
-    smult h (offset_poly p h) + pCons a (offset_poly p h)"
-  by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def)
-
-lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
-  by (simp add: offset_poly_pCons offset_poly_0)
-
-lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
-  apply (induct p)
-  apply (simp add: offset_poly_0)
-  apply (simp add: offset_poly_pCons algebra_simps)
-  done
-
-lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
-  by (induct p arbitrary: a) (simp, force)
-
-lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
-  apply (safe intro!: offset_poly_0)
-  apply (induct p)
-  apply simp
-  apply (simp add: offset_poly_pCons)
-  apply (frule offset_poly_eq_0_lemma, simp)
-  done
-
-lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
-  apply (induct p)
-  apply (simp add: offset_poly_0)
-  apply (case_tac "p = 0")
-  apply (simp add: offset_poly_0 offset_poly_pCons)
-  apply (simp add: offset_poly_pCons)
-  apply (subst degree_add_eq_right)
-  apply (rule le_less_trans [OF degree_smult_le])
-  apply (simp add: offset_poly_eq_0_iff)
-  apply (simp add: offset_poly_eq_0_iff)
-  done
-
-definition "psize p = (if p = 0 then 0 else Suc (degree p))"
-
-lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
-  unfolding psize_def by simp
-
-lemma poly_offset:
-  fixes p :: "'a::comm_ring_1 poly"
-  shows "\<exists>q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))"
-proof (intro exI conjI)
-  show "psize (offset_poly p a) = psize p"
-    unfolding psize_def
-    by (simp add: offset_poly_eq_0_iff degree_offset_poly)
-  show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
-    by (simp add: poly_offset_poly)
-qed
-
-text \<open>An alternative useful formulation of completeness of the reals\<close>
-lemma real_sup_exists:
-  assumes ex: "\<exists>x. P x"
-    and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
-  shows "\<exists>s::real. \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
-proof
-  from bz have "bdd_above (Collect P)"
-    by (force intro: less_imp_le)
-  then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
-    using ex bz by (subst less_cSup_iff) auto
-qed
-
-
-subsection \<open>Fundamental theorem of algebra\<close>
-
-lemma unimodular_reduce_norm:
-  assumes md: "cmod z = 1"
-  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + \<i>) < 1 \<or> cmod (z - \<i>) < 1"
-proof -
-  obtain x y where z: "z = Complex x y "
-    by (cases z) auto
-  from md z have xy: "x\<^sup>2 + y\<^sup>2 = 1"
-    by (simp add: cmod_def)
-  have False if "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + \<i>) \<ge> 1" "cmod (z - \<i>) \<ge> 1"
-  proof -
-    from that z xy have "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1"
-      by (simp_all add: cmod_def power2_eq_square algebra_simps)
-    then have "\<bar>2 * x\<bar> \<le> 1" "\<bar>2 * y\<bar> \<le> 1"
-      by simp_all
-    then have "\<bar>2 * x\<bar>\<^sup>2 \<le> 1\<^sup>2" "\<bar>2 * y\<bar>\<^sup>2 \<le> 1\<^sup>2"
-      by - (rule power_mono, simp, simp)+
-    then have th0: "4 * x\<^sup>2 \<le> 1" "4 * y\<^sup>2 \<le> 1"
-      by (simp_all add: power_mult_distrib)
-    from add_mono[OF th0] xy show ?thesis
-      by simp
-  qed
-  then show ?thesis
-    unfolding linorder_not_le[symmetric] by blast
-qed
-
-text \<open>Hence we can always reduce modulus of \<open>1 + b z^n\<close> if nonzero\<close>
-lemma reduce_poly_simple:
-  assumes b: "b \<noteq> 0"
-    and n: "n \<noteq> 0"
-  shows "\<exists>z. cmod (1 + b * z^n) < 1"
-  using n
-proof (induct n rule: nat_less_induct)
-  fix n
-  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)"
-  assume n: "n \<noteq> 0"
-  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
-  show "\<exists>z. ?P z n"
-  proof cases
-    assume "even n"
-    then have "\<exists>m. n = 2 * m"
-      by presburger
-    then obtain m where m: "n = 2 * m"
-      by blast
-    from n m have "m \<noteq> 0" "m < n"
-      by presburger+
-    with IH[rule_format, of m] obtain z where z: "?P z m"
-      by blast
-    from z have "?P (csqrt z) n"
-      by (simp add: m power_mult)
-    then show ?thesis ..
-  next
-    assume "odd n"
-    then have "\<exists>m. n = Suc (2 * m)"
-      by presburger+
-    then obtain m where m: "n = Suc (2 * m)"
-      by blast
-    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
-      using b by (simp add: norm_divide)
-    from unimodular_reduce_norm[OF th0] \<open>odd n\<close>
-    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
-      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1")
-      apply (rule_tac x="1" in exI)
-      apply simp
-      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1")
-      apply (rule_tac x="-1" in exI)
-      apply simp
-      apply (cases "cmod (complex_of_real (cmod b) / b + \<i>) < 1")
-      apply (cases "even m")
-      apply (rule_tac x="\<i>" in exI)
-      apply (simp add: m power_mult)
-      apply (rule_tac x="- \<i>" in exI)
-      apply (simp add: m power_mult)
-      apply (cases "even m")
-      apply (rule_tac x="- \<i>" in exI)
-      apply (simp add: m power_mult)
-      apply (auto simp add: m power_mult)
-      apply (rule_tac x="\<i>" in exI)
-      apply (auto simp add: m power_mult)
-      done
-    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1"
-      by blast
-    let ?w = "v / complex_of_real (root n (cmod b))"
-    from odd_real_root_pow[OF \<open>odd n\<close>, of "cmod b"]
-    have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
-      by (simp add: power_divide of_real_power[symmetric])
-    have th2:"cmod (complex_of_real (cmod b) / b) = 1"
-      using b by (simp add: norm_divide)
-    then have th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0"
-      by simp
-    have th4: "cmod (complex_of_real (cmod b) / b) *
-        cmod (1 + b * (v ^ n / complex_of_real (cmod b))) <
-        cmod (complex_of_real (cmod b) / b) * 1"
-      apply (simp only: norm_mult[symmetric] distrib_left)
-      using b v
-      apply (simp add: th2)
-      done
-    from mult_left_less_imp_less[OF th4 th3]
-    have "?P ?w n" unfolding th1 .
-    then show ?thesis ..
-  qed
-qed
-
-text \<open>Bolzano-Weierstrass type property for closed disc in complex plane.\<close>
-
-lemma metric_bound_lemma: "cmod (x - y) \<le> \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
-  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y"]
-  unfolding cmod_def by simp
-
-lemma bolzano_weierstrass_complex_disc:
-  assumes r: "\<forall>n. cmod (s n) \<le> r"
-  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
-proof -
-  from seq_monosub[of "Re \<circ> s"]
-  obtain f where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))"
-    unfolding o_def by blast
-  from seq_monosub[of "Im \<circ> s \<circ> f"]
-  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s (f (g n))))"
-    unfolding o_def by blast
-  let ?h = "f \<circ> g"
-  from r[rule_format, of 0] have rp: "r \<ge> 0"
-    using norm_ge_zero[of "s 0"] by arith
-  have th: "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>"
-  proof
-    fix n
-    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]
-    show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
-  qed
-  have conv1: "convergent (\<lambda>n. Re (s (f n)))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (metis gt_ex le_less_linear less_trans order.trans th)
-    apply (rule f(2))
-    done
-  have th: "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>"
-  proof
-    fix n
-    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]
-    show "\<bar>Im (s n)\<bar> \<le> r + 1"
-      by arith
-  qed
-
-  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (metis gt_ex le_less_linear less_trans order.trans th)
-    apply (rule g(2))
-    done
-
-  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
-    by blast
-  then have x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Re (s (f n)) - x\<bar> < r"
-    unfolding LIMSEQ_iff real_norm_def .
-
-  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
-    by blast
-  then have y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Im (s (f (g n))) - y\<bar> < r"
-    unfolding LIMSEQ_iff real_norm_def .
-  let ?w = "Complex x y"
-  from f(1) g(1) have hs: "subseq ?h"
-    unfolding subseq_def by auto
-  have "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" if "e > 0" for e
-  proof -
-    from that have e2: "e/2 > 0"
-      by simp
-    from x[rule_format, OF e2] y[rule_format, OF e2]
-    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2"
-      and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2"
-      by blast
-    have "cmod (s (?h n) - ?w) < e" if "n \<ge> N1 + N2" for n
-    proof -
-      from that have nN1: "g n \<ge> N1" and nN2: "n \<ge> N2"
-        using seq_suble[OF g(1), of n] by arith+
-      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
-      show ?thesis
-        using metric_bound_lemma[of "s (f (g n))" ?w] by simp
-    qed
-    then show ?thesis by blast
-  qed
-  with hs show ?thesis by blast
-qed
-
-text \<open>Polynomial is continuous.\<close>
-
-lemma poly_cont:
-  fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
-  assumes ep: "e > 0"
-  shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e"
-proof -
-  obtain q where q: "degree q = degree p" "poly q x = poly p (z + x)" for x
-  proof
-    show "degree (offset_poly p z) = degree p"
-      by (rule degree_offset_poly)
-    show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
-      by (rule poly_offset_poly)
-  qed
-  have th: "\<And>w. poly q (w - z) = poly p w"
-    using q(2)[of "w - z" for w] by simp
-  show ?thesis unfolding th[symmetric]
-  proof (induct q)
-    case 0
-    then show ?case
-      using ep by auto
-  next
-    case (pCons c cs)
-    from poly_bound_exists[of 1 "cs"]
-    obtain m where m: "m > 0" "norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m" for z
-      by blast
-    from ep m(1) have em0: "e/m > 0"
-      by (simp add: field_simps)
-    have one0: "1 > (0::real)"
-      by arith
-    from real_lbound_gt_zero[OF one0 em0]
-    obtain d where d: "d > 0" "d < 1" "d < e / m"
-      by blast
-    from d(1,3) m(1) have dm: "d * m > 0" "d * m < e"
-      by (simp_all add: field_simps)
-    show ?case
-    proof (rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
-      fix d w
-      assume H: "d > 0" "d < 1" "d < e/m" "w \<noteq> z" "norm (w - z) < d"
-      then have d1: "norm (w-z) \<le> 1" "d \<ge> 0"
-        by simp_all
-      from H(3) m(1) have dme: "d*m < e"
-        by (simp add: field_simps)
-      from H have th: "norm (w - z) \<le> d"
-        by simp
-      from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
-      show "norm (w - z) * norm (poly cs (w - z)) < e"
-        by simp
-    qed
-  qed
-qed
-
-text \<open>Hence a polynomial attains minimum on a closed disc
-  in the complex plane.\<close>
-lemma poly_minimum_modulus_disc: "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
-proof -
-  show ?thesis
-  proof (cases "r \<ge> 0")
-    case False
-    then show ?thesis
-      by (metis norm_ge_zero order.trans)
-  next
-    case True
-    then have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))"
-      by simp
-    then have mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"
-      by blast
-    have False if "cmod z \<le> r" "cmod (poly p z) = - x" "\<not> x < 1" for x z
-    proof -
-      from that have "- x < 0 "
-        by arith
-      with that(2) norm_ge_zero[of "poly p z"] show ?thesis
-        by simp
-    qed
-    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z"
-      by blast
-    from real_sup_exists[OF mth1 mth2] obtain s where
-      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow> y < s"
-      by blast
-    let ?m = "- s"
-    have s1[unfolded minus_minus]:
-      "(\<exists>z x. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" for y
-      using s[rule_format, of "-y"]
-      unfolding minus_less_iff[of y] equation_minus_iff by blast
-    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
-      by auto
-    have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" for n
-      using s1[rule_format, of "?m + 1/real (Suc n)"] by simp
-    then have th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
-    from choice[OF th] obtain g where
-        g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)"
-      by blast
-    from bolzano_weierstrass_complex_disc[OF g(1)]
-    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
-      by blast
-    {
-      fix w
-      assume wr: "cmod w \<le> r"
-      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
-      {
-        assume e: "?e > 0"
-        then have e2: "?e/2 > 0"
-          by simp
-        from poly_cont[OF e2, of z p] obtain d where
-            d: "d > 0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2"
-          by blast
-        have th1: "cmod(poly p w - poly p z) < ?e / 2" if w: "cmod (w - z) < d" for w
-          using d(2)[rule_format, of w] w e by (cases "w = z") simp_all
-        from fz(2) d(1) obtain N1 where N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d"
-          by blast
-        from reals_Archimedean2[of "2/?e"] obtain N2 :: nat where N2: "2/?e < real N2"
-          by blast
-        have th2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2"
-          using N1[rule_format, of "N1 + N2"] th1 by simp
-        have th0: "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False"
-          for a b e2 m :: real
-          by arith
-        have ath: "m \<le> x \<Longrightarrow> x < m + e \<Longrightarrow> \<bar>x - m\<bar> < e" for m x e :: real
-          by arith
-        from s1m[OF g(1)[rule_format]] have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
-        from seq_suble[OF fz(1), of "N1 + N2"]
-        have th00: "real (Suc (N1 + N2)) \<le> real (Suc (f (N1 + N2)))"
-          by simp
-        have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1 + N2)) > 0"
-          using N2 by auto
-        from frac_le[OF th000 th00]
-        have th00: "?m + 1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))"
-          by simp
-        from g(2)[rule_format, of "f (N1 + N2)"]
-        have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
-        from order_less_le_trans[OF th01 th00]
-        have th32: "cmod (poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
-        from N2 have "2/?e < real (Suc (N1 + N2))"
-          by arith
-        with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
-        have "?e/2 > 1/ real (Suc (N1 + N2))"
-          by (simp add: inverse_eq_divide)
-        with ath[OF th31 th32] have thc1: "\<bar>cmod (poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2"
-          by arith
-        have ath2: "\<bar>a - b\<bar> \<le> c \<Longrightarrow> \<bar>b - m\<bar> \<le> \<bar>a - m\<bar> + c" for a b c m :: real
-          by arith
-        have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar> \<le>
-            cmod (poly p (g (f (N1 + N2))) - poly p z)"
-          by (simp add: norm_triangle_ineq3)
-        from ath2[OF th22, of ?m]
-        have thc2: "2 * (?e/2) \<le>
-            \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)"
-          by simp
-        from th0[OF th2 thc1 thc2] have False .
-      }
-      then have "?e = 0"
-        by auto
-      then have "cmod (poly p z) = ?m"
-        by simp
-      with s1m[OF wr] have "cmod (poly p z) \<le> cmod (poly p w)"
-        by simp
-    }
-    then show ?thesis by blast
-  qed
-qed
-
-text \<open>Nonzero polynomial in z goes to infinity as z does.\<close>
-
-lemma poly_infinity:
-  fixes p:: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
-  assumes ex: "p \<noteq> 0"
-  shows "\<exists>r. \<forall>z. r \<le> norm z \<longrightarrow> d \<le> norm (poly (pCons a p) z)"
-  using ex
-proof (induct p arbitrary: a d)
-  case 0
-  then show ?case by simp
-next
-  case (pCons c cs a d)
-  show ?case
-  proof (cases "cs = 0")
-    case False
-    with pCons.hyps obtain r where r: "\<forall>z. r \<le> norm z \<longrightarrow> d + norm a \<le> norm (poly (pCons c cs) z)"
-      by blast
-    let ?r = "1 + \<bar>r\<bar>"
-    have "d \<le> norm (poly (pCons a (pCons c cs)) z)" if "1 + \<bar>r\<bar> \<le> norm z" for z
-    proof -
-      have r0: "r \<le> norm z"
-        using that by arith
-      from r[rule_format, OF r0] have th0: "d + norm a \<le> 1 * norm(poly (pCons c cs) z)"
-        by arith
-      from that have z1: "norm z \<ge> 1"
-        by arith
-      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
-      have th1: "d \<le> norm(z * poly (pCons c cs) z) - norm a"
-        unfolding norm_mult by (simp add: algebra_simps)
-      from norm_diff_ineq[of "z * poly (pCons c cs) z" a]
-      have th2: "norm (z * poly (pCons c cs) z) - norm a \<le> norm (poly (pCons a (pCons c cs)) z)"
-        by (simp add: algebra_simps)
-      from th1 th2 show ?thesis
-        by arith
-    qed
-    then show ?thesis by blast
-  next
-    case True
-    with pCons.prems have c0: "c \<noteq> 0"
-      by simp
-    have "d \<le> norm (poly (pCons a (pCons c cs)) z)"
-      if h: "(\<bar>d\<bar> + norm a) / norm c \<le> norm z" for z :: 'a
-    proof -
-      from c0 have "norm c > 0"
-        by simp
-      from h c0 have th0: "\<bar>d\<bar> + norm a \<le> norm (z * c)"
-        by (simp add: field_simps norm_mult)
-      have ath: "\<And>mzh mazh ma. mzh \<le> mazh + ma \<Longrightarrow> \<bar>d\<bar> + ma \<le> mzh \<Longrightarrow> d \<le> mazh"
-        by arith
-      from norm_diff_ineq[of "z * c" a] have th1: "norm (z * c) \<le> norm (a + z * c) + norm a"
-        by (simp add: algebra_simps)
-      from ath[OF th1 th0] show ?thesis
-        using True by simp
-    qed
-    then show ?thesis by blast
-  qed
-qed
-
-text \<open>Hence polynomial's modulus attains its minimum somewhere.\<close>
-lemma poly_minimum_modulus: "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
-proof (induct p)
-  case 0
-  then show ?case by simp
-next
-  case (pCons c cs)
-  show ?case
-  proof (cases "cs = 0")
-    case False
-    from poly_infinity[OF False, of "cmod (poly (pCons c cs) 0)" c]
-    obtain r where r: "cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)"
-      if "r \<le> cmod z" for z
-      by blast
-    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>"
-      by arith
-    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
-    obtain v where v: "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)"
-      if "cmod w \<le> \<bar>r\<bar>" for w
-      by blast
-    have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)" if z: "r \<le> cmod z" for z
-      using v[of 0] r[OF z] by simp
-    with v ath[of r] show ?thesis
-      by blast
-  next
-    case True
-    with pCons.hyps show ?thesis
-      by simp
-  qed
-qed
-
-text \<open>Constant function (non-syntactic characterization).\<close>
-definition "constant f \<longleftrightarrow> (\<forall>x y. f x = f y)"
-
-lemma nonconstant_length: "\<not> constant (poly p) \<Longrightarrow> psize p \<ge> 2"
-  by (induct p) (auto simp: constant_def psize_def)
-
-lemma poly_replicate_append: "poly (monom 1 n * p) (x::'a::comm_ring_1) = x^n * poly p x"
-  by (simp add: poly_monom)
-
-text \<open>Decomposition of polynomial, skipping zero coefficients after the first.\<close>
-
-lemma poly_decompose_lemma:
-  assumes nz: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly p z = (0::'a::idom))"
-  shows "\<exists>k a q. a \<noteq> 0 \<and> Suc (psize q + k) = psize p \<and> (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
-  unfolding psize_def
-  using nz
-proof (induct p)
-  case 0
-  then show ?case by simp
-next
-  case (pCons c cs)
-  show ?case
-  proof (cases "c = 0")
-    case True
-    from pCons.hyps pCons.prems True show ?thesis
-      apply auto
-      apply (rule_tac x="k+1" in exI)
-      apply (rule_tac x="a" in exI)
-      apply clarsimp
-      apply (rule_tac x="q" in exI)
-      apply auto
-      done
-  next
-    case False
-    show ?thesis
-      apply (rule exI[where x=0])
-      apply (rule exI[where x=c])
-      apply (auto simp: False)
-      done
-  qed
-qed
-
-lemma poly_decompose:
-  assumes nc: "\<not> constant (poly p)"
-  shows "\<exists>k a q. a \<noteq> (0::'a::idom) \<and> k \<noteq> 0 \<and>
-               psize q + k + 1 = psize p \<and>
-              (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
-  using nc
-proof (induct p)
-  case 0
-  then show ?case
-    by (simp add: constant_def)
-next
-  case (pCons c cs)
-  have "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)"
-  proof
-    assume "\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
-    then have "poly (pCons c cs) x = poly (pCons c cs) y" for x y
-      by (cases "x = 0") auto
-    with pCons.prems show False
-      by (auto simp add: constant_def)
-  qed
-  from poly_decompose_lemma[OF this]
-  show ?case
-    apply clarsimp
-    apply (rule_tac x="k+1" in exI)
-    apply (rule_tac x="a" in exI)
-    apply simp
-    apply (rule_tac x="q" in exI)
-    apply (auto simp add: psize_def split: if_splits)
-    done
-qed
-
-text \<open>Fundamental theorem of algebra\<close>
-
-lemma fundamental_theorem_of_algebra:
-  assumes nc: "\<not> constant (poly p)"
-  shows "\<exists>z::complex. poly p z = 0"
-  using nc
-proof (induct "psize p" arbitrary: p rule: less_induct)
-  case less
-  let ?p = "poly p"
-  let ?ths = "\<exists>z. ?p z = 0"
-
-  from nonconstant_length[OF less(2)] have n2: "psize p \<ge> 2" .
-  from poly_minimum_modulus obtain c where c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)"
-    by blast
-
-  show ?ths
-  proof (cases "?p c = 0")
-    case True
-    then show ?thesis by blast
-  next
-    case False
-    from poly_offset[of p c] obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)"
-      by blast
-    have False if h: "constant (poly q)"
-    proof -
-      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x"
-        by auto
-      have "?p x = ?p y" for x y
-      proof -
-        from th have "?p x = poly q (x - c)"
-          by auto
-        also have "\<dots> = poly q (y - c)"
-          using h unfolding constant_def by blast
-        also have "\<dots> = ?p y"
-          using th by auto
-        finally show ?thesis .
-      qed
-      with less(2) show ?thesis
-        unfolding constant_def by blast
-    qed
-    then have qnc: "\<not> constant (poly q)"
-      by blast
-    from q(2) have pqc0: "?p c = poly q 0"
-      by simp
-    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)"
-      by simp
-    let ?a0 = "poly q 0"
-    from False pqc0 have a00: "?a0 \<noteq> 0"
-      by simp
-    from a00 have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
-      by simp
-    let ?r = "smult (inverse ?a0) q"
-    have lgqr: "psize q = psize ?r"
-      using a00
-      unfolding psize_def degree_def
-      by (simp add: poly_eq_iff)
-    have False if h: "\<And>x y. poly ?r x = poly ?r y"
-    proof -
-      have "poly q x = poly q y" for x y
-      proof -
-        from qr[rule_format, of x] have "poly q x = poly ?r x * ?a0"
-          by auto
-        also have "\<dots> = poly ?r y * ?a0"
-          using h by simp
-        also have "\<dots> = poly q y"
-          using qr[rule_format, of y] by simp
-        finally show ?thesis .
-      qed
-      with qnc show ?thesis
-        unfolding constant_def by blast
-    qed
-    then have rnc: "\<not> constant (poly ?r)"
-      unfolding constant_def by blast
-    from qr[rule_format, of 0] a00 have r01: "poly ?r 0 = 1"
-      by auto
-    have mrmq_eq: "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" for w
-    proof -
-      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
-        using qr[rule_format, of w] a00 by (simp add: divide_inverse ac_simps)
-      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
-        using a00 unfolding norm_divide by (simp add: field_simps)
-      finally show ?thesis .
-    qed
-    from poly_decompose[OF rnc] obtain k a s where
-      kas: "a \<noteq> 0" "k \<noteq> 0" "psize s + k + 1 = psize ?r"
-        "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
-    have "\<exists>w. cmod (poly ?r w) < 1"
-    proof (cases "psize p = k + 1")
-      case True
-      with kas(3) lgqr[symmetric] q(1) have s0: "s = 0"
-        by auto
-      have hth[symmetric]: "cmod (poly ?r w) = cmod (1 + a * w ^ k)" for w
-        using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)
-      from reduce_poly_simple[OF kas(1,2)] show ?thesis
-        unfolding hth by blast
-    next
-      case False note kn = this
-      from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p"
-        by simp
-      have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
-        unfolding constant_def poly_pCons poly_monom
-        using kas(1)
-        apply simp
-        apply (rule exI[where x=0])
-        apply (rule exI[where x=1])
-        apply simp
-        done
-      from kas(1) kas(2) have th02: "k + 1 = psize (pCons 1 (monom a (k - 1)))"
-        by (simp add: psize_def degree_monom_eq)
-      from less(1) [OF k1n [simplified th02] th01]
-      obtain w where w: "1 + w^k * a = 0"
-        unfolding poly_pCons poly_monom
-        using kas(2) by (cases k) (auto simp add: algebra_simps)
-      from poly_bound_exists[of "cmod w" s] obtain m where
-        m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
-      have w0: "w \<noteq> 0"
-        using kas(2) w by (auto simp add: power_0_left)
-      from w have "(1 + w ^ k * a) - 1 = 0 - 1"
-        by simp
-      then have wm1: "w^k * a = - 1"
-        by simp
-      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
-        using norm_ge_zero[of w] w0 m(1)
-        by (simp add: inverse_eq_divide zero_less_mult_iff)
-      with real_lbound_gt_zero[OF zero_less_one] obtain t where
-        t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
-      let ?ct = "complex_of_real t"
-      let ?w = "?ct * w"
-      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w"
-        using kas(1) by (simp add: algebra_simps power_mult_distrib)
-      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
-        unfolding wm1 by simp
-      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) =
-        cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
-        by metis
-      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
-      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)"
-        unfolding norm_of_real by simp
-      have ath: "\<And>x t::real. 0 \<le> x \<Longrightarrow> x < t \<Longrightarrow> t \<le> 1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1"
-        by arith
-      have "t * cmod w \<le> 1 * cmod w"
-        apply (rule mult_mono)
-        using t(1,2)
-        apply auto
-        done
-      then have tw: "cmod ?w \<le> cmod w"
-        using t(1) by (simp add: norm_mult)
-      from t inv0 have "t * (cmod w ^ (k + 1) * m) < 1"
-        by (simp add: field_simps)
-      with zero_less_power[OF t(1), of k] have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
-        by simp
-      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k + 1) * cmod (poly s ?w)))"
-        using w0 t(1)
-        by (simp add: algebra_simps power_mult_distrib norm_power norm_mult)
-      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
-        using t(1,2) m(2)[rule_format, OF tw] w0
-        by auto
-      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k"
-        by simp
-      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
-        by auto
-      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
-      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
-      from th11 th12 have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"
-        by arith
-      then have "cmod (poly ?r ?w) < 1"
-        unfolding kas(4)[rule_format, of ?w] r01 by simp
-      then show ?thesis
-        by blast
-    qed
-    with cq0 q(2) show ?thesis
-      unfolding mrmq_eq not_less[symmetric] by auto
-  qed
-qed
-
-text \<open>Alternative version with a syntactic notion of constant polynomial.\<close>
-
-lemma fundamental_theorem_of_algebra_alt:
-  assumes nc: "\<not> (\<exists>a l. a \<noteq> 0 \<and> l = 0 \<and> p = pCons a l)"
-  shows "\<exists>z. poly p z = (0::complex)"
-  using nc
-proof (induct p)
-  case 0
-  then show ?case by simp
-next
-  case (pCons c cs)
-  show ?case
-  proof (cases "c = 0")
-    case True
-    then show ?thesis by auto
-  next
-    case False
-    have "\<not> constant (poly (pCons c cs))"
-    proof
-      assume nc: "constant (poly (pCons c cs))"
-      from nc[unfolded constant_def, rule_format, of 0]
-      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
-      then have "cs = 0"
-      proof (induct cs)
-        case 0
-        then show ?case by simp
-      next
-        case (pCons d ds)
-        show ?case
-        proof (cases "d = 0")
-          case True
-          then show ?thesis
-            using pCons.prems pCons.hyps by simp
-        next
-          case False
-          from poly_bound_exists[of 1 ds] obtain m where
-            m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
-          have dm: "cmod d / m > 0"
-            using False m(1) by (simp add: field_simps)
-          from real_lbound_gt_zero[OF dm zero_less_one]
-          obtain x where x: "x > 0" "x < cmod d / m" "x < 1"
-            by blast
-          let ?x = "complex_of_real x"
-          from x have cx: "?x \<noteq> 0" "cmod ?x \<le> 1"
-            by simp_all
-          from pCons.prems[rule_format, OF cx(1)]
-          have cth: "cmod (?x*poly ds ?x) = cmod d"
-            by (simp add: eq_diff_eq[symmetric])
-          from m(2)[rule_format, OF cx(2)] x(1)
-          have th0: "cmod (?x*poly ds ?x) \<le> x*m"
-            by (simp add: norm_mult)
-          from x(2) m(1) have "x * m < cmod d"
-            by (simp add: field_simps)
-          with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d"
-            by auto
-          with cth show ?thesis
-            by blast
-        qed
-      qed
-      then show False
-        using pCons.prems False by blast
-    qed
-    then show ?thesis
-      by (rule fundamental_theorem_of_algebra)
-  qed
-qed
-
-
-subsection \<open>Nullstellensatz, degrees and divisibility of polynomials\<close>
-
-lemma nullstellensatz_lemma:
-  fixes p :: "complex poly"
-  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
-    and "degree p = n"
-    and "n \<noteq> 0"
-  shows "p dvd (q ^ n)"
-  using assms
-proof (induct n arbitrary: p q rule: nat_less_induct)
-  fix n :: nat
-  fix p q :: "complex poly"
-  assume IH: "\<forall>m<n. \<forall>p q.
-                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
-                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
-    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
-    and dpn: "degree p = n"
-    and n0: "n \<noteq> 0"
-  from dpn n0 have pne: "p \<noteq> 0" by auto
-  show "p dvd (q ^ n)"
-  proof (cases "\<exists>a. poly p a = 0")
-    case True
-    then obtain a where a: "poly p a = 0" ..
-    have ?thesis if oa: "order a p \<noteq> 0"
-    proof -
-      let ?op = "order a p"
-      from pne have ap: "([:- a, 1:] ^ ?op) dvd p" "\<not> [:- a, 1:] ^ (Suc ?op) dvd p"
-        using order by blast+
-      note oop = order_degree[OF pne, unfolded dpn]
-      show ?thesis
-      proof (cases "q = 0")
-        case True
-        with n0 show ?thesis by (simp add: power_0_left)
-      next
-        case False
-        from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
-        obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
-        from ap(1) obtain s where s: "p = [:- a, 1:] ^ ?op * s"
-          by (rule dvdE)
-        have sne: "s \<noteq> 0"
-          using s pne by auto
-        show ?thesis
-        proof (cases "degree s = 0")
-          case True
-          then obtain k where kpn: "s = [:k:]"
-            by (cases s) (auto split: if_splits)
-          from sne kpn have k: "k \<noteq> 0" by simp
-          let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
-          have "q ^ n = p * ?w"
-            apply (subst r)
-            apply (subst s)
-            apply (subst kpn)
-            using k oop [of a]
-            apply (subst power_mult_distrib)
-            apply simp
-            apply (subst power_add [symmetric])
-            apply simp
-            done
-          then show ?thesis
-            unfolding dvd_def by blast
-        next
-          case False
-          with sne dpn s oa have dsn: "degree s < n"
-            apply auto
-            apply (erule ssubst)
-            apply (simp add: degree_mult_eq degree_linear_power)
-            done
-          have "poly r x = 0" if h: "poly s x = 0" for x
-          proof -
-            have xa: "x \<noteq> a"
-            proof
-              assume "x = a"
-              from h[unfolded this poly_eq_0_iff_dvd] obtain u where u: "s = [:- a, 1:] * u"
-                by (rule dvdE)
-              have "p = [:- a, 1:] ^ (Suc ?op) * u"
-                apply (subst s)
-                apply (subst u)
-                apply (simp only: power_Suc ac_simps)
-                done
-              with ap(2)[unfolded dvd_def] show False
-                by blast
-            qed
-            from h have "poly p x = 0"
-              by (subst s) simp
-            with pq0 have "poly q x = 0"
-              by blast
-            with r xa show ?thesis
-              by auto
-          qed
-          with IH[rule_format, OF dsn, of s r] False have "s dvd (r ^ (degree s))"
-            by blast
-          then obtain u where u: "r ^ (degree s) = s * u" ..
-          then have u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
-            by (simp only: poly_mult[symmetric] poly_power[symmetric])
-          let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
-          from oop[of a] dsn have "q ^ n = p * ?w"
-            apply -
-            apply (subst s)
-            apply (subst r)
-            apply (simp only: power_mult_distrib)
-            apply (subst mult.assoc [where b=s])
-            apply (subst mult.assoc [where a=u])
-            apply (subst mult.assoc [where b=u, symmetric])
-            apply (subst u [symmetric])
-            apply (simp add: ac_simps power_add [symmetric])
-            done
-          then show ?thesis
-            unfolding dvd_def by blast
-        qed
-      qed
-    qed
-    then show ?thesis
-      using a order_root pne by blast
-  next
-    case False
-    with fundamental_theorem_of_algebra_alt[of p]
-    obtain c where ccs: "c \<noteq> 0" "p = pCons c 0"
-      by blast
-    then have pp: "poly p x = c" for x
-      by simp
-    let ?w = "[:1/c:] * (q ^ n)"
-    from ccs have "(q ^ n) = (p * ?w)"
-      by simp
-    then show ?thesis
-      unfolding dvd_def by blast
-  qed
-qed
-
-lemma nullstellensatz_univariate:
-  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow>
-    p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
-proof -
-  consider "p = 0" | "p \<noteq> 0" "degree p = 0" | n where "p \<noteq> 0" "degree p = Suc n"
-    by (cases "degree p") auto
-  then show ?thesis
-  proof cases
-    case p: 1
-    then have eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
-      by (auto simp add: poly_all_0_iff_0)
-    {
-      assume "p dvd (q ^ (degree p))"
-      then obtain r where r: "q ^ (degree p) = p * r" ..
-      from r p have False by simp
-    }
-    with eq p show ?thesis by blast
-  next
-    case dp: 2
-    then obtain k where k: "p = [:k:]" "k \<noteq> 0"
-      by (cases p) (simp split: if_splits)
-    then have th1: "\<forall>x. poly p x \<noteq> 0"
-      by simp
-    from k dp(2) have "q ^ (degree p) = p * [:1/k:]"
-      by (simp add: one_poly_def)
-    then have th2: "p dvd (q ^ (degree p))" ..
-    from dp(1) th1 th2 show ?thesis
-      by blast
-  next
-    case dp: 3
-    have False if dvd: "p dvd (q ^ (Suc n))" and h: "poly p x = 0" "poly q x \<noteq> 0" for x
-    proof -
-      from dvd obtain u where u: "q ^ (Suc n) = p * u" ..
-      from h have "poly (q ^ (Suc n)) x \<noteq> 0"
-        by simp
-      with u h(1) show ?thesis
-        by (simp only: poly_mult) simp
-    qed
-    with dp nullstellensatz_lemma[of p q "degree p"] show ?thesis
-      by auto
-  qed
-qed
-
-text \<open>Useful lemma\<close>
-lemma constant_degree:
-  fixes p :: "'a::{idom,ring_char_0} poly"
-  shows "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
-proof
-  show ?rhs if ?lhs
-  proof -
-    from that[unfolded constant_def, rule_format, of _ "0"]
-    have th: "poly p = poly [:poly p 0:]"
-      by auto
-    then have "p = [:poly p 0:]"
-      by (simp add: poly_eq_poly_eq_iff)
-    then have "degree p = degree [:poly p 0:]"
-      by simp
-    then show ?thesis
-      by simp
-  qed
-  show ?lhs if ?rhs
-  proof -
-    from that obtain k where "p = [:k:]"
-      by (cases p) (simp split: if_splits)
-    then show ?thesis
-      unfolding constant_def by auto
-  qed
-qed
-
-text \<open>Arithmetic operations on multivariate polynomials.\<close>
-
-lemma mpoly_base_conv:
-  fixes x :: "'a::comm_ring_1"
-  shows "0 = poly 0 x" "c = poly [:c:] x" "x = poly [:0,1:] x"
-  by simp_all
-
-lemma mpoly_norm_conv:
-  fixes x :: "'a::comm_ring_1"
-  shows "poly [:0:] x = poly 0 x" "poly [:poly 0 y:] x = poly 0 x"
-  by simp_all
-
-lemma mpoly_sub_conv:
-  fixes x :: "'a::comm_ring_1"
-  shows "poly p x - poly q x = poly p x + -1 * poly q x"
-  by simp
-
-lemma poly_pad_rule: "poly p x = 0 \<Longrightarrow> poly (pCons 0 p) x = 0"
-  by simp
-
-lemma poly_cancel_eq_conv:
-  fixes x :: "'a::field"
-  shows "x = 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> y = 0 \<longleftrightarrow> a * y - b * x = 0"
-  by auto
-
-lemma poly_divides_pad_rule:
-  fixes p:: "('a::comm_ring_1) poly"
-  assumes pq: "p dvd q"
-  shows "p dvd (pCons 0 q)"
-proof -
-  have "pCons 0 q = q * [:0,1:]" by simp
-  then have "q dvd (pCons 0 q)" ..
-  with pq show ?thesis by (rule dvd_trans)
-qed
-
-lemma poly_divides_conv0:
-  fixes p:: "'a::field poly"
-  assumes lgpq: "degree q < degree p"
-    and lq: "p \<noteq> 0"
-  shows "p dvd q \<longleftrightarrow> q = 0" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?rhs
-  then have "q = p * 0" by simp
-  then show ?lhs ..
-next
-  assume l: ?lhs
-  show ?rhs
-  proof (cases "q = 0")
-    case True
-    then show ?thesis by simp
-  next
-    assume q0: "q \<noteq> 0"
-    from l q0 have "degree p \<le> degree q"
-      by (rule dvd_imp_degree_le)
-    with lgpq show ?thesis by simp
-  qed
-qed
-
-lemma poly_divides_conv1:
-  fixes p :: "'a::field poly"
-  assumes a0: "a \<noteq> 0"
-    and pp': "p dvd p'"
-    and qrp': "smult a q - p' = r"
-  shows "p dvd q \<longleftrightarrow> p dvd r" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  from pp' obtain t where t: "p' = p * t" ..
-  show ?rhs if ?lhs
-  proof -
-    from that obtain u where u: "q = p * u" ..
-    have "r = p * (smult a u - t)"
-      using u qrp' [symmetric] t by (simp add: algebra_simps)
-    then show ?thesis ..
-  qed
-  show ?lhs if ?rhs
-  proof -
-    from that obtain u where u: "r = p * u" ..
-    from u [symmetric] t qrp' [symmetric] a0
-    have "q = p * smult (1/a) (u + t)"
-      by (simp add: algebra_simps)
-    then show ?thesis ..
-  qed
-qed
-
-lemma basic_cqe_conv1:
-  "(\<exists>x. poly p x = 0 \<and> poly 0 x \<noteq> 0) \<longleftrightarrow> False"
-  "(\<exists>x. poly 0 x \<noteq> 0) \<longleftrightarrow> False"
-  "(\<exists>x. poly [:c:] x \<noteq> 0) \<longleftrightarrow> c \<noteq> 0"
-  "(\<exists>x. poly 0 x = 0) \<longleftrightarrow> True"
-  "(\<exists>x. poly [:c:] x = 0) \<longleftrightarrow> c = 0"
-  by simp_all
-
-lemma basic_cqe_conv2:
-  assumes l: "p \<noteq> 0"
-  shows "\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)"
-proof -
-  have False if "h \<noteq> 0" "t = 0" and "pCons a (pCons b p) = pCons h t" for h t
-    using l that by simp
-  then have th: "\<not> (\<exists> h t. h \<noteq> 0 \<and> t = 0 \<and> pCons a (pCons b p) = pCons h t)"
-    by blast
-  from fundamental_theorem_of_algebra_alt[OF th] show ?thesis
-    by auto
-qed
-
-lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<longleftrightarrow> p \<noteq> 0"
-  by (metis poly_all_0_iff_0)
-
-lemma basic_cqe_conv3:
-  fixes p q :: "complex poly"
-  assumes l: "p \<noteq> 0"
-  shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<longleftrightarrow> \<not> (pCons a p) dvd (q ^ psize p)"
-proof -
-  from l have dp: "degree (pCons a p) = psize p"
-    by (simp add: psize_def)
-  from nullstellensatz_univariate[of "pCons a p" q] l
-  show ?thesis
-    by (metis dp pCons_eq_0_iff)
-qed
-
-lemma basic_cqe_conv4:
-  fixes p q :: "complex poly"
-  assumes h: "\<And>x. poly (q ^ n) x = poly r x"
-  shows "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
-proof -
-  from h have "poly (q ^ n) = poly r"
-    by auto
-  then have "(q ^ n) = r"
-    by (simp add: poly_eq_poly_eq_iff)
-  then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
-    by simp
-qed
-
-lemma poly_const_conv:
-  fixes x :: "'a::comm_ring_1"
-  shows "poly [:c:] x = y \<longleftrightarrow> c = y"
-  by simp
-
-end
--- a/src/HOL/Library/Library.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Library/Library.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -25,13 +25,10 @@
   Extended_Real
   Finite_Map
   Float
-  Formal_Power_Series
-  Fraction_Field
   FSet
   FuncSet
   Function_Division
   Function_Growth
-  Fundamental_Theorem_Algebra
   Fun_Lexorder
   Groups_Big_Fun
   Indicator_Function
@@ -59,8 +56,6 @@
   Perm
   Permutation
   Permutations
-  Polynomial
-  Polynomial_FPS
   Preorder
   Product_Plus
   Quadratic_Discriminant
--- a/src/HOL/Library/Normalized_Fraction.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,338 +0,0 @@
-(*  Title:      HOL/Library/Normalized_Fraction.thy
-    Author:     Manuel Eberl
-*)
-
-theory Normalized_Fraction
-imports 
-  Main 
-  "~~/src/HOL/Number_Theory/Euclidean_Algorithm" 
-  Fraction_Field
-begin
-
-definition quot_to_fract :: "'a :: {idom} \<times> 'a \<Rightarrow> 'a fract" where
-  "quot_to_fract = (\<lambda>(a,b). Fraction_Field.Fract a b)"
-
-definition normalize_quot :: "'a :: {ring_gcd,idom_divide} \<times> 'a \<Rightarrow> 'a \<times> 'a" where
-  "normalize_quot = 
-     (\<lambda>(a,b). if b = 0 then (0,1) else let d = gcd a b * unit_factor b in (a div d, b div d))" 
-
-definition normalized_fracts :: "('a :: {ring_gcd,idom_divide} \<times> 'a) set" where
-  "normalized_fracts = {(a,b). coprime a b \<and> unit_factor b = 1}"
-  
-lemma not_normalized_fracts_0_denom [simp]: "(a, 0) \<notin> normalized_fracts"
-  by (auto simp: normalized_fracts_def)
-
-lemma unit_factor_snd_normalize_quot [simp]:
-  "unit_factor (snd (normalize_quot x)) = 1"
-  by (simp add: normalize_quot_def case_prod_unfold Let_def dvd_unit_factor_div
-                mult_unit_dvd_iff unit_factor_mult unit_factor_gcd)
-  
-lemma snd_normalize_quot_nonzero [simp]: "snd (normalize_quot x) \<noteq> 0"
-  using unit_factor_snd_normalize_quot[of x] 
-  by (auto simp del: unit_factor_snd_normalize_quot)
-  
-lemma normalize_quot_aux:
-  fixes a b
-  assumes "b \<noteq> 0"
-  defines "d \<equiv> gcd a b * unit_factor b"
-  shows   "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d"
-          "d dvd a" "d dvd b" "d \<noteq> 0"
-proof -
-  from assms show "d dvd a" "d dvd b"
-    by (simp_all add: d_def mult_unit_dvd_iff)
-  thus "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d \<noteq> 0"
-    by (auto simp: normalize_quot_def Let_def d_def \<open>b \<noteq> 0\<close>)
-qed
-
-lemma normalize_quotE:
-  assumes "b \<noteq> 0"
-  obtains d where "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d"
-                  "d dvd a" "d dvd b" "d \<noteq> 0"
-  using that[OF normalize_quot_aux[OF assms]] .
-  
-lemma normalize_quotE':
-  assumes "snd x \<noteq> 0"
-  obtains d where "fst x = fst (normalize_quot x) * d" "snd x = snd (normalize_quot x) * d"
-                  "d dvd fst x" "d dvd snd x" "d \<noteq> 0"
-proof -
-  from normalize_quotE[OF assms, of "fst x"] guess d .
-  from this show ?thesis unfolding prod.collapse by (intro that[of d])
-qed
-  
-lemma coprime_normalize_quot:
-  "coprime (fst (normalize_quot x)) (snd (normalize_quot x))"
-  by (simp add: normalize_quot_def case_prod_unfold Let_def
-        div_mult_unit2 gcd_div_unit1 gcd_div_unit2 div_gcd_coprime)
-
-lemma normalize_quot_in_normalized_fracts [simp]: "normalize_quot x \<in> normalized_fracts"
-  by (simp add: normalized_fracts_def coprime_normalize_quot case_prod_unfold)
-
-lemma normalize_quot_eq_iff:
-  assumes "b \<noteq> 0" "d \<noteq> 0"
-  shows   "normalize_quot (a,b) = normalize_quot (c,d) \<longleftrightarrow> a * d = b * c"
-proof -
-  define x y where "x = normalize_quot (a,b)" and "y = normalize_quot (c,d)" 
-  from normalize_quotE[OF assms(1), of a] normalize_quotE[OF assms(2), of c]
-    obtain d1 d2 
-      where "a = fst x * d1" "b = snd x * d1" "c = fst y * d2" "d = snd y * d2" "d1 \<noteq> 0" "d2 \<noteq> 0"
-    unfolding x_def y_def by metis
-  hence "a * d = b * c \<longleftrightarrow> fst x * snd y = snd x * fst y" by simp
-  also have "\<dots> \<longleftrightarrow> fst x = fst y \<and> snd x = snd y"
-    by (intro coprime_crossproduct') (simp_all add: x_def y_def coprime_normalize_quot)
-  also have "\<dots> \<longleftrightarrow> x = y" using prod_eqI by blast
-  finally show "x = y \<longleftrightarrow> a * d = b * c" ..
-qed
-
-lemma normalize_quot_eq_iff':
-  assumes "snd x \<noteq> 0" "snd y \<noteq> 0"
-  shows   "normalize_quot x = normalize_quot y \<longleftrightarrow> fst x * snd y = snd x * fst y"
-  using assms by (cases x, cases y, hypsubst) (subst normalize_quot_eq_iff, simp_all)
-
-lemma normalize_quot_id: "x \<in> normalized_fracts \<Longrightarrow> normalize_quot x = x"
-  by (auto simp: normalized_fracts_def normalize_quot_def case_prod_unfold)
-
-lemma normalize_quot_idem [simp]: "normalize_quot (normalize_quot x) = normalize_quot x"
-  by (rule normalize_quot_id) simp_all
-
-lemma fractrel_iff_normalize_quot_eq:
-  "fractrel x y \<longleftrightarrow> normalize_quot x = normalize_quot y \<and> snd x \<noteq> 0 \<and> snd y \<noteq> 0"
-  by (cases x, cases y) (auto simp: fractrel_def normalize_quot_eq_iff)
-  
-lemma fractrel_normalize_quot_left:
-  assumes "snd x \<noteq> 0"
-  shows   "fractrel (normalize_quot x) y \<longleftrightarrow> fractrel x y"
-  using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto
-
-lemma fractrel_normalize_quot_right:
-  assumes "snd x \<noteq> 0"
-  shows   "fractrel y (normalize_quot x) \<longleftrightarrow> fractrel y x"
-  using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto
-
-  
-lift_definition quot_of_fract :: "'a :: {ring_gcd,idom_divide} fract \<Rightarrow> 'a \<times> 'a" 
-    is normalize_quot
-  by (subst (asm) fractrel_iff_normalize_quot_eq) simp_all
-  
-lemma quot_to_fract_quot_of_fract [simp]: "quot_to_fract (quot_of_fract x) = x"
-  unfolding quot_to_fract_def
-proof transfer
-  fix x :: "'a \<times> 'a" assume rel: "fractrel x x"
-  define x' where "x' = normalize_quot x"
-  obtain a b where [simp]: "x = (a, b)" by (cases x)
-  from rel have "b \<noteq> 0" by simp
-  from normalize_quotE[OF this, of a] guess d .
-  hence "a = fst x' * d" "b = snd x' * d" "d \<noteq> 0" "snd x' \<noteq> 0" by (simp_all add: x'_def)
-  thus "fractrel (case x' of (a, b) \<Rightarrow> if b = 0 then (0, 1) else (a, b)) x"
-    by (auto simp add: case_prod_unfold)
-qed
-
-lemma quot_of_fract_quot_to_fract: "quot_of_fract (quot_to_fract x) = normalize_quot x"
-proof (cases "snd x = 0")
-  case True
-  thus ?thesis unfolding quot_to_fract_def
-    by transfer (simp add: case_prod_unfold normalize_quot_def)
-next
-  case False
-  thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold)
-qed
-
-lemma quot_of_fract_quot_to_fract': 
-  "x \<in> normalized_fracts \<Longrightarrow> quot_of_fract (quot_to_fract x) = x"
-  unfolding quot_to_fract_def by transfer (auto simp: normalize_quot_id)
-
-lemma quot_of_fract_in_normalized_fracts [simp]: "quot_of_fract x \<in> normalized_fracts"
-  by transfer simp
-
-lemma normalize_quotI:
-  assumes "a * d = b * c" "b \<noteq> 0" "(c, d) \<in> normalized_fracts"
-  shows   "normalize_quot (a, b) = (c, d)"
-proof -
-  from assms have "normalize_quot (a, b) = normalize_quot (c, d)"
-    by (subst normalize_quot_eq_iff) auto
-  also have "\<dots> = (c, d)" by (intro normalize_quot_id) fact
-  finally show ?thesis .
-qed
-
-lemma td_normalized_fract:
-  "type_definition quot_of_fract quot_to_fract normalized_fracts"
-  by standard (simp_all add: quot_of_fract_quot_to_fract')
-
-lemma quot_of_fract_add_aux:
-  assumes "snd x \<noteq> 0" "snd y \<noteq> 0" 
-  shows   "(fst x * snd y + fst y * snd x) * (snd (normalize_quot x) * snd (normalize_quot y)) =
-             snd x * snd y * (fst (normalize_quot x) * snd (normalize_quot y) +
-             snd (normalize_quot x) * fst (normalize_quot y))"
-proof -
-  from normalize_quotE'[OF assms(1)] guess d . note d = this
-  from normalize_quotE'[OF assms(2)] guess e . note e = this
-  show ?thesis by (simp_all add: d e algebra_simps)
-qed
-
-
-locale fract_as_normalized_quot
-begin
-setup_lifting td_normalized_fract
-end
-
-
-lemma quot_of_fract_add:
-  "quot_of_fract (x + y) = 
-     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y
-      in  normalize_quot (a * d + b * c, b * d))"
-  by transfer (insert quot_of_fract_add_aux, 
-               simp_all add: Let_def case_prod_unfold normalize_quot_eq_iff)
-
-lemma quot_of_fract_uminus:
-  "quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))"
-  by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff)
-
-lemma quot_of_fract_diff:
-  "quot_of_fract (x - y) = 
-     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y
-      in  normalize_quot (a * d - b * c, b * d))" (is "_ = ?rhs")
-proof -
-  have "x - y = x + -y" by simp
-  also have "quot_of_fract \<dots> = ?rhs"
-    by (simp only: quot_of_fract_add quot_of_fract_uminus Let_def case_prod_unfold) simp_all
-  finally show ?thesis .
-qed
-
-lemma normalize_quot_mult_coprime:
-  assumes "coprime a b" "coprime c d" "unit_factor b = 1" "unit_factor d = 1"
-  defines "e \<equiv> fst (normalize_quot (a, d))" and "f \<equiv> snd (normalize_quot (a, d))"
-     and  "g \<equiv> fst (normalize_quot (c, b))" and "h \<equiv> snd (normalize_quot (c, b))"
-  shows   "normalize_quot (a * c, b * d) = (e * g, f * h)"
-proof (rule normalize_quotI)
-  from assms have "b \<noteq> 0" "d \<noteq> 0" by auto
-  from normalize_quotE[OF \<open>b \<noteq> 0\<close>, of c] guess k . note k = this [folded assms]
-  from normalize_quotE[OF \<open>d \<noteq> 0\<close>, of a] guess l . note l = this [folded assms]
-  from k l show "a * c * (f * h) = b * d * (e * g)" by (simp_all)
-  from assms have [simp]: "unit_factor f = 1" "unit_factor h = 1"
-    by simp_all
-  from assms have "coprime e f" "coprime g h" by (simp_all add: coprime_normalize_quot)
-  with k l assms(1,2) show "(e * g, f * h) \<in> normalized_fracts"
-    by (simp add: normalized_fracts_def unit_factor_mult coprime_mul_eq coprime_mul_eq')
-qed (insert assms(3,4), auto)
-
-lemma normalize_quot_mult:
-  assumes "snd x \<noteq> 0" "snd y \<noteq> 0"
-  shows   "normalize_quot (fst x * fst y, snd x * snd y) = normalize_quot 
-             (fst (normalize_quot x) * fst (normalize_quot y),
-              snd (normalize_quot x) * snd (normalize_quot y))"
-proof -
-  from normalize_quotE'[OF assms(1)] guess d . note d = this
-  from normalize_quotE'[OF assms(2)] guess e . note e = this
-  show ?thesis by (simp_all add: d e algebra_simps normalize_quot_eq_iff)
-qed
-
-lemma quot_of_fract_mult:
-  "quot_of_fract (x * y) = 
-     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y;
-          (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b)
-      in  (e*g, f*h))"
-  by transfer (simp_all add: Let_def case_prod_unfold normalize_quot_mult_coprime [symmetric]
-                 coprime_normalize_quot normalize_quot_mult [symmetric])
-  
-lemma normalize_quot_0 [simp]: 
-    "normalize_quot (0, x) = (0, 1)" "normalize_quot (x, 0) = (0, 1)"
-  by (simp_all add: normalize_quot_def)
-  
-lemma normalize_quot_eq_0_iff [simp]: "fst (normalize_quot x) = 0 \<longleftrightarrow> fst x = 0 \<or> snd x = 0"
-  by (auto simp: normalize_quot_def case_prod_unfold Let_def div_mult_unit2 dvd_div_eq_0_iff)
-  find_theorems "_ div _ = 0"
-  
-lemma fst_quot_of_fract_0_imp: "fst (quot_of_fract x) = 0 \<Longrightarrow> snd (quot_of_fract x) = 1"
-  by transfer auto
-
-lemma normalize_quot_swap:
-  assumes "a \<noteq> 0" "b \<noteq> 0"
-  defines "a' \<equiv> fst (normalize_quot (a, b))" and "b' \<equiv> snd (normalize_quot (a, b))"
-  shows   "normalize_quot (b, a) = (b' div unit_factor a', a' div unit_factor a')"
-proof (rule normalize_quotI)
-  from normalize_quotE[OF assms(2), of a] guess d . note d = this [folded assms(3,4)]
-  show "b * (a' div unit_factor a') = a * (b' div unit_factor a')"
-    using assms(1,2) d 
-    by (simp add: div_unit_factor [symmetric] unit_div_mult_swap mult_ac del: div_unit_factor)
-  have "coprime a' b'" by (simp add: a'_def b'_def coprime_normalize_quot)
-  thus "(b' div unit_factor a', a' div unit_factor a') \<in> normalized_fracts"
-    using assms(1,2) d by (auto simp: normalized_fracts_def gcd_div_unit1 gcd_div_unit2 gcd.commute)
-qed fact+
-  
-lemma quot_of_fract_inverse:
-  "quot_of_fract (inverse x) = 
-     (let (a,b) = quot_of_fract x; d = unit_factor a 
-      in  if d = 0 then (0, 1) else (b div d, a div d))"
-proof (transfer, goal_cases)
-  case (1 x)
-  from normalize_quot_swap[of "fst x" "snd x"] show ?case
-    by (auto simp: Let_def case_prod_unfold)
-qed
-
-lemma normalize_quot_div_unit_left:
-  fixes x y u
-  assumes "is_unit u"
-  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
-  shows "normalize_quot (x div u, y) = (x' div u, y')"
-proof (cases "y = 0")
-  case False
-  from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)]
-  from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot)
-  with False d \<open>is_unit u\<close> show ?thesis
-    by (intro normalize_quotI)
-       (auto simp: normalized_fracts_def unit_div_mult_swap unit_div_commute unit_div_cancel
-          gcd_div_unit1)
-qed (simp_all add: assms)
-
-lemma normalize_quot_div_unit_right:
-  fixes x y u
-  assumes "is_unit u"
-  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
-  shows "normalize_quot (x, y div u) = (x' * u, y')"
-proof (cases "y = 0")
-  case False
-  from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)]
-  from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot)
-  with False d \<open>is_unit u\<close> show ?thesis
-    by (intro normalize_quotI)
-       (auto simp: normalized_fracts_def unit_div_mult_swap unit_div_commute unit_div_cancel
-          gcd_mult_unit1 unit_div_eq_0_iff mult.assoc [symmetric])
-qed (simp_all add: assms)
-
-lemma normalize_quot_normalize_left:
-  fixes x y u
-  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
-  shows "normalize_quot (normalize x, y) = (x' div unit_factor x, y')"
-  using normalize_quot_div_unit_left[of "unit_factor x" x y]
-  by (cases "x = 0") (simp_all add: assms)
-  
-lemma normalize_quot_normalize_right:
-  fixes x y u
-  defines "x' \<equiv> fst (normalize_quot (x, y))" and "y' \<equiv> snd (normalize_quot (x, y))"
-  shows "normalize_quot (x, normalize y) = (x' * unit_factor y, y')"
-  using normalize_quot_div_unit_right[of "unit_factor y" x y]
-  by (cases "y = 0") (simp_all add: assms)
-  
-lemma quot_of_fract_0 [simp]: "quot_of_fract 0 = (0, 1)"
-  by transfer auto
-
-lemma quot_of_fract_1 [simp]: "quot_of_fract 1 = (1, 1)"
-  by transfer (rule normalize_quotI, simp_all add: normalized_fracts_def)
-
-lemma quot_of_fract_divide:
-  "quot_of_fract (x / y) = (if y = 0 then (0, 1) else
-     (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y;
-          (e,f) = normalize_quot (a,c); (g,h) = normalize_quot (d,b)
-      in  (e * g, f * h)))" (is "_ = ?rhs")
-proof (cases "y = 0")
-  case False
-  hence A: "fst (quot_of_fract y) \<noteq> 0" by transfer auto
-  have "x / y = x * inverse y" by (simp add: divide_inverse)
-  also from False A have "quot_of_fract \<dots> = ?rhs"
-    by (simp only: quot_of_fract_mult quot_of_fract_inverse)
-       (simp_all add: Let_def case_prod_unfold fst_quot_of_fract_0_imp
-          normalize_quot_div_unit_left normalize_quot_div_unit_right 
-          normalize_quot_normalize_right normalize_quot_normalize_left)
-  finally show ?thesis .
-qed simp_all
-
-end
--- a/src/HOL/Library/Polynomial.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4453 +0,0 @@
-(*  Title:      HOL/Library/Polynomial.thy
-    Author:     Brian Huffman
-    Author:     Clemens Ballarin
-    Author:     Amine Chaieb
-    Author:     Florian Haftmann
-*)
-
-section \<open>Polynomials as type over a ring structure\<close>
-
-theory Polynomial
-  imports
-    "~~/src/HOL/Deriv"
-    More_List
-    Infinite_Set
-begin
-
-subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
-
-definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixr "##" 65)
-  where "x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)"
-
-lemma cCons_0_Nil_eq [simp]: "0 ## [] = []"
-  by (simp add: cCons_def)
-
-lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys"
-  by (simp add: cCons_def)
-
-lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys"
-  by (simp add: cCons_def)
-
-lemma cCons_not_0_eq [simp]: "x \<noteq> 0 \<Longrightarrow> x ## xs = x # xs"
-  by (simp add: cCons_def)
-
-lemma strip_while_not_0_Cons_eq [simp]:
-  "strip_while (\<lambda>x. x = 0) (x # xs) = x ## strip_while (\<lambda>x. x = 0) xs"
-proof (cases "x = 0")
-  case False
-  then show ?thesis by simp
-next
-  case True
-  show ?thesis
-  proof (induct xs rule: rev_induct)
-    case Nil
-    with True show ?case by simp
-  next
-    case (snoc y ys)
-    then show ?case
-      by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons)
-  qed
-qed
-
-lemma tl_cCons [simp]: "tl (x ## xs) = xs"
-  by (simp add: cCons_def)
-
-
-subsection \<open>Definition of type \<open>poly\<close>\<close>
-
-typedef (overloaded) 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
-  morphisms coeff Abs_poly
-  by (auto intro!: ALL_MOST)
-
-setup_lifting type_definition_poly
-
-lemma poly_eq_iff: "p = q \<longleftrightarrow> (\<forall>n. coeff p n = coeff q n)"
-  by (simp add: coeff_inject [symmetric] fun_eq_iff)
-
-lemma poly_eqI: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q"
-  by (simp add: poly_eq_iff)
-
-lemma MOST_coeff_eq_0: "\<forall>\<^sub>\<infinity> n. coeff p n = 0"
-  using coeff [of p] by simp
-
-
-subsection \<open>Degree of a polynomial\<close>
-
-definition degree :: "'a::zero poly \<Rightarrow> nat"
-  where "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)"
-
-lemma coeff_eq_0:
-  assumes "degree p < n"
-  shows "coeff p n = 0"
-proof -
-  have "\<exists>n. \<forall>i>n. coeff p i = 0"
-    using MOST_coeff_eq_0 by (simp add: MOST_nat)
-  then have "\<forall>i>degree p. coeff p i = 0"
-    unfolding degree_def by (rule LeastI_ex)
-  with assms show ?thesis by simp
-qed
-
-lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p"
-  by (erule contrapos_np, rule coeff_eq_0, simp)
-
-lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n"
-  unfolding degree_def by (erule Least_le)
-
-lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0"
-  unfolding degree_def by (drule not_less_Least, simp)
-
-
-subsection \<open>The zero polynomial\<close>
-
-instantiation poly :: (zero) zero
-begin
-
-lift_definition zero_poly :: "'a poly"
-  is "\<lambda>_. 0"
-  by (rule MOST_I) simp
-
-instance ..
-
-end
-
-lemma coeff_0 [simp]: "coeff 0 n = 0"
-  by transfer rule
-
-lemma degree_0 [simp]: "degree 0 = 0"
-  by (rule order_antisym [OF degree_le le0]) simp
-
-lemma leading_coeff_neq_0:
-  assumes "p \<noteq> 0"
-  shows "coeff p (degree p) \<noteq> 0"
-proof (cases "degree p")
-  case 0
-  from \<open>p \<noteq> 0\<close> obtain n where "coeff p n \<noteq> 0"
-    by (auto simp add: poly_eq_iff)
-  then have "n \<le> degree p"
-    by (rule le_degree)
-  with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close> show "coeff p (degree p) \<noteq> 0"
-    by simp
-next
-  case (Suc n)
-  from \<open>degree p = Suc n\<close> have "n < degree p"
-    by simp
-  then have "\<exists>i>n. coeff p i \<noteq> 0"
-    by (rule less_degree_imp)
-  then obtain i where "n < i" and "coeff p i \<noteq> 0"
-    by blast
-  from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i"
-    by simp
-  also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p"
-    by (rule le_degree)
-  finally have "degree p = i" .
-  with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp
-qed
-
-lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0"
-  by (cases "p = 0") (simp_all 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)
-  from \<open>degree p \<le> n\<close> have "\<forall>i>n. coeff p i = 0"
-    by (simp add: coeff_eq_0)
-  with \<open>coeff p n = 0\<close> have "\<forall>i\<ge>n. coeff p i = 0"
-    by (simp add: le_less)
-  with \<open>n = Suc m\<close> have "\<forall>i>m. coeff p i = 0"
-    by (simp add: less_eq_Suc_le)
-  then have "degree p \<le> m"
-    by (rule degree_le)
-  with \<open>n = Suc m\<close> have "degree p < n"
-    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>
-
-lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>a p. case_nat a (coeff p)"
-  by (rule MOST_SucD) (simp add: MOST_coeff_eq_0)
-
-lemmas coeff_pCons = pCons.rep_eq
-
-lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a"
-  by transfer simp
-
-lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n"
-  by (simp add: coeff_pCons)
-
-lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)"
-  by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split)
-
-lemma degree_pCons_eq: "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)"
-  apply (rule order_antisym [OF degree_pCons_le])
-  apply (rule le_degree, simp)
-  done
-
-lemma degree_pCons_0: "degree (pCons a 0) = 0"
-  apply (rule order_antisym [OF _ le0])
-  apply (rule degree_le, simp add: coeff_pCons split: nat.split)
-  done
-
-lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))"
-  apply (cases "p = 0", simp_all)
-  apply (rule order_antisym [OF _ le0])
-  apply (rule degree_le, simp add: coeff_pCons split: nat.split)
-  apply (rule order_antisym [OF degree_pCons_le])
-  apply (rule le_degree, simp)
-  done
-
-lemma pCons_0_0 [simp]: "pCons 0 0 = 0"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q"
-proof safe
-  assume "pCons a p = pCons b q"
-  then have "coeff (pCons a p) 0 = coeff (pCons b q) 0"
-    by simp
-  then show "a = b"
-    by simp
-next
-  assume "pCons a p = pCons b q"
-  then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n
-    by simp
-  then show "p = q"
-    by (simp add: poly_eq_iff)
-qed
-
-lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0"
-  using pCons_eq_iff [of a p 0 0] by simp
-
-lemma pCons_cases [cases type: poly]:
-  obtains (pCons) a q where "p = pCons a q"
-proof
-  show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))"
-    by transfer
-      (simp_all add: MOST_inj[where f=Suc and P="\<lambda>n. p n = 0" for p] fun_eq_iff Abs_poly_inverse
-        split: nat.split)
-qed
-
-lemma pCons_induct [case_names 0 pCons, induct type: poly]:
-  assumes zero: "P 0"
-  assumes pCons: "\<And>a p. a \<noteq> 0 \<or> p \<noteq> 0 \<Longrightarrow> P p \<Longrightarrow> P (pCons a p)"
-  shows "P p"
-proof (induct p rule: measure_induct_rule [where f=degree])
-  case (less p)
-  obtain a q where "p = pCons a q" by (rule pCons_cases)
-  have "P q"
-  proof (cases "q = 0")
-    case True
-    then show "P q" by (simp add: zero)
-  next
-    case False
-    then have "degree (pCons a q) = Suc (degree q)"
-      by (rule degree_pCons_eq)
-    with \<open>p = pCons a q\<close> have "degree q < degree p"
-      by simp
-    then show "P q"
-      by (rule less.hyps)
-  qed
-  have "P (pCons a q)"
-  proof (cases "a \<noteq> 0 \<or> q \<noteq> 0")
-    case True
-    with \<open>P q\<close> show ?thesis by (auto intro: pCons)
-  next
-    case False
-    with zero show ?thesis by simp
-  qed
-  with \<open>p = pCons a q\<close> show ?case
-    by simp
-qed
-
-lemma degree_eq_zeroE:
-  fixes p :: "'a::zero poly"
-  assumes "degree p = 0"
-  obtains a where "p = pCons a 0"
-proof -
-  obtain a q where p: "p = pCons a q"
-    by (cases p)
-  with assms have "q = 0"
-    by (cases "q = 0") simp_all
-  with p have "p = pCons a 0"
-    by simp
-  then show thesis ..
-qed
-
-
-subsection \<open>Quickcheck generator for polynomials\<close>
-
-quickcheck_generator poly constructors: "0 :: _ poly", pCons
-
-
-subsection \<open>List-style syntax for polynomials\<close>
-
-syntax "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
-translations
-  "[:x, xs:]" \<rightleftharpoons> "CONST pCons x [:xs:]"
-  "[:x:]" \<rightleftharpoons> "CONST pCons x 0"
-  "[:x:]" \<leftharpoondown> "CONST pCons x (_constrain 0 t)"
-
-
-subsection \<open>Representation of polynomials by lists of coefficients\<close>
-
-primrec Poly :: "'a::zero list \<Rightarrow> 'a poly"
-  where
-    [code_post]: "Poly [] = 0"
-  | [code_post]: "Poly (a # as) = pCons a (Poly as)"
-
-lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0"
-  by (induct n) simp_all
-
-lemma Poly_eq_0: "Poly as = 0 \<longleftrightarrow> (\<exists>n. as = replicate n 0)"
-  by (induct as) (auto simp add: Cons_replicate_eq)
-
-lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as"
-  by (induct as) simp_all
-
-lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as"
-  using Poly_append_replicate_zero [of as 1] by simp
-
-lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)"
-  by (simp add: cCons_def)
-
-lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \<Longrightarrow> Poly (rev (tl as)) = Poly (rev as)"
-  by (cases as) simp_all
-
-lemma degree_Poly: "degree (Poly xs) \<le> length xs"
-  by (induct xs) simp_all
-
-lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs"
-  by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits)
-
-definition coeffs :: "'a poly \<Rightarrow> 'a::zero list"
-  where "coeffs p = (if p = 0 then [] else map (\<lambda>i. coeff p i) [0 ..< Suc (degree p)])"
-
-lemma coeffs_eq_Nil [simp]: "coeffs p = [] \<longleftrightarrow> p = 0"
-  by (simp add: coeffs_def)
-
-lemma not_0_coeffs_not_Nil: "p \<noteq> 0 \<Longrightarrow> coeffs p \<noteq> []"
-  by simp
-
-lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []"
-  by simp
-
-lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p"
-proof -
-  have *: "\<forall>m\<in>set ms. m > 0 \<Longrightarrow> map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)"
-    for ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a"
-    by (induct ms) (auto split: nat.split)
-  show ?thesis
-    by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc)
-qed
-
-lemma length_coeffs: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = degree p + 1"
-  by (simp add: coeffs_def)
-
-lemma coeffs_nth: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeffs p ! n = coeff p n"
-  by (auto simp: coeffs_def simp del: upt_Suc)
-
-lemma coeff_in_coeffs: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeff p n \<in> set (coeffs p)"
-  using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs)
-
-lemma not_0_cCons_eq [simp]: "p \<noteq> 0 \<Longrightarrow> a ## coeffs p = a # coeffs p"
-  by (simp add: cCons_def)
-
-lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p"
-  by (induct p) auto
-
-lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as"
-proof (induct as)
-  case Nil
-  then show ?case by simp
-next
-  case (Cons a as)
-  from replicate_length_same [of as 0] have "(\<forall>n. as \<noteq> replicate n 0) \<longleftrightarrow> (\<exists>a\<in>set as. a \<noteq> 0)"
-    by (auto dest: sym [of _ as])
-  with Cons show ?case by auto
-qed
-
-lemma no_trailing_coeffs [simp]:
-  "no_trailing (HOL.eq 0) (coeffs p)"
-  by (induct p)  auto
-
-lemma strip_while_coeffs [simp]:
-  "strip_while (HOL.eq 0) (coeffs p) = coeffs p"
-  by simp
-
-lemma coeffs_eq_iff: "p = q \<longleftrightarrow> coeffs p = coeffs q"
-  (is "?P \<longleftrightarrow> ?Q")
-proof
-  assume ?P
-  then show ?Q by simp
-next
-  assume ?Q
-  then have "Poly (coeffs p) = Poly (coeffs q)" by simp
-  then show ?P by simp
-qed
-
-lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p"
-  by (simp add: fun_eq_iff coeff_Poly_eq [symmetric])
-
-lemma [code]: "coeff p = nth_default 0 (coeffs p)"
-  by (simp add: nth_default_coeffs_eq)
-
-lemma coeffs_eqI:
-  assumes coeff: "\<And>n. coeff p n = nth_default 0 xs n"
-  assumes zero: "no_trailing (HOL.eq 0) xs"
-  shows "coeffs p = xs"
-proof -
-  from coeff have "p = Poly xs"
-    by (simp add: poly_eq_iff)
-  with zero show ?thesis by simp
-qed
-
-lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1"
-  by (simp add: coeffs_def)
-
-lemma length_coeffs_degree: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = Suc (degree p)"
-  by (induct p) (auto simp: cCons_def)
-
-lemma [code abstract]: "coeffs 0 = []"
-  by (fact coeffs_0_eq_Nil)
-
-lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p"
-  by (fact coeffs_pCons_eq_cCons)
-
-instantiation poly :: ("{zero, equal}") equal
-begin
-
-definition [code]: "HOL.equal (p::'a poly) q \<longleftrightarrow> HOL.equal (coeffs p) (coeffs q)"
-
-instance
-  by standard (simp add: equal equal_poly_def coeffs_eq_iff)
-
-end
-
-lemma [code nbe]: "HOL.equal (p :: _ poly) p \<longleftrightarrow> True"
-  by (fact equal_refl)
-
-definition is_zero :: "'a::zero poly \<Rightarrow> bool"
-  where [code]: "is_zero p \<longleftrightarrow> List.null (coeffs p)"
-
-lemma is_zero_null [code_abbrev]: "is_zero p \<longleftrightarrow> p = 0"
-  by (simp add: is_zero_def null_def)
-
-
-subsubsection \<open>Reconstructing the polynomial from the list\<close>
-  \<comment> \<open>contributed by Sebastiaan J.C. Joosten and René Thiemann\<close>
-
-definition poly_of_list :: "'a::comm_monoid_add list \<Rightarrow> 'a poly"
-  where [simp]: "poly_of_list = Poly"
-
-lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as"
-  by simp
-
-
-subsection \<open>Fold combinator for polynomials\<close>
-
-definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
-  where "fold_coeffs f p = foldr f (coeffs p)"
-
-lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id"
-  by (simp add: fold_coeffs_def)
-
-lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p"
-  by (simp add: fold_coeffs_def cCons_def fun_eq_iff)
-
-lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id"
-  by (simp add: fold_coeffs_def)
-
-lemma fold_coeffs_pCons_coeff_not_0_eq [simp]:
-  "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p"
-  by (simp add: fold_coeffs_def)
-
-lemma fold_coeffs_pCons_not_0_0_eq [simp]:
-  "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"
-  where "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" \<comment> \<open>The Horner Schema\<close>
-
-lemma poly_0 [simp]: "poly 0 x = 0"
-  by (simp add: poly_def)
-
-lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x"
-  by (cases "p = 0 \<and> a = 0") (auto simp add: poly_def)
-
-lemma poly_altdef: "poly p x = (\<Sum>i\<le>degree p. coeff p i * x ^ i)"
-  for x :: "'a::{comm_semiring_0,semiring_1}"
-proof (induction p rule: pCons_induct)
-  case 0
-  then show ?case
-    by simp
-next
-  case (pCons a p)
-  show ?case
-  proof (cases "p = 0")
-    case True
-    then show ?thesis by simp
-  next
-    case False
-    let ?p' = "pCons a p"
-    note poly_pCons[of a p x]
-    also note pCons.IH
-    also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) =
-        coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)"
-      by (simp add: field_simps sum_distrib_left coeff_pCons)
-    also note sum_atMost_Suc_shift[symmetric]
-    also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric]
-    finally show ?thesis .
-  qed
-qed
-
-lemma poly_0_coeff_0: "poly p 0 = coeff p 0"
-  by (cases p) (auto simp: poly_altdef)
-
-
-subsection \<open>Monomials\<close>
-
-lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly"
-  is "\<lambda>a m n. if m = n then a else 0"
-  by (simp add: MOST_iff_cofinite)
-
-lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)"
-  by transfer rule
-
-lemma monom_0: "monom a 0 = pCons a 0"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma monom_eq_0 [simp]: "monom 0 n = 0"
-  by (rule poly_eqI) simp
-
-lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0"
-  by (simp add: poly_eq_iff)
-
-lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b"
-  by (simp add: poly_eq_iff)
-
-lemma degree_monom_le: "degree (monom a n) \<le> n"
-  by (rule degree_le, simp)
-
-lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n"
-  apply (rule order_antisym [OF degree_monom_le])
-  apply (rule le_degree)
-  apply simp
-  done
-
-lemma coeffs_monom [code abstract]:
-  "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])"
-  by (induct n) (simp_all add: monom_0 monom_Suc)
-
-lemma fold_coeffs_monom [simp]: "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (monom a n) = f 0 ^^ n \<circ> f a"
-  by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff)
-
-lemma poly_monom: "poly (monom a n) x = a * x ^ n"
-  for a x :: "'a::comm_semiring_1"
-  by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_def)
-
-lemma monom_eq_iff': "monom c n = monom d m \<longleftrightarrow>  c = d \<and> (c = 0 \<or> n = m)"
-  by (auto simp: poly_eq_iff)
-
-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
-begin
-
-lift_definition plus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>p q n. coeff p n + coeff q n"
-proof -
-  fix q p :: "'a poly"
-  show "\<forall>\<^sub>\<infinity>n. coeff p n + coeff q n = 0"
-    using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp
-qed
-
-lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n"
-  by (simp add: plus_poly.rep_eq)
-
-instance
-proof
-  fix p q r :: "'a poly"
-  show "(p + q) + r = p + (q + r)"
-    by (simp add: poly_eq_iff add.assoc)
-  show "p + q = q + p"
-    by (simp add: poly_eq_iff add.commute)
-  show "0 + p = p"
-    by (simp add: poly_eq_iff)
-qed
-
-end
-
-instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add
-begin
-
-lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>p q n. coeff p n - coeff q n"
-proof -
-  fix q p :: "'a poly"
-  show "\<forall>\<^sub>\<infinity>n. coeff p n - coeff q n = 0"
-    using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp
-qed
-
-lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n"
-  by (simp add: minus_poly.rep_eq)
-
-instance
-proof
-  fix p q r :: "'a poly"
-  show "p + q - p = q"
-    by (simp add: poly_eq_iff)
-  show "p - q - r = p - (q + r)"
-    by (simp add: poly_eq_iff diff_diff_eq)
-qed
-
-end
-
-instantiation poly :: (ab_group_add) ab_group_add
-begin
-
-lift_definition uminus_poly :: "'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>p n. - coeff p n"
-proof -
-  fix p :: "'a poly"
-  show "\<forall>\<^sub>\<infinity>n. - coeff p n = 0"
-    using MOST_coeff_eq_0 by simp
-qed
-
-lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n"
-  by (simp add: uminus_poly.rep_eq)
-
-instance
-proof
-  fix p q :: "'a poly"
-  show "- p + p = 0"
-    by (simp add: poly_eq_iff)
-  show "p - q = p + - q"
-    by (simp add: poly_eq_iff)
-qed
-
-end
-
-lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)"
-  by (rule degree_le) (auto simp add: coeff_eq_0)
-
-lemma degree_add_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p + q) \<le> n"
-  by (auto intro: order_trans degree_add_le_max)
-
-lemma degree_add_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p + q) < n"
-  by (auto intro: le_less_trans degree_add_le_max)
-
-lemma degree_add_eq_right: "degree p < degree q \<Longrightarrow> degree (p + q) = degree q"
-  apply (cases "q = 0")
-   apply simp
-  apply (rule order_antisym)
-   apply (simp add: degree_add_le)
-  apply (rule le_degree)
-  apply (simp add: coeff_eq_0)
-  done
-
-lemma degree_add_eq_left: "degree q < degree p \<Longrightarrow> degree (p + q) = degree p"
-  using degree_add_eq_right [of q p] by (simp add: add.commute)
-
-lemma degree_minus [simp]: "degree (- p) = degree p"
-  by (simp add: degree_def)
-
-lemma lead_coeff_add_le: "degree p < degree q \<Longrightarrow> lead_coeff (p + q) = lead_coeff q"
-  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: "degree (p - q) \<le> max (degree p) (degree q)"
-  for p q :: "'a::ab_group_add poly"
-  using degree_add_le [where p=p and q="-q"] by simp
-
-lemma degree_diff_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p - q) \<le> n"
-  for p q :: "'a::ab_group_add poly"
-  using degree_add_le [of p n "- q"] by simp
-
-lemma degree_diff_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p - q) < n"
-  for p q :: "'a::ab_group_add poly"
-  using degree_add_less [of p n "- q"] by simp
-
-lemma add_monom: "monom a n + monom b n = monom (a + b) n"
-  by (rule poly_eqI) simp
-
-lemma diff_monom: "monom a n - monom b n = monom (a - b) n"
-  by (rule poly_eqI) simp
-
-lemma minus_monom: "- monom a n = monom (- a) n"
-  by (rule poly_eqI) simp
-
-lemma coeff_sum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)"
-  by (induct A rule: infinite_finite_induct) simp_all
-
-lemma monom_sum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)"
-  by (rule poly_eqI) (simp add: coeff_sum)
-
-fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-  where
-    "plus_coeffs xs [] = xs"
-  | "plus_coeffs [] ys = ys"
-  | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys"
-
-lemma coeffs_plus_eq_plus_coeffs [code abstract]:
-  "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)"
-proof -
-  have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n"
-    for xs ys :: "'a list" and n
-  proof (induct xs ys arbitrary: n rule: plus_coeffs.induct)
-    case (3 x xs y ys n)
-    then show ?case
-      by (cases n) (auto simp add: cCons_def)
-  qed simp_all
-  have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)"
-    if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys"
-    for xs ys :: "'a list"
-    using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def)
-  show ?thesis
-    by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **)
-qed
-
-lemma coeffs_uminus [code abstract]:
-  "coeffs (- p) = map uminus (coeffs p)"
-proof -
-  have eq_0: "HOL.eq 0 \<circ> uminus = HOL.eq (0::'a)"
-    by (simp add: fun_eq_iff)
-  show ?thesis
-    by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0)
-qed
-
-lemma [code]: "p - q = p + - q"
-  for p q :: "'a::ab_group_add poly"
-  by (fact diff_conv_add_uminus)
-
-lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x"
-  apply (induct p arbitrary: q)
-   apply simp
-  apply (case_tac q, simp, simp add: algebra_simps)
-  done
-
-lemma poly_minus [simp]: "poly (- p) x = - poly p x"
-  for x :: "'a::comm_ring"
-  by (induct p) simp_all
-
-lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x"
-  for x :: "'a::comm_ring"
-  using poly_add [of p "- q" x] by simp
-
-lemma poly_sum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
-  by (induct A rule: infinite_finite_induct) simp_all
-
-lemma degree_sum_le: "finite S \<Longrightarrow> (\<And>p. p \<in> S \<Longrightarrow> degree (f p) \<le> n) \<Longrightarrow> degree (sum f S) \<le> n"
-proof (induct S rule: finite_induct)
-  case empty
-  then show ?case by simp
-next
-  case (insert p S)
-  then have "degree (sum f S) \<le> n" "degree (f p) \<le> n"
-    by auto
-  then show ?case
-    unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le)
-qed
-
-lemma poly_as_sum_of_monoms':
-  assumes "degree p \<le> n"
-  shows "(\<Sum>i\<le>n. monom (coeff p i) i) = p"
-proof -
-  have eq: "\<And>i. {..n} \<inter> {i} = (if i \<le> n then {i} else {})"
-    by auto
-  from assms show ?thesis
-    by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq
-        if_distrib[where f="\<lambda>x. x * a" for a])
-qed
-
-lemma poly_as_sum_of_monoms: "(\<Sum>i\<le>degree p. monom (coeff p i) i) = p"
-  by (intro poly_as_sum_of_monoms' order_refl)
-
-lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)"
-  by (induct xs) (simp_all add: monom_0 monom_Suc)
-
-
-subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close>
-
-lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>a p n. a * coeff p n"
-proof -
-  fix a :: 'a and p :: "'a poly"
-  show "\<forall>\<^sub>\<infinity> i. a * coeff p i = 0"
-    using MOST_coeff_eq_0[of p] by eventually_elim simp
-qed
-
-lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n"
-  by (simp add: smult.rep_eq)
-
-lemma degree_smult_le: "degree (smult a p) \<le> degree p"
-  by (rule degree_le) (simp add: coeff_eq_0)
-
-lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p"
-  by (rule poly_eqI) (simp add: mult.assoc)
-
-lemma smult_0_right [simp]: "smult a 0 = 0"
-  by (rule poly_eqI) simp
-
-lemma smult_0_left [simp]: "smult 0 p = 0"
-  by (rule poly_eqI) simp
-
-lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p"
-  by (rule poly_eqI) simp
-
-lemma smult_add_right: "smult a (p + q) = smult a p + smult a q"
-  by (rule poly_eqI) (simp add: algebra_simps)
-
-lemma smult_add_left: "smult (a + b) p = smult a p + smult b p"
-  by (rule poly_eqI) (simp add: algebra_simps)
-
-lemma smult_minus_right [simp]: "smult a (- p) = - smult a p"
-  for a :: "'a::comm_ring"
-  by (rule poly_eqI) simp
-
-lemma smult_minus_left [simp]: "smult (- a) p = - smult a p"
-  for a :: "'a::comm_ring"
-  by (rule poly_eqI) simp
-
-lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q"
-  for a :: "'a::comm_ring"
-  by (rule poly_eqI) (simp add: algebra_simps)
-
-lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p"
-  for a b :: "'a::comm_ring"
-  by (rule poly_eqI) (simp add: algebra_simps)
-
-lemmas smult_distribs =
-  smult_add_left smult_add_right
-  smult_diff_left smult_diff_right
-
-lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)"
-  by (rule poly_eqI) (simp add: coeff_pCons split: nat.split)
-
-lemma smult_monom: "smult a (monom b n) = monom (a * b) n"
-  by (induct n) (simp_all add: monom_0 monom_Suc)
-
-lemma smult_Poly: "smult c (Poly xs) = Poly (map (op * c) xs)"
-  by (auto simp: poly_eq_iff nth_default_def)
-
-lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)"
-  for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
-  by (cases "a = 0") (simp_all add: degree_def)
-
-lemma smult_eq_0_iff [simp]: "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0"
-  for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
-  by (simp add: poly_eq_iff)
-
-lemma coeffs_smult [code abstract]:
-  "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))"
-  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-proof -
-  have eq_0: "HOL.eq 0 \<circ> times a = HOL.eq (0::'a)" if "a \<noteq> 0"
-    using that by (simp add: fun_eq_iff)
-  show ?thesis
-    by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0)
-qed  
-
-lemma smult_eq_iff:
-  fixes b :: "'a :: field"
-  assumes "b \<noteq> 0"
-  shows "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
-    (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?lhs
-  also from assms have "smult (inverse b) \<dots> = q"
-    by simp
-  finally show ?rhs
-    by (simp add: field_simps)
-next
-  assume ?rhs
-  with assms show ?lhs by auto
-qed
-
-instantiation poly :: (comm_semiring_0) comm_semiring_0
-begin
-
-definition "p * q = fold_coeffs (\<lambda>a p. smult a q + pCons 0 p) p 0"
-
-lemma mult_poly_0_left: "(0::'a poly) * q = 0"
-  by (simp add: times_poly_def)
-
-lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)"
-  by (cases "p = 0 \<and> a = 0") (auto simp add: times_poly_def)
-
-lemma mult_poly_0_right: "p * (0::'a poly) = 0"
-  by (induct p) (simp_all add: mult_poly_0_left)
-
-lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)"
-  by (induct p) (simp_all add: mult_poly_0_left algebra_simps)
-
-lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right
-
-lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)"
-  by (induct p) (simp_all add: mult_poly_0 smult_add_right)
-
-lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)"
-  by (induct q) (simp_all add: mult_poly_0 smult_add_right)
-
-lemma mult_poly_add_left: "(p + q) * r = p * r + q * r"
-  for p q r :: "'a poly"
-  by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps)
-
-instance
-proof
-  fix p q r :: "'a poly"
-  show 0: "0 * p = 0"
-    by (rule mult_poly_0_left)
-  show "p * 0 = 0"
-    by (rule mult_poly_0_right)
-  show "(p + q) * r = p * r + q * r"
-    by (rule mult_poly_add_left)
-  show "(p * q) * r = p * (q * r)"
-    by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left)
-  show "p * q = q * p"
-    by (induct p) (simp_all add: mult_poly_0)
-qed
-
-end
-
-lemma coeff_mult_degree_sum:
-  "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)"
-  by (induct p) (simp_all add: coeff_eq_0)
-
-instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors
-proof
-  fix p q :: "'a poly"
-  assume "p \<noteq> 0" and "q \<noteq> 0"
-  have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)"
-    by (rule coeff_mult_degree_sum)
-  also from \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "coeff p (degree p) * coeff q (degree q) \<noteq> 0"
-    by simp
-  finally have "\<exists>n. coeff (p * q) n \<noteq> 0" ..
-  then show "p * q \<noteq> 0"
-    by (simp add: poly_eq_iff)
-qed
-
-instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel ..
-
-lemma coeff_mult: "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))"
-proof (induct p arbitrary: n)
-  case 0
-  show ?case by simp
-next
-  case (pCons a p n)
-  then show ?case
-    by (cases n) (simp_all add: sum_atMost_Suc_shift del: sum_atMost_Suc)
-qed
-
-lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q"
-  apply (rule degree_le)
-  apply (induct p)
-   apply simp
-  apply (simp add: coeff_eq_0 coeff_pCons split: nat.split)
-  done
-
-lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)"
-  by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc)
-
-instantiation poly :: (comm_semiring_1) comm_semiring_1
-begin
-
-definition one_poly_def: "1 = pCons 1 0"
-
-instance
-proof
-  show "1 * p = p" for p :: "'a poly"
-    by (simp add: one_poly_def)
-  show "0 \<noteq> (1::'a poly)"
-    by (simp add: one_poly_def)
-qed
-
-end
-
-instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors ..
-instance poly :: (comm_ring) comm_ring ..
-instance poly :: (comm_ring_1) comm_ring_1 ..
-instance poly :: (comm_ring_1) comm_semiring_1_cancel ..
-
-lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)"
-  by (simp add: one_poly_def coeff_pCons split: nat.split)
-
-lemma monom_eq_1 [simp]: "monom 1 0 = 1"
-  by (simp add: monom_0 one_poly_def)
-
-lemma monom_eq_1_iff: "monom c n = 1 \<longleftrightarrow> c = 1 \<and> n = 0"
-  using monom_eq_const_iff[of c n 1] by (auto simp: one_poly_def)
-
-lemma monom_altdef: "monom c n = smult c ([:0, 1:]^n)"
-  by (induct n) (simp_all add: monom_0 monom_Suc one_poly_def)
-
-lemma degree_1 [simp]: "degree 1 = 0"
-  unfolding one_poly_def by (rule degree_pCons_0)
-
-lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]"
-  by (simp add: one_poly_def)
-
-lemma degree_power_le: "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 (induct n) (simp_all add: coeff_mult)
-
-lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x"
-  by (induct p) (simp_all add: algebra_simps)
-
-lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x"
-  by (induct p) (simp_all add: algebra_simps)
-
-lemma poly_1 [simp]: "poly 1 x = 1"
-  by (simp add: one_poly_def)
-
-lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n"
-  for p :: "'a::comm_semiring_1 poly"
-  by (induct n) simp_all
-
-lemma poly_prod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)"
-  by (induct A rule: infinite_finite_induct) simp_all
-
-lemma degree_prod_sum_le: "finite S \<Longrightarrow> degree (prod f S) \<le> sum (degree o f) S"
-proof (induct S rule: finite_induct)
-  case empty
-  then show ?case by simp
-next
-  case (insert a S)
-  show ?case
-    unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)]
-    by (rule le_trans[OF degree_mult_le]) (use insert in auto)
-qed
-
-lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
-  by (induct 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 "p = monom 1 n * r"
-    by (rule dvdE)
-  then show "\<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
-  then have 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 (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero)
-  then show "monom 1 n dvd p" by simp
-qed
-
-
-subsection \<open>Mapping polynomials\<close>
-
-definition map_poly :: "('a :: zero \<Rightarrow> 'b :: zero) \<Rightarrow> 'a poly \<Rightarrow> 'b poly"
-  where "map_poly f p = Poly (map f (coeffs p))"
-
-lemma map_poly_0 [simp]: "map_poly f 0 = 0"
-  by (simp add: map_poly_def)
-
-lemma map_poly_1: "map_poly f 1 = [:f 1:]"
-  by (simp add: map_poly_def)
-
-lemma map_poly_1' [simp]: "f 1 = 1 \<Longrightarrow> map_poly f 1 = 1"
-  by (simp add: map_poly_def one_poly_def)
-
-lemma coeff_map_poly:
-  assumes "f 0 = 0"
-  shows "coeff (map_poly f p) n = f (coeff p n)"
-  by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0
-      simp del: upt_Suc)
-
-lemma coeffs_map_poly [code abstract]:
-  "coeffs (map_poly f p) = strip_while (op = 0) (map f (coeffs p))"
-  by (simp add: map_poly_def)
-
-lemma coeffs_map_poly':
-  assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
-  shows "coeffs (map_poly f p) = map f (coeffs p)"
-  using assms by (simp add: coeffs_map_poly no_trailing_map strip_while_idem_iff)
-    (metis comp_def no_leading_def no_trailing_coeffs)
-
-lemma set_coeffs_map_poly:
-  "(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)"
-  by (simp add: coeffs_map_poly')
-
-lemma degree_map_poly:
-  assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
-  shows "degree (map_poly f p) = degree p"
-  by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms)
-
-lemma map_poly_eq_0_iff:
-  assumes "f 0 = 0" "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
-  shows "map_poly f p = 0 \<longleftrightarrow> p = 0"
-proof -
-  have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n
-  proof -
-    have "coeff (map_poly f p) n = f (coeff p n)"
-      by (simp add: coeff_map_poly assms)
-    also have "\<dots> = 0 \<longleftrightarrow> coeff p n = 0"
-    proof (cases "n < length (coeffs p)")
-      case True
-      then have "coeff p n \<in> set (coeffs p)"
-        by (auto simp: coeffs_def simp del: upt_Suc)
-      with assms show "f (coeff p n) = 0 \<longleftrightarrow> coeff p n = 0"
-        by auto
-    next
-      case False
-      then show ?thesis
-        by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def)
-    qed
-    finally show ?thesis .
-  qed
-  then show ?thesis by (auto simp: poly_eq_iff)
-qed
-
-lemma map_poly_smult:
-  assumes "f 0 = 0""\<And>c x. f (c * x) = f c * f x"
-  shows "map_poly f (smult c p) = smult (f c) (map_poly f p)"
-  by (intro poly_eqI) (simp_all add: assms coeff_map_poly)
-
-lemma map_poly_pCons:
-  assumes "f 0 = 0"
-  shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)"
-  by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits)
-
-lemma map_poly_map_poly:
-  assumes "f 0 = 0" "g 0 = 0"
-  shows "map_poly f (map_poly g p) = map_poly (f \<circ> g) p"
-  by (intro poly_eqI) (simp add: coeff_map_poly assms)
-
-lemma map_poly_id [simp]: "map_poly id p = p"
-  by (simp add: map_poly_def)
-
-lemma map_poly_id' [simp]: "map_poly (\<lambda>x. x) p = p"
-  by (simp add: map_poly_def)
-
-lemma map_poly_cong:
-  assumes "(\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = g x)"
-  shows "map_poly f p = map_poly g p"
-proof -
-  from assms have "map f (coeffs p) = map g (coeffs p)"
-    by (intro map_cong) simp_all
-  then show ?thesis
-    by (simp only: coeffs_eq_iff coeffs_map_poly)
-qed
-
-lemma map_poly_monom: "f 0 = 0 \<Longrightarrow> map_poly f (monom c n) = monom (f c) n"
-  by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-lemma map_poly_idI:
-  assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
-  shows "map_poly f p = p"
-  using map_poly_cong[OF assms, of _ id] by simp
-
-lemma map_poly_idI':
-  assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
-  shows "p = map_poly f p"
-  using map_poly_cong[OF assms, of _ id] by simp
-
-lemma smult_conv_map_poly: "smult c p = map_poly (\<lambda>x. c * x) p"
-  by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-
-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 (induct n)
-  case 0
-  then show ?case by simp
-next
-  case (Suc n)
-  then have "of_nat (Suc n) = 1 + (of_nat n :: 'a poly)"
-    by simp
-  also have "(of_nat n :: 'a poly) = [: of_nat n :]"
-    by (subst Suc) (rule refl)
-  also have "1 = [:1:]"
-    by (simp add: one_poly_def)
-  finally show ?case
-    by (subst (asm) add_pCons) simp
-qed
-
-lemma degree_of_nat [simp]: "degree (of_nat n) = 0"
-  by (simp add: of_nat_poly)
-
-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>
-
-lemma dvd_smult:
-  assumes "p dvd q"
-  shows "p dvd smult a q"
-proof -
-  from assms obtain k where "q = p * k" ..
-  then have "smult a q = p * smult a k" by simp
-  then show "p dvd smult a q" ..
-qed
-
-lemma dvd_smult_cancel: "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q"
-  for a :: "'a::field"
-  by (drule dvd_smult [where a="inverse a"]) simp
-
-lemma dvd_smult_iff: "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q"
-  for a :: "'a::field"
-  by (safe elim!: dvd_smult dvd_smult_cancel)
-
-lemma smult_dvd_cancel:
-  assumes "smult a p dvd q"
-  shows "p dvd q"
-proof -
-  from assms obtain k where "q = smult a p * k" ..
-  then have "q = p * smult a k" by simp
-  then show "p dvd q" ..
-qed
-
-lemma smult_dvd: "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q"
-  for a :: "'a::field"
-  by (rule smult_dvd_cancel [where a="inverse a"]) simp
-
-lemma smult_dvd_iff: "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)"
-  for a :: "'a::field"
-  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 *: "[:c:] * p dvd 1"
-    then show "p dvd 1"
-      by (rule dvd_mult_right)
-    from * obtain q where q: "1 = [:c:] * p * q"
-      by (rule 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 q [symmetric]
-    finally have "c dvd coeff 1 0" .
-    then show "c dvd 1" by simp
-  next
-    assume "c dvd 1" "p dvd 1"
-    from this(1) obtain d where "1 = c * d"
-      by (rule dvdE)
-    then have "1 = [:c:] * [:d:]"
-      by (simp add: one_poly_def mult_ac)
-    then have "[: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>
-
-instance poly :: (idom) idom ..
-
-lemma degree_mult_eq: "p \<noteq> 0 \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree (p * q) = degree p + degree q"
-  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum)
-
-lemma degree_mult_eq_0:
-  "degree (p * q) = 0 \<longleftrightarrow> p = 0 \<or> q = 0 \<or> (p \<noteq> 0 \<and> q \<noteq> 0 \<and> degree p = 0 \<and> degree q = 0)"
-  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  by (auto simp: degree_mult_eq)
-
-lemma degree_mult_right_le:
-  fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  assumes "q \<noteq> 0"
-  shows "degree p \<le> degree (p * q)"
-  using assms by (cases "p = 0") (simp_all add: degree_mult_eq)
-
-lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)"
-  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  by (cases "p = 0 \<or> q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac)
-
-lemma dvd_imp_degree_le: "p dvd q \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree p \<le> degree q"
-  for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  by (erule dvdE, hypsubst, subst degree_mult_eq) auto
-
-lemma divides_degree:
-  fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  assumes "p dvd q"
-  shows "degree p \<le> degree q \<or> q = 0"
-  by (metis dvd_imp_degree_le assms)
-
-lemma const_poly_dvd_iff:
-  fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
-  shows "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)"
-proof (cases "c = 0 \<or> p = 0")
-  case True
-  then show ?thesis
-    by (auto intro!: poly_eqI)
-next
-  case False
-  show ?thesis
-  proof
-    assume "[:c:] dvd p"
-    then show "\<forall>n. c dvd coeff p n"
-      by (auto elim!: dvdE simp: coeffs_def)
-  next
-    assume *: "\<forall>n. c dvd coeff p n"
-    define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a
-    have mydiv: "x = y * mydiv x y" if "y dvd x" for x y
-      using that unfolding mydiv_def dvd_def by (rule someI_ex)
-    define q where "q = Poly (map (\<lambda>a. mydiv a c) (coeffs p))"
-    from False * have "p = q * [:c:]"
-      by (intro poly_eqI)
-        (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth
-          intro!: coeff_eq_0 mydiv)
-    then show "[:c:] dvd p"
-      by (simp only: dvd_triv_right)
-  qed
-qed
-
-lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \<longleftrightarrow> a dvd b"
-  for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
-  by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits)
-
-lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
-  for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly"
-  by (cases "p = 0 \<or> q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq)
-
-lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p"
-  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-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) = lead_coeff p ^ n"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  by (induct n) (simp_all add: lead_coeff_mult)
-
-
-subsection \<open>Polynomials form an ordered integral domain\<close>
-
-definition pos_poly :: "'a::linordered_semidom poly \<Rightarrow> bool"
-  where "pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)"
-
-lemma pos_poly_pCons: "pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)"
-  by (simp add: pos_poly_def)
-
-lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0"
-  by (simp add: pos_poly_def)
-
-lemma pos_poly_add: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p + q)"
-  apply (induct p arbitrary: q)
-   apply simp
-  apply (case_tac q)
-  apply (force simp add: pos_poly_pCons add_pos_pos)
-  done
-
-lemma pos_poly_mult: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p * q)"
-  unfolding pos_poly_def
-  apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0")
-   apply (simp add: degree_mult_eq coeff_mult_degree_sum)
-  apply auto
-  done
-
-lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)"
-  for p :: "'a::linordered_idom poly"
-  by (induct p) (auto simp: pos_poly_pCons)
-
-lemma last_coeffs_eq_coeff_degree: "p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)"
-  by (simp add: coeffs_def)
-
-lemma pos_poly_coeffs [code]: "pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?rhs
-  then show ?lhs
-    by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree)
-next
-  assume ?lhs
-  then have *: "0 < coeff p (degree p)"
-    by (simp add: pos_poly_def)
-  then have "p \<noteq> 0"
-    by auto
-  with * show ?rhs
-    by (simp add: last_coeffs_eq_coeff_degree)
-qed
-
-instantiation poly :: (linordered_idom) linordered_idom
-begin
-
-definition "x < y \<longleftrightarrow> pos_poly (y - x)"
-
-definition "x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)"
-
-definition "\<bar>x::'a poly\<bar> = (if x < 0 then - x else x)"
-
-definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
-
-instance
-proof
-  fix x y z :: "'a poly"
-  show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
-    unfolding less_eq_poly_def less_poly_def
-    apply safe
-     apply simp
-    apply (drule (1) pos_poly_add)
-    apply simp
-    done
-  show "x \<le> x"
-    by (simp add: less_eq_poly_def)
-  show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
-    unfolding less_eq_poly_def
-    apply safe
-    apply (drule (1) pos_poly_add)
-    apply (simp add: algebra_simps)
-    done
-  show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
-    unfolding less_eq_poly_def
-    apply safe
-    apply (drule (1) pos_poly_add)
-    apply simp
-    done
-  show "x \<le> y \<Longrightarrow> z + x \<le> z + y"
-    unfolding less_eq_poly_def
-    apply safe
-    apply (simp add: algebra_simps)
-    done
-  show "x \<le> y \<or> y \<le> x"
-    unfolding less_eq_poly_def
-    using pos_poly_total [of "x - y"]
-    by auto
-  show "x < y \<Longrightarrow> 0 < z \<Longrightarrow> z * x < z * y"
-    by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult)
-  show "\<bar>x\<bar> = (if x < 0 then - x else x)"
-    by (rule abs_poly_def)
-  show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
-    by (rule sgn_poly_def)
-qed
-
-end
-
-text \<open>TODO: Simplification rules for comparisons\<close>
-
-
-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>
-
-definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a"
-  where "synthetic_divmod p c = fold_coeffs (\<lambda>a (q, r). (pCons r q, a + c * r)) p (0, 0)"
-
-definition synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
-  where "synthetic_div p c = fst (synthetic_divmod p c)"
-
-lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)"
-  by (simp add: synthetic_divmod_def)
-
-lemma synthetic_divmod_pCons [simp]:
-  "synthetic_divmod (pCons a p) c = (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)"
-  by (cases "p = 0 \<and> a = 0") (auto simp add: synthetic_divmod_def)
-
-lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0"
-  by (simp add: synthetic_div_def)
-
-lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0"
-  by (induct p arbitrary: a) simp_all
-
-lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c"
-  by (induct p) (simp_all add: split_def)
-
-lemma synthetic_div_pCons [simp]:
-  "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)"
-  by (simp add: synthetic_div_def split_def snd_synthetic_divmod)
-
-lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0"
-proof (induct p)
-  case 0
-  then show ?case by simp
-next
-  case (pCons a p)
-  then show ?case by (cases p) simp
-qed
-
-lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1"
-  by (induct p) (simp_all add: synthetic_div_eq_0_iff)
-
-lemma synthetic_div_correct:
-  "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)"
-  by (induct p) simp_all
-
-lemma synthetic_div_unique: "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c"
-  apply (induct p arbitrary: q r)
-   apply simp
-   apply (frule synthetic_div_unique_lemma)
-   apply simp
-  apply (case_tac q, force)
-  done
-
-lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p"
-  for c :: "'a::comm_ring_1"
-  using synthetic_div_correct [of p c] by (simp add: algebra_simps)
-
-
-subsubsection \<open>Polynomial roots\<close>
-
-lemma poly_eq_0_iff_dvd: "poly p c = 0 \<longleftrightarrow> [:- c, 1:] dvd p"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-  for c :: "'a::comm_ring_1"
-proof
-  assume ?lhs
-  with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp
-  then show ?rhs ..
-next
-  assume ?rhs
-  then obtain k where "p = [:-c, 1:] * k" by (rule dvdE)
-  then show ?lhs by simp
-qed
-
-lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \<longleftrightarrow> poly p (- c) = 0"
-  for c :: "'a::comm_ring_1"
-  by (simp add: poly_eq_0_iff_dvd)
-
-lemma poly_roots_finite: "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
-  for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly"
-proof (induct n \<equiv> "degree p" arbitrary: p)
-  case 0
-  then obtain a where "a \<noteq> 0" and "p = [:a:]"
-    by (cases p) (simp split: if_splits)
-  then show "finite {x. poly p x = 0}"
-    by simp
-next
-  case (Suc n)
-  show "finite {x. poly p x = 0}"
-  proof (cases "\<exists>x. poly p x = 0")
-    case False
-    then show "finite {x. poly p x = 0}" by simp
-  next
-    case True
-    then obtain a where "poly p a = 0" ..
-    then have "[:-a, 1:] dvd p"
-      by (simp only: poly_eq_0_iff_dvd)
-    then obtain k where k: "p = [:-a, 1:] * k" ..
-    with \<open>p \<noteq> 0\<close> have "k \<noteq> 0"
-      by auto
-    with k have "degree p = Suc (degree k)"
-      by (simp add: degree_mult_eq del: mult_pCons_left)
-    with \<open>Suc n = degree p\<close> have "n = degree k"
-      by simp
-    from this \<open>k \<noteq> 0\<close> have "finite {x. poly k x = 0}"
-      by (rule Suc.hyps)
-    then have "finite (insert a {x. poly k x = 0})"
-      by simp
-    then show "finite {x. poly p x = 0}"
-      by (simp add: k Collect_disj_eq del: mult_pCons_left)
-  qed
-qed
-
-lemma poly_eq_poly_eq_iff: "poly p = poly q \<longleftrightarrow> p = q"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-  for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly"
-proof
-  assume ?rhs
-  then show ?lhs by simp
-next
-  assume ?lhs
-  have "poly p = poly 0 \<longleftrightarrow> p = 0" for p :: "'a poly"
-    apply (cases "p = 0")
-     apply simp_all
-    apply (drule poly_roots_finite)
-    apply (auto simp add: infinite_UNIV_char_0)
-    done
-  from \<open>?lhs\<close> and this [of "p - q"] show ?rhs
-    by auto
-qed
-
-lemma poly_all_0_iff_0: "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0"
-  for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly"
-  by (auto simp add: poly_eq_poly_eq_iff [symmetric])
-
-
-subsubsection \<open>Order of polynomial roots\<close>
-
-definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat"
-  where "order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)"
-
-lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1"
-  for a :: "'a::comm_semiring_1"
-  apply (induct n)
-   apply simp_all
-  apply (subst coeff_eq_0)
-   apply (auto intro: le_less_trans degree_power_le)
-  done
-
-lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n"
-  for a :: "'a::comm_semiring_1"
-  apply (rule order_antisym)
-   apply (rule ord_le_eq_trans [OF degree_power_le])
-   apply simp
-  apply (rule le_degree)
-  apply (simp add: coeff_linear_power)
-  done
-
-lemma order_1: "[:-a, 1:] ^ order a p dvd p"
-  apply (cases "p = 0")
-   apply simp
-  apply (cases "order a p")
-   apply simp
-  apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)")
-   apply (drule not_less_Least)
-   apply simp
-  apply (fold order_def)
-  apply simp
-  done
-
-lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
-  unfolding order_def
-  apply (rule LeastI_ex)
-  apply (rule_tac x="degree p" in exI)
-  apply (rule notI)
-  apply (drule (1) dvd_imp_degree_le)
-  apply (simp only: degree_linear_power)
-  done
-
-lemma order: "p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
-  by (rule conjI [OF order_1 order_2])
-
-lemma order_degree:
-  assumes p: "p \<noteq> 0"
-  shows "order a p \<le> degree p"
-proof -
-  have "order a p = degree ([:-a, 1:] ^ order a p)"
-    by (simp only: degree_linear_power)
-  also from order_1 p have "\<dots> \<le> degree p"
-    by (rule dvd_imp_degree_le)
-  finally show ?thesis .
-qed
-
-lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0"
-  apply (cases "p = 0")
-   apply simp_all
-  apply (rule iffI)
-   apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right)
-  unfolding poly_eq_0_iff_dvd
-  apply (metis dvd_power dvd_trans order_1)
-  done
-
-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"
-    by (simp add: t_def 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 True
-  then show ?thesis
-    by simp
-next
-  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 have "order x [:c:] = 0"
-    by (rule order_0I) (use assms in auto)
-  finally show ?thesis
-    by simp
-qed
-
-(* 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
-  then show ?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 notI)
-    assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
-    then have "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:])"
-      by (rule dvd_imp_degree_le) auto
-    then show False
-      by auto
-  next
-    fix y
-    assume *: "\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
-    show "1 \<le> y"
-    proof (rule ccontr)
-      assume "\<not> 1 \<le> y"
-      then have "y = 0" by auto
-      then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto
-      with * show False by auto
-    qed
-  qed
-  ultimately show ?case
-    using Suc by auto
-qed
-
-lemma order_0_monom [simp]: "c \<noteq> 0 \<Longrightarrow> order 0 (monom c n) = n"
-  using 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. \<open>f / gcd f f'\<close>.\<close>
-
-lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"
-  apply (cases "p = 0")
-  apply 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, 1:] ^ order a p dvd p"
-    and **: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p"
-    by (auto dest: order)
-  from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" ..
-  with ** 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
-  with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q]
-  have "\<not> [:- a, 1:] dvd q" by auto
-  with q show ?thesis by blast
-qed
-
-lemma monom_1_dvd_iff: "p \<noteq> 0 \<Longrightarrow> monom 1 n dvd p \<longleftrightarrow> n \<le> order 0 p"
-  using order_divides[of 0 n p] by (simp add: monom_altdef)
-
-
-subsection \<open>Additional induction rules on polynomials\<close>
-
-text \<open>
-  An induction rule for induction over the roots of a polynomial with a certain property.
-  (e.g. all positive roots)
-\<close>
-lemma poly_root_induct [case_names 0 no_roots root]:
-  fixes p :: "'a :: idom poly"
-  assumes "Q 0"
-    and "\<And>p. (\<And>a. P a \<Longrightarrow> poly p a \<noteq> 0) \<Longrightarrow> Q p"
-    and "\<And>a p. P a \<Longrightarrow> Q p \<Longrightarrow> Q ([:a, -1:] * p)"
-  shows "Q p"
-proof (induction "degree p" arbitrary: p rule: less_induct)
-  case (less p)
-  show ?case
-  proof (cases "p = 0")
-    case True
-    with assms(1) show ?thesis by simp
-  next
-    case False
-    show ?thesis
-    proof (cases "\<exists>a. P a \<and> poly p a = 0")
-      case False
-      then show ?thesis by (intro assms(2)) blast
-    next
-      case True
-      then obtain a where a: "P a" "poly p a = 0"
-        by blast
-      then have "-[:-a, 1:] dvd p"
-        by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd)
-      then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp
-      with False have "q \<noteq> 0" by auto
-      have "degree p = Suc (degree q)"
-        by (subst q, subst degree_mult_eq) (simp_all add: \<open>q \<noteq> 0\<close>)
-      then have "Q q" by (intro less) simp
-      with a(1) have "Q ([:a, -1:] * q)"
-        by (rule assms(3))
-      with q show ?thesis by simp
-    qed
-  qed
-qed
-
-lemma dropWhile_replicate_append:
-  "dropWhile (op = a) (replicate n a @ ys) = dropWhile (op = a) ys"
-  by (induct n) simp_all
-
-lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs"
-  by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append)
-
-text \<open>
-  An induction rule for simultaneous induction over two polynomials,
-  prepending one coefficient in each step.
-\<close>
-lemma poly_induct2 [case_names 0 pCons]:
-  assumes "P 0 0" "\<And>a p b q. P p q \<Longrightarrow> P (pCons a p) (pCons b q)"
-  shows "P p q"
-proof -
-  define n where "n = max (length (coeffs p)) (length (coeffs q))"
-  define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)"
-  define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)"
-  have "length xs = length ys"
-    by (simp add: xs_def ys_def n_def)
-  then have "P (Poly xs) (Poly ys)"
-    by (induct rule: list_induct2) (simp_all add: assms)
-  also have "Poly xs = p"
-    by (simp add: xs_def Poly_append_replicate_0)
-  also have "Poly ys = q"
-    by (simp add: ys_def Poly_append_replicate_0)
-  finally show ?thesis .
-qed
-
-
-subsection \<open>Composition of polynomials\<close>
-
-(* Several lemmas contributed by René Thiemann and Akihisa Yamada *)
-
-definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  where "pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0"
-
-notation pcompose (infixl "\<circ>\<^sub>p" 71)
-
-lemma pcompose_0 [simp]: "pcompose 0 q = 0"
-  by (simp add: pcompose_def)
-
-lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q"
-  by (cases "p = 0 \<and> a = 0") (auto simp add: pcompose_def)
-
-lemma pcompose_1: "pcompose 1 p = 1"
-  for p :: "'a::comm_semiring_1 poly"
-  by (auto simp: one_poly_def pcompose_pCons)
-
-lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)"
-  by (induct p) (simp_all add: pcompose_pCons)
-
-lemma degree_pcompose_le: "degree (pcompose p q) \<le> degree p * degree q"
-  apply (induct p)
-   apply simp
-  apply (simp add: pcompose_pCons)
-  apply clarify
-  apply (rule degree_add_le)
-   apply simp
-  apply (rule order_trans [OF degree_mult_le])
-  apply simp
-  done
-
-lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r"
-  for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly"
-proof (induction p q rule: poly_induct2)
-  case 0
-  then show ?case by simp
-next
-  case (pCons a p b q)
-  have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r"
-    by (simp_all add: pcompose_pCons pCons.IH algebra_simps)
-  also have "[:a + b:] = [:a:] + [:b:]" by simp
-  also have "\<dots> + r * pcompose p r + r * pcompose q r =
-    pcompose (pCons a p) r + pcompose (pCons b q) r"
-    by (simp only: pcompose_pCons add_ac)
-  finally show ?case .
-qed
-
-lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r"
-  for p r :: "'a::comm_ring poly"
-  by (induct p) (simp_all add: pcompose_pCons)
-
-lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r"
-  for p q r :: "'a::comm_ring poly"
-  using pcompose_add[of p "-q"] by (simp add: pcompose_uminus)
-
-lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)"
-  for p r :: "'a::comm_semiring_0 poly"
-  by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right)
-
-lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r"
-  for p q r :: "'a::comm_semiring_0 poly"
-  by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps)
-
-lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r"
-  for p q r :: "'a::comm_semiring_0 poly"
-  by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult)
-
-lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p"
-  for p :: "'a::comm_semiring_1 poly"
-  by (induct p) (simp_all add: pcompose_pCons)
-
-lemma pcompose_sum: "pcompose (sum f A) p = sum (\<lambda>i. pcompose (f i) p) A"
-  by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add)
-
-lemma pcompose_prod: "pcompose (prod f A) p = prod (\<lambda>i. pcompose (f i) p) A"
-  by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult)
-
-lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]"
-  by (subst pcompose_pCons) simp
-
-lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]"
-  by (induct p) (auto simp add: pcompose_pCons)
-
-lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q"
-  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-proof (induct p)
-  case 0
-  then show ?case by auto
-next
-  case (pCons a p)
-  consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0"
-    by blast
-  then show ?case
-  proof cases
-    case prems: 1
-    show ?thesis
-    proof (cases "p = 0")
-      case True
-      then show ?thesis by auto
-    next
-      case False
-      from prems have "degree q = 0 \<or> pcompose p q = 0"
-        by (auto simp add: degree_mult_eq_0)
-      moreover have False if "pcompose p q = 0" "degree q \<noteq> 0"
-      proof -
-        from pCons.hyps(2) that have "degree p = 0"
-          by auto
-        then obtain a1 where "p = [:a1:]"
-          by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases)
-        with \<open>pcompose p q = 0\<close> \<open>p \<noteq> 0\<close> show False
-          by auto
-      qed
-      ultimately have "degree (pCons a p) * degree q = 0"
-        by auto
-      moreover have "degree (pcompose (pCons a p) q) = 0"
-      proof -
-        from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))"
-          by simp
-        also have "\<dots> \<ge> degree ([:a:] + q * pcompose p q)"
-          by (rule degree_add_le_max)
-        finally show ?thesis
-          by (auto simp add: pcompose_pCons)
-      qed
-      ultimately show ?thesis by simp
-    qed
-  next
-    case prems: 2
-    then have "p \<noteq> 0" "q \<noteq> 0" "pcompose p q \<noteq> 0"
-      by auto
-    from prems degree_add_eq_right [of "[:a:]"]
-    have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)"
-      by (auto simp: pcompose_pCons)
-    with pCons.hyps(2) degree_mult_eq[OF \<open>q\<noteq>0\<close> \<open>pcompose p q\<noteq>0\<close>] show ?thesis
-      by auto
-  qed
-qed
-
-lemma pcompose_eq_0:
-  fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  assumes "pcompose p q = 0" "degree q > 0"
-  shows "p = 0"
-proof -
-  from assms degree_pcompose [of p q] have "degree p = 0"
-    by auto
-  then obtain a where "p = [:a:]"
-    by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases)
-  with assms(1) have "a = 0"
-    by auto
-  with \<open>p = [:a:]\<close> show ?thesis
-    by simp
-qed
-
-lemma lead_coeff_comp:
-  fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  assumes "degree q > 0"
-  shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)"
-proof (induct p)
-  case 0
-  then show ?case by auto
-next
-  case (pCons a p)
-  consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0"
-    by blast
-  then show ?case
-  proof cases
-    case prems: 1
-    then have "pcompose p q = 0"
-      by (metis assms degree_0 degree_mult_eq_0 neq0_conv)
-    with pcompose_eq_0[OF _ \<open>degree q > 0\<close>] have "p = 0"
-      by simp
-    then show ?thesis
-      by auto
-  next
-    case prems: 2
-    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 "\<dots> = 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 "\<dots> = lead_coeff p * lead_coeff q ^ (degree p + 1)"
-      by (auto simp: mult_ac)
-    finally show ?thesis by auto
-  qed
-qed
-
-
-subsection \<open>Shifting polynomials\<close>
-
-definition poly_shift :: "nat \<Rightarrow> 'a::zero poly \<Rightarrow> 'a poly"
-  where "poly_shift n p = Abs_poly (\<lambda>i. coeff p (i + n))"
-
-lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)"
-  by (auto simp add: nth_default_def add_ac)
-
-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)
-  then have "\<forall>k>m. coeff p (k + n) = 0"
-    by auto
-  then have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0"
-    by (auto simp: MOST_nat)
-  then show ?thesis
-    by (simp add: poly_shift_def poly.Abs_poly_inverse)
-qed
-
-lemma poly_shift_id [simp]: "poly_shift 0 = (\<lambda>x. x)"
-  by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift)
-
-lemma poly_shift_0 [simp]: "poly_shift n 0 = 0"
-  by (simp add: poly_eq_iff coeff_poly_shift)
-
-lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)"
-  by (simp add: poly_eq_iff coeff_poly_shift)
-
-lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \<ge> n then monom c (m - n) else 0)"
-  by (auto simp add: poly_eq_iff coeff_poly_shift)
-
-lemma coeffs_shift_poly [code abstract]:
-  "coeffs (poly_shift n p) = drop n (coeffs p)"
-proof (cases "p = 0")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  then show ?thesis
-    by (intro coeffs_eqI)
-      (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq)
-qed
-
-
-subsection \<open>Truncating polynomials\<close>
-
-definition poly_cutoff
-  where "poly_cutoff n p = Abs_poly (\<lambda>k. if k < n then coeff p k else 0)"
-
-lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)"
-  unfolding poly_cutoff_def
-  by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n])
-
-lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0"
-  by (simp add: poly_eq_iff coeff_poly_cutoff)
-
-lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)"
-  by (simp add: poly_eq_iff coeff_poly_cutoff)
-
-lemma coeffs_poly_cutoff [code abstract]:
-  "coeffs (poly_cutoff n p) = strip_while (op = 0) (take n (coeffs p))"
-proof (cases "strip_while (op = 0) (take n (coeffs p)) = []")
-  case True
-  then have "coeff (poly_cutoff n p) k = 0" for k
-    unfolding coeff_poly_cutoff
-    by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth)
-  then have "poly_cutoff n p = 0"
-    by (simp add: poly_eq_iff)
-  then show ?thesis
-    by (subst True) simp_all
-next
-  case False
-  have "no_trailing (op = 0) (strip_while (op = 0) (take n (coeffs p)))"
-    by simp
-  with False have "last (strip_while (op = 0) (take n (coeffs p))) \<noteq> 0"
-    unfolding no_trailing_unfold by auto
-  then show ?thesis
-    by (intro coeffs_eqI)
-      (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq)
-qed
-
-
-subsection \<open>Reflecting polynomials\<close>
-
-definition reflect_poly :: "'a::zero poly \<Rightarrow> 'a poly"
-  where "reflect_poly p = Poly (rev (coeffs p))"
-
-lemma coeffs_reflect_poly [code abstract]:
-  "coeffs (reflect_poly p) = rev (dropWhile (op = 0) (coeffs p))"
-  by (simp add: reflect_poly_def)
-
-lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0"
-  by (simp add: reflect_poly_def)
-
-lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1"
-  by (simp add: reflect_poly_def one_poly_def)
-
-lemma coeff_reflect_poly:
-  "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))"
-  by (cases "p = 0")
-    (auto simp add: reflect_poly_def nth_default_def
-      rev_nth degree_eq_length_coeffs coeffs_nth not_less
-      dest: le_imp_less_Suc)
-
-lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0"
-  by (simp add: coeff_reflect_poly)
-
-lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0"
-  by (simp add: coeff_reflect_poly poly_0_coeff_0)
-
-lemma reflect_poly_pCons':
-  "p \<noteq> 0 \<Longrightarrow> reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))"
-  by (intro poly_eqI)
-    (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split)
-
-lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]"
-  by (cases "a = 0") (simp_all add: reflect_poly_def)
-
-lemma poly_reflect_poly_nz:
-  "x \<noteq> 0 \<Longrightarrow> poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)"
-  for x :: "'a::field"
-  by (induct 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)
-
-lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p"
-  by (simp add: poly_0_coeff_0)
-
-lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> reflect_poly (reflect_poly p) = p"
-  by (cases p rule: pCons_cases) (simp add: reflect_poly_def )
-
-lemma degree_reflect_poly_le: "degree (reflect_poly p) \<le> degree p"
-  by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono)
-
-lemma reflect_poly_pCons: "a \<noteq> 0 \<Longrightarrow> reflect_poly (pCons a p) = Poly (rev (a # coeffs p))"
-  by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly)
-
-lemma degree_reflect_poly_eq [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> degree (reflect_poly p) = degree p"
-  by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs)
-
-(* TODO: does this work with zero divisors as well? Probably not. *)
-lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q"
-  for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-proof (cases "p = 0 \<or> q = 0")
-  case False
-  then have [simp]: "p \<noteq> 0" "q \<noteq> 0" by auto
-  show ?thesis
-  proof (rule poly_eqI)
-    show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i
-    proof (cases "i \<le> degree (p * q)")
-      case True
-      define A where "A = {..i} \<inter> {i - degree q..degree p}"
-      define B where "B = {..degree p} \<inter> {degree p - i..degree (p*q) - i}"
-      let ?f = "\<lambda>j. degree p - j"
-
-      from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)"
-        by (simp add: coeff_reflect_poly)
-      also have "\<dots> = (\<Sum>j\<le>degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))"
-        by (simp add: coeff_mult)
-      also have "\<dots> = (\<Sum>j\<in>B. coeff p j * coeff q (degree (p * q) - i - j))"
-        by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0)
-      also from True have "\<dots> = (\<Sum>j\<in>A. coeff p (degree p - j) * coeff q (degree q - (i - j)))"
-        by (intro sum.reindex_bij_witness[of _ ?f ?f])
-          (auto simp: A_def B_def degree_mult_eq add_ac)
-      also have "\<dots> =
-        (\<Sum>j\<le>i.
-          if j \<in> {i - degree q..degree p}
-          then coeff p (degree p - j) * coeff q (degree q - (i - j))
-          else 0)"
-        by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def)
-      also have "\<dots> = coeff (reflect_poly p * reflect_poly q) i"
-        by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong)
-      finally show ?thesis .
-    qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral)
-  qed
-qed auto
-
-lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)"
-  for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  using reflect_poly_mult[of "[:c:]" p] by simp
-
-lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  by (induct n) (simp_all add: reflect_poly_mult)
-
-lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\<lambda>x. reflect_poly (f x)) A"
-  for f :: "_ \<Rightarrow> _::{comm_semiring_0,semiring_no_zero_divisors} poly"
-  by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult)
-
-lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)"
-  for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list"
-  by (induct xs) (simp_all add: reflect_poly_mult)
-
-lemma reflect_poly_Poly_nz:
-  "no_trailing (HOL.eq 0) xs \<Longrightarrow> reflect_poly (Poly xs) = Poly (rev xs)"
-  by (simp add: reflect_poly_def)
-
-lemmas reflect_poly_simps =
-  reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult
-  reflect_poly_power reflect_poly_prod reflect_poly_prod_list
-
-
-subsection \<open>Derivatives\<close>
-
-function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \<Rightarrow> 'a poly"
-  where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))"
-  by (auto intro: pCons_cases)
-
-termination pderiv
-  by (relation "measure degree") simp_all
-
-declare pderiv.simps[simp del]
-
-lemma pderiv_0 [simp]: "pderiv 0 = 0"
-  using pderiv.simps [of 0 0] by simp
-
-lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)"
-  by (simp add: pderiv.simps)
-
-lemma pderiv_1 [simp]: "pderiv 1 = 0"
-  by (simp add: one_poly_def pderiv_pCons)
-
-lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0"
-  and pderiv_numeral [simp]: "pderiv (numeral m) = 0"
-  by (simp_all add: of_nat_poly numeral_poly pderiv_pCons)
-
-lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)"
-  by (induct p arbitrary: n)
-    (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split)
-
-fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \<Rightarrow> 'a list \<Rightarrow> 'a list"
-  where
-    "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)"
-  | "pderiv_coeffs_code f [] = []"
-
-definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \<Rightarrow> 'a list"
-  where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)"
-
-(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *)
-lemma pderiv_coeffs_code:
-  "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n"
-proof (induct xs arbitrary: f n)
-  case Nil
-  then show ?case by simp
-next
-  case (Cons x xs)
-  show ?case
-  proof (cases n)
-    case 0
-    then show ?thesis
-      by (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0") (auto simp: cCons_def)
-  next
-    case n: (Suc m)
-    show ?thesis
-    proof (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0")
-      case False
-      then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n =
-          nth_default 0 (pderiv_coeffs_code (f + 1) xs) m"
-        by (auto simp: cCons_def n)
-      also have "\<dots> = (f + of_nat n) * nth_default 0 xs m"
-        by (simp add: Cons n add_ac)
-      finally show ?thesis
-        by (simp add: n)
-    next
-      case True
-      have empty: "pderiv_coeffs_code g xs = [] \<Longrightarrow> g + of_nat m = 0 \<or> nth_default 0 xs m = 0" for g
-      proof (induct xs arbitrary: g m)
-        case Nil
-        then show ?case by simp
-      next
-        case (Cons x xs)
-        from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \<or> x = 0"
-          by (auto simp: cCons_def split: if_splits)
-        note IH = Cons(1)[OF empty]
-        from IH[of m] IH[of "m - 1"] g show ?case
-          by (cases m) (auto simp: field_simps)
-      qed
-      from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0"
-        by (auto simp: cCons_def n)
-      moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0"
-        by (simp add: n) (use empty[of "f+1"] in \<open>auto simp: field_simps\<close>)
-      ultimately show ?thesis by simp
-    qed
-  qed
-qed
-
-lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda>i. f (Suc i)) [0 ..< n]"
-  by (induct n arbitrary: f) auto
-
-lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)"
-  unfolding pderiv_coeffs_def
-proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases)
-  case (1 n)
-  have id: "coeff p (Suc n) = nth_default 0 (map (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n"
-    by (cases "n < degree p") (auto simp: nth_default_def coeff_eq_0)
-  show ?case
-    unfolding coeffs_def map_upt_Suc by (auto simp: id)
-next
-  case 2
-  obtain n :: 'a and xs where defs: "tl (coeffs p) = xs" "1 = n"
-    by simp
-  from 2 show ?case
-    unfolding defs by (induct xs arbitrary: n) (auto simp: cCons_def)
-qed
-
-lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
-  apply (rule iffI)
-   apply (cases p)
-   apply simp
-   apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc)
-  apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0)
-  done
-
-lemma degree_pderiv: "degree (pderiv p) = degree p - 1"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
-  apply (rule order_antisym [OF degree_le])
-   apply (simp add: coeff_pderiv coeff_eq_0)
-  apply (cases "degree p")
-   apply simp
-  apply (rule le_degree)
-  apply (simp add: coeff_pderiv del: of_nat_Suc)
-  apply (metis degree_0 leading_coeff_0_iff nat.distinct(1))
-  done
-
-lemma not_dvd_pderiv:
-  fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
-  assumes "degree p \<noteq> 0"
-  shows "\<not> p dvd pderiv p"
-proof
-  assume dvd: "p dvd pderiv p"
-  then obtain q where p: "pderiv p = p * q"
-    unfolding dvd_def by auto
-  from dvd have le: "degree p \<le> degree (pderiv p)"
-    by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff)
-  from assms and this [unfolded degree_pderiv]
-    show False by auto
-qed
-
-lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \<longleftrightarrow> degree p = 0"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
-  using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric])
-
-lemma pderiv_singleton [simp]: "pderiv [:a:] = 0"
-  by (simp add: pderiv_pCons)
-
-lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q"
-  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
-
-lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p"
-  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
-
-lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q"
-  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
-
-lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)"
-  by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps)
-
-lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p"
-  by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps)
-
-lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p"
-  apply (induct n)
-   apply simp
-  apply (subst power_Suc)
-  apply (subst pderiv_mult)
-  apply (erule ssubst)
-  apply (simp only: of_nat_Suc smult_add_left smult_1_left)
-  apply (simp add: algebra_simps)
-  done
-
-lemma pderiv_prod: "pderiv (prod f (as)) = (\<Sum>a\<in>as. prod f (as - {a}) * pderiv (f a))"
-proof (induct as rule: infinite_finite_induct)
-  case (insert a as)
-  then have id: "prod f (insert a as) = f a * prod f as"
-    "\<And>g. sum g (insert a as) = g a + sum g as"
-    "insert a as - {a} = as"
-    by auto
-  have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \<in> as" for b
-  proof -
-    from \<open>a \<notin> as\<close> that have *: "insert a as - {b} = insert a (as - {b})"
-      by auto
-    show ?thesis
-      unfolding * by (subst prod.insert) (use insert in auto)
-  qed
-  then show ?case
-    unfolding id pderiv_mult insert(3) sum_distrib_left
-    by (auto simp add: ac_simps intro!: sum.cong)
-qed auto
-
-lemma DERIV_pow2: "DERIV (\<lambda>x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"
-  by (rule DERIV_cong, rule DERIV_pow) simp
-declare DERIV_pow2 [simp] DERIV_pow [simp]
-
-lemma DERIV_add_const: "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. a + f x :: 'a::real_normed_field) x :> D"
-  by (rule DERIV_cong, rule DERIV_add) auto
-
-lemma poly_DERIV [simp]: "DERIV (\<lambda>x. poly p x) x :> poly (pderiv p) x"
-  by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons)
-
-lemma continuous_on_poly [continuous_intros]:
-  fixes p :: "'a :: {real_normed_field} poly"
-  assumes "continuous_on A f"
-  shows "continuous_on A (\<lambda>x. poly p (f x))"
-proof -
-  have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))"
-    by (intro continuous_intros assms)
-  also have "\<dots> = (\<lambda>x. poly p (f x))"
-    by (rule ext) (simp add: poly_altdef mult_ac)
-  finally show ?thesis .
-qed
-
-text \<open>Consequences of the derivative theorem above.\<close>
-
-lemma poly_differentiable[simp]: "(\<lambda>x. poly p x) differentiable (at x)"
-  for x :: real
-  by (simp add: real_differentiable_def) (blast intro: poly_DERIV)
-
-lemma poly_isCont[simp]: "isCont (\<lambda>x. poly p x) x"
-  for x :: real
-  by (rule poly_DERIV [THEN DERIV_isCont])
-
-lemma poly_IVT_pos: "a < b \<Longrightarrow> poly p a < 0 \<Longrightarrow> 0 < poly p b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0"
-  for a b :: real
-  using IVT_objl [of "poly p" a 0 b] by (auto simp add: order_le_less)
-
-lemma poly_IVT_neg: "a < b \<Longrightarrow> 0 < poly p a \<Longrightarrow> poly p b < 0 \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0"
-  for a b :: real
-  using poly_IVT_pos [where p = "- p"] by simp
-
-lemma poly_IVT: "a < b \<Longrightarrow> poly p a * poly p b < 0 \<Longrightarrow> \<exists>x>a. x < b \<and> poly p x = 0"
-  for p :: "real poly"
-  by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos)
-
-lemma poly_MVT: "a < b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p b - poly p a = (b - a) * poly (pderiv p) x"
-  for a b :: real
-  using MVT [of a b "poly p"]
-  apply auto
-  apply (rule_tac x = z in exI)
-  apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique])
-  done
-
-lemma poly_MVT':
-  fixes a b :: real
-  assumes "{min a b..max a b} \<subseteq> A"
-  shows "\<exists>x\<in>A. poly p b - poly p a = (b - a) * poly (pderiv p) x"
-proof (cases a b rule: linorder_cases)
-  case less
-  from poly_MVT[OF less, of p] guess x by (elim exE conjE)
-  then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms])
-next
-  case greater
-  from poly_MVT[OF greater, of p] guess x by (elim exE conjE)
-  then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms])
-qed (use assms in auto)
-
-lemma poly_pinfty_gt_lc:
-  fixes p :: "real poly"
-  assumes "lead_coeff p > 0"
-  shows "\<exists>n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p"
-  using assms
-proof (induct p)
-  case 0
-  then show ?case by auto
-next
-  case (pCons a p)
-  from this(1) consider "a \<noteq> 0" "p = 0" | "p \<noteq> 0" by auto
-  then show ?case
-  proof cases
-    case 1
-    then show ?thesis by auto
-  next
-    case 2
-    with pCons obtain n1 where gte_lcoeff: "\<forall>x\<ge>n1. lead_coeff p \<le> poly p x"
-      by auto
-    from pCons(3) \<open>p \<noteq> 0\<close> have gt_0: "lead_coeff p > 0" by auto
-    define n where "n = max n1 (1 + \<bar>a\<bar> / lead_coeff p)"
-    have "lead_coeff (pCons a p) \<le> poly (pCons a p) x" if "n \<le> x" for x
-    proof -
-      from gte_lcoeff that have "lead_coeff p \<le> poly p x"
-        by (auto simp: n_def)
-      with gt_0 have "\<bar>a\<bar> / lead_coeff p \<ge> \<bar>a\<bar> / poly p x" and "poly p x > 0"
-        by (auto intro: frac_le)
-      with \<open>n \<le> x\<close>[unfolded n_def] have "x \<ge> 1 + \<bar>a\<bar> / poly p x"
-        by auto
-      with \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x > 0\<close> \<open>p \<noteq> 0\<close>
-      show "lead_coeff (pCons a p) \<le> poly (pCons a p) x"
-        by (auto simp: field_simps)
-    qed
-    then show ?thesis by blast
-  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)"
-  by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons)
-
-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: "\<not> [:- a, 1:] dvd q"
-  shows "n = Suc (order a (pderiv p))"
-proof -
-  from assms have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0"
-    by auto
-  from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
-    by (cases n) auto
-  have *: "k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l" for k 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
-    show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
-      apply (subst lemma_order_pderiv1)
-      apply (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one)
-      done
-  qed
-  then show ?thesis
-    by (metis \<open>n = Suc n'\<close> pe)
-qed
-
-lemma order_pderiv: "pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a p = Suc (order a (pderiv p))"
-  for p :: "'a::field_char_0 poly"
-  apply (cases "p = 0")
-   apply 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:
-  fixes p :: "'a::field_char_0 poly"
-  assumes "pderiv p \<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: "\<not> ?thesis"
-  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)
-  from \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> have "order a (pderiv p) = order a e + order a d"
-    by (simp add: order_mult)
-  from \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> have "order a p = Suc (order a (pderiv p))"
-    by (rule order_pderiv)
-  from \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> have "d \<noteq> 0"
-    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
-  with \<open>d \<noteq> 0\<close> have "order a (pderiv p) \<le> order a d"
-    by (simp add: order_divides)
-  show ?thesis
-    using \<open>order a p = order a q + order a d\<close>
-      and \<open>order a (pderiv p) = order a e + order a d\<close>
-      and \<open>order a p = Suc (order a (pderiv p))\<close>
-      and \<open>order a (pderiv p) \<le> order a d\<close>
-    by auto
-qed
-
-lemma poly_squarefree_decomp_order2:
-  "pderiv p \<noteq> 0 \<Longrightarrow> p = q * d \<Longrightarrow> pderiv p = e * d \<Longrightarrow>
-    d = r * p + s * pderiv p \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-  for p :: "'a::field_char_0 poly"
-  by (blast intro: poly_squarefree_decomp_order)
-
-lemma order_pderiv2:
-  "pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a (pderiv p) = n \<longleftrightarrow> order a p = Suc n"
-  for p :: "'a::field_char_0 poly"
-  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:]"
-  for p :: "'a::{semidom,semiring_char_0} poly"
-  by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits)
-
-lemma rsquarefree_roots: "rsquarefree p \<longleftrightarrow> (\<forall>a. \<not> (poly p a = 0 \<and> poly (pderiv p) a = 0))"
-  for p :: "'a::field_char_0 poly"
-  apply (simp add: rsquarefree_def)
-  apply (case_tac "p = 0")
-   apply simp
-  apply 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:
-  fixes p :: "'a::field_char_0 poly"
-  assumes "pderiv p \<noteq> 0"
-    and "p = q * d"
-    and "pderiv p = e * d"
-    and "d = r * p + s * pderiv p"
-  shows "rsquarefree q \<and> (\<forall>a. poly q a = 0 \<longleftrightarrow> 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
-  from assms have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-    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>
-
-text \<open>
-  Algebraic numbers can be defined in two equivalent ways: all real numbers that are
-  roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry
-  uses the rational definition, but we need the integer definition.
-
-  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>
-
-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)"
-
-lemma algebraicI: "(\<And>i. coeff p i \<in> \<int>) \<Longrightarrow> p \<noteq> 0 \<Longrightarrow> poly p x = 0 \<Longrightarrow> algebraic x"
-  unfolding algebraic_def by blast
-
-lemma algebraicE:
-  assumes "algebraic x"
-  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 algebraic_altdef: "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
-  for p :: "'a::field_char_0 poly"
-proof safe
-  fix p
-  assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0"
-  define cs where "cs = coeffs p"
-  from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'"
-    unfolding Rats_def by blast
-  then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i
-    by (subst (asm) bchoice_iff) blast
-  define cs' where "cs' = map (quotient_of \<circ> f) (coeffs p)"
-  define d where "d = Lcm (set (map snd cs'))"
-  define p' where "p' = smult (of_int d) p"
-
-  have "coeff p' n \<in> \<int>" for n
-  proof (cases "n \<le> degree p")
-    case True
-    define c where "c = coeff p n"
-    define a where "a = fst (quotient_of (f (coeff p n)))"
-    define b where "b = snd (quotient_of (f (coeff p n)))"
-    have b_pos: "b > 0"
-      unfolding b_def using quotient_of_denom_pos' by simp
-    have "coeff p' n = of_int d * coeff p n"
-      by (simp add: p'_def)
-    also have "coeff p n = of_rat (of_int a / of_int b)"
-      unfolding a_def b_def
-      by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric])
-    also have "of_int d * \<dots> = of_rat (of_int (a*d) / of_int b)"
-      by (simp add: of_rat_mult of_rat_divide)
-    also from nz True have "b \<in> snd ` set cs'"
-      by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc)
-    then have "b dvd (a * d)"
-      by (simp add: d_def)
-    then have "of_int (a * d) / of_int b \<in> (\<int> :: rat set)"
-      by (rule of_int_divide_in_Ints)
-    then have "of_rat (of_int (a * d) / of_int b) \<in> \<int>" by (elim Ints_cases) auto
-    finally show ?thesis .
-  next
-    case False
-    then show ?thesis
-      by (auto simp: p'_def not_le coeff_eq_0)
-  qed
-  moreover have "set (map snd cs') \<subseteq> {0<..}"
-    unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc)
-  then have "d \<noteq> 0"
-    unfolding d_def by (induct cs') simp_all
-  with nz have "p' \<noteq> 0" by (simp add: p'_def)
-  moreover from root have "poly p' x = 0"
-    by (simp add: p'_def)
-  ultimately show "algebraic x"
-    unfolding algebraic_def by blast
-next
-  assume "algebraic x"
-  then obtain p where p: "coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0" for i
-    by (force simp: algebraic_def)
-  moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i
-    by (elim Ints_cases) simp
-  ultimately show "\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0" by auto
-qed
-
-
-subsection \<open>Content and primitive part of a polynomial\<close>
-
-definition content :: "'a::semiring_gcd poly \<Rightarrow> 'a"
-  where "content p = gcd_list (coeffs p)"
-
-lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0"
-  by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps)
-
-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: "[:c:] dvd p \<longleftrightarrow> c dvd content p"
-  for c :: "'a::semiring_gcd"
-proof (cases "p = 0")
-  case True
-  then show ?thesis by simp
-next
-  case 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"
-    then show "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_fin_iff dvd_mult_unit_iff)
-  finally show ?thesis .
-qed
-
-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"
-proof (cases "p = 0")
-  case True
-  then show ?thesis
-    by simp
-next
-  case False
-  then show ?thesis
-    by (cases "n \<le> degree p")
-      (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd)
-qed
-
-lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
-  by (simp add: content_def Gcd_fin_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)"
-  then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
-  then show "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_fin_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 poly \<Rightarrow> 'a poly"
-  where "primitive_part p = 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]: "smult (content p) (primitive_part p) = p"
-  for p :: "'a :: semiring_gcd poly"
-proof (cases "p = 0")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  then show ?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
-
-lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
-proof (cases "p = 0")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  then have "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
-
-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 (primitive_part p) * content p"
-    by (simp del: content_times_primitive_part add: ac_simps)
-  finally have "1 * content p = content (primitive_part p) * content p"
-    by simp
-  then have "1 * content p div content p = content (primitive_part p) * content p div content p"
-    by simp
-  with assms show ?thesis
-    by simp
-qed
-
-lemma content_decompose:
-  obtains p' :: "'a::semiring_gcd poly" where "p = smult (content p) p'" "content p' = 1"
-proof (cases "p = 0")
-  case True
-  then show ?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 (rule 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 True
-  then show ?thesis by simp
-next
-  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
-
-
-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'"
-    and "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> d * r = 0"
-  shows "q' = q + r"
-  using assms(3-)
-proof (induct n arbitrary: q r dr)
-  case (Suc n)
-  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)
-  from Suc(4) have dr: "dr = n + degree d" by auto
-  from d have lc: "lc \<noteq> 0" by auto
-  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
-  proof (cases "?qq = 0")
-    case True
-    then show ?thesis by simp
-  next
-    case False
-    then have n: "n = degree b"
-      by (simp add: degree_monom_eq)
-    show ?thesis
-      unfolding n dr by (simp add: coeff_mult_degree_sum)
-  qed
-  also have "\<dots> = lc * coeff b n"
-    by (simp add: d)
-  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
-    with Suc(2) show ?thesis
-      unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps)
-  next
-    case False
-    from dr Suc(2) have "degree r \<le> n"
-      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
-    then have right: "lc * coeff r n = 0"
-      by (simp add: coeff_eq_0)
-    have "coeff (d * r) dr = coeff (d * r) (degree d + n)"
-      by (simp add: dr ac_simps)
-    also from r_n have "\<dots> = 0"
-      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
-      by (simp only: 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]
-  from Suc(4) have dr: "dr = n + degree d" by auto
-  from Suc(2) have deg_rr: "degree ?rr \<le> dr" 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
-  next
-    case _: Suc
-    with Suc(4) show ?thesis by auto
-  qed
-  from IH[OF deg_rrr this] show ?case
-    by simp
-next
-  case 0
-  show ?case
-  proof (cases "r = 0")
-    case True
-    with 0 show ?thesis by auto
-  next
-    case False
-    from d False have "degree (d * r) = degree d + degree r"
-      by (subst degree_mult_eq) auto
-    with 0 d show ?thesis 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 0
-  then show ?case by simp
-next
-  case Suc
-  show ?case
-    unfolding divide_poly_main.simps[of _ _ r] Let_def
-    by (simp add: Suc del: divide_poly_main.simps)
-qed
-
-lemma divide_poly:
-  assumes g: "g \<noteq> 0"
-  shows "(f * g) div g = (f :: 'a poly)"
-proof -
-  have len: "length (coeffs f) = Suc (degree f)" if "f \<noteq> 0" for f :: "'a poly"
-    using that unfolding degree_eq_length_coeffs by auto
-  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"
-    by (simp add: divide_poly_def Let_def ac_simps)
-  note main = divide_poly_main[OF g refl le_refl 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 simp
-qed
-
-lemma divide_poly_0: "f div 0 = 0"
-  for f :: "'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 True
-  then show ?thesis
-    by (auto intro!: poly_eqI simp: coeff_map_poly)
-next
-  case False
-  from assms obtain q where p: "p = [:c:] * q" by (rule dvdE)
-  moreover {
-    have "smult c q = [:c:] * q"
-      by simp
-    also have "\<dots> div [:c:] = q"
-      by (rule nonzero_mult_div_cancel_left) (use False in auto)
-    finally have "smult c q div [:c:] = q" .
-  }
-  ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
-qed
-
-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: "a \<noteq> 0 \<Longrightarrow> is_unit [:a:]"
-  for a :: "'a::field"
-  by (simp add: is_unit_monom_0 monom_0 [symmetric])
-
-lemma is_unit_iff_degree:
-  fixes p :: "'a::field poly"
-  assumes "p \<noteq> 0"
-  shows "is_unit p \<longleftrightarrow> degree p = 0"
-    (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?rhs
-  then obtain a where "p = [:a:]"
-    by (rule degree_eq_zeroE)
-  with assms show ?lhs
-    by (simp add: is_unit_triv)
-next
-  assume ?lhs
-  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 ?rhs by simp
-qed
-
-lemma is_unit_pCons_iff: "is_unit (pCons a p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
-  for p :: "'a::field poly"
-  by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree)
-
-lemma is_unit_monom_trival: "is_unit p \<Longrightarrow> monom (coeff p (degree p)) 0 = p"
-  for p :: "'a::field poly"
-  by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
-
-lemma is_unit_const_poly_iff: "[:c:] dvd 1 \<longleftrightarrow> c dvd 1"
-  for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
-  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:]" .
-  with \<open>p dvd 1\<close> have "c dvd 1"
-    by (simp add: is_unit_const_poly_iff)
-  with c show thesis ..
-qed
-
-lemma is_unit_polyE':
-  fixes p :: "'a::field poly"
-  assumes "is_unit p"
-  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: "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
-  for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
-  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')"
-    and "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 assms(3-)
-proof (induct n arbitrary: q r dr)
-  case 0
-  then show ?case by auto
-next
-  case (Suc n)
-  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)
-  from Suc(4) have dr: "dr = n + degree d" by auto
-  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
-  proof (cases "?qq = 0")
-    case True
-    then show ?thesis by auto
-  next
-    case False
-    then have n: "n = degree b"
-      by (simp add: degree_monom_eq)
-    show ?thesis
-      unfolding n dr by (simp add: coeff_mult_degree_sum)
-  qed
-  also have "\<dots> = lc * coeff b n" by (simp add: d)
-  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
-  from Suc(4) have dr: "dr = n + degree d" by auto
-  have deg_rr: "degree ?rr \<le> dr"
-    using Suc(2) 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
-    then have "\<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
-  next
-    case _: Suc
-    with Suc(4) show ?thesis by auto
-  qed
-  note IH = Suc(1)[OF deg_rrr res this]
-  show ?case
-  proof (intro conjI)
-    from IH show "r' = 0 \<or> degree r' < degree d"
-      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
-
-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]
-  from g 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"
-    by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs)
-  note main' = main[OF this]
-  then 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,
-        cases "f = 0"; cases "coeffs g", use g in 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_all 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 :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly"
-  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"
-  from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)"
-    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"
-    by (auto simp: a_def)
-  show "r = 0 \<or> degree r < degree g" by fact
-  from g have "a \<noteq> 0"
-    by (auto simp: a_def)
-  with id show "\<exists>a q. a \<noteq> 0 \<and> smult a f = g * q + r"
-    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 \<dots> = 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"
-    by (simp only: Suc[unfolded divide_poly_main.simps Let_def])
-  also have "\<dots> = 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
-
-
-subsubsection \<open>Division in polynomials over fields\<close>
-
-lemma pseudo_divmod_field:
-  fixes g :: "'a::field poly"
-  assumes g: "g \<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"
-    by (auto simp: c_def)
-  from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r"
-    by auto
-  also have "smult (1 / c) \<dots> = 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:
-  fixes d :: "'a::field poly"
-  assumes d: "d \<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 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"
-    by (auto simp: f'_def)
-  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 :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom
-begin
-
-definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
-  where "unit_factor_poly p = [:unit_factor (lead_coeff p):]"
-
-definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
-  where "normalize p = p div [:unit_factor (lead_coeff p):]"
-
-instance
-proof
-  fix p :: "'a poly"
-  show "unit_factor p * normalize p = p"
-  proof (cases "p = 0")
-    case True
-    then show ?thesis
-      by (simp add: unit_factor_poly_def normalize_poly_def)
-  next
-    case False
-    then have "lead_coeff p \<noteq> 0"
-      by simp
-    then have *: "unit_factor (lead_coeff p) \<noteq> 0"
-      using unit_factor_is_unit [of "lead_coeff p"] by auto
-    then have "unit_factor (lead_coeff p) dvd 1"
-      by (auto intro: unit_factor_is_unit)
-    then have **: "unit_factor (lead_coeff p) dvd c" for c
-      by (rule dvd_trans) simp
-    have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c
-    proof -
-      from ** obtain b where "c = unit_factor (lead_coeff p) * b" ..
-      with False * show ?thesis by simp
-    qed
-    have "p div [:unit_factor (lead_coeff p):] =
-      map_poly (\<lambda>c. c div unit_factor (lead_coeff p)) p"
-      by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **)
-    then show ?thesis
-      by (simp add: normalize_poly_def unit_factor_poly_def
-        smult_conv_map_poly map_poly_map_poly o_def ***)
-  qed
-next
-  fix p :: "'a poly"
-  assume "is_unit p"
-  then obtain c where p: "p = [:c:]" "c dvd 1"
-    by (auto simp: is_unit_poly_iff)
-  then show "unit_factor p = p"
-    by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor)
-next
-  fix p :: "'a poly"
-  assume "p \<noteq> 0"
-  then show "is_unit (unit_factor p)"
-    by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit)
-qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
-
-end
-
-lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
-proof -
-  have "[:unit_factor (lead_coeff p):] dvd p"
-    by (metis unit_factor_poly_def unit_factor_self)
-  then show ?thesis
-    by (simp add: normalize_poly_def div_const_poly_conv_map_poly)
-qed
-
-lemma coeff_normalize [simp]:
-  "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)"
-  by (simp add: normalize_poly_eq_map_poly coeff_map_poly)
-
-class field_unit_factor = field + unit_factor +
-  assumes unit_factor_field [simp]: "unit_factor = id"
-begin
-
-subclass semidom_divide_unit_factor
-proof
-  fix a
-  assume "a \<noteq> 0"
-  then have "1 = a * inverse a" by simp
-  then have "a dvd 1" ..
-  then show "unit_factor a dvd 1" by simp
-qed simp_all
-
-end
-
-lemma unit_factor_pCons:
-  "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] 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_eq_map_poly degree_monom_eq)
-
-lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]"
-  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_eq_map_poly 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
-
-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
-
-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)"
-  by (simp add: eucl_rel_poly_iff)
-
-lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)"
-  by (simp add: eucl_rel_poly_iff)
-
-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 -
-  from assms have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
-    by (simp_all add: eucl_rel_poly_iff)
-  from b x have "?x = ?q * y + ?r" by simp
-  moreover
-  have "?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)
-      from r show "degree (pCons a r) \<le> degree y"
-        by auto
-      show "degree (smult b y) \<le> degree y"
-        by (rule degree_smult_le)
-    qed
-    from \<open>y \<noteq> 0\<close> show "coeff ?r (degree y) = 0"
-      by (simp add: b)
-  qed
-  ultimately 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 classical)
-    assume "\<not> ?thesis"
-    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 from \<open>q1 \<noteq> q2\<close> have "\<dots> = degree ((q1 - q2) * y)"
-      by (simp add: degree_mult_eq)
-    also from q3 have "\<dots> = degree (r2 - r1)"
-      by simp
-    finally have "degree (r2 - r1) < degree (r2 - r1)" .
-    then show ?thesis 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) semidom_modulo
-begin
-
-definition modulo_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  where mod_poly_def: "f mod g =
-    (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)"
-
-instance
-proof
-  fix x y :: "'a poly"
-  show "x div y * y + x mod y = x"
-  proof (cases "y = 0")
-    case True
-    then show ?thesis
-      by (simp add: divide_poly_0 mod_poly_def)
-  next
-    case False
-    then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y =
-        (x div y, x mod y)"
-      by (simp add: divide_poly_field mod_poly_def pseudo_mod_def)
-    with False pseudo_divmod [OF False this] show ?thesis
-      by (simp add: power_mult_distrib [symmetric] ac_simps)
-  qed
-qed
-
-end
-
-lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)"
-  unfolding eucl_rel_poly_iff
-proof
-  show "x = x div y * y + x mod y"
-    by (simp add: div_mult_mod_eq)
-  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
-      by (simp add: mod_poly_def)
-  qed
-qed
-
-lemma div_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x div y = q"
-  for x :: "'a::field poly"
-  by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly])
-
-lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x mod y = r"
-  for x :: "'a::field poly"
-  by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly])
-
-instance poly :: (field) ring_div
-proof
-  fix x y z :: "'a poly"
-  assume "y \<noteq> 0"
-  with eucl_rel_poly [of x y] have "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
-    by (simp add: eucl_rel_poly_iff distrib_right)
-  then show "(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: 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
-
-lemma div_pCons_eq:
-  "pCons a p div q =
-    (if q = 0 then 0
-     else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))"
-  using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p]
-  by (auto intro: div_poly_eq)
-
-lemma mod_pCons_eq:
-  "pCons a p mod q =
-    (if q = 0 then pCons a p
-     else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)"
-  using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p]
-  by (auto intro: mod_poly_eq)
-
-lemma div_mod_fold_coeffs:
-  "(p div q, p mod 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))"
-  by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def)
-
-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:
-  fixes x :: "'a::field poly"
-  assumes "degree x < degree y"
-  shows "x div y = 0"
-proof -
-  from assms have "eucl_rel_poly x y (0, x)"
-    by (simp add: eucl_rel_poly_iff)
-  then show "x div y = 0"
-    by (rule div_poly_eq)
-qed
-
-lemma mod_poly_less:
-  assumes "degree x < degree y"
-  shows "x mod y = x"
-proof -
-  from assms have "eucl_rel_poly x y (0, x)"
-    by (simp add: eucl_rel_poly_iff)
-  then show "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)"
-  by (simp add: eucl_rel_poly_iff smult_add_right)
-
-lemma div_smult_left: "(smult a x) div y = smult a (x div y)"
-  for x y :: "'a::field poly"
-  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]: "(- x) div y = - (x div y)"
-  for x y :: "'a::field poly"
-  using div_smult_left [of "- 1::'a"] by simp
-
-lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)"
-  for x y :: "'a::field poly"
-  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: algebra_simps degree_add_less)
-
-lemma poly_div_add_left: "(x + y) div z = x div z + y div z"
-  for x y z :: "'a::field poly"
-  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
-  by (rule div_poly_eq)
-
-lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z"
-  for x y z :: "'a::field poly"
-  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
-  by (rule mod_poly_eq)
-
-lemma poly_div_diff_left: "(x - y) div z = x div z - y div z"
-  for x y z :: "'a::field poly"
-  by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left)
-
-lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z"
-  for x y z :: "'a::field poly"
-  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)"
-  by (simp add: eucl_rel_poly_iff)
-
-lemma div_smult_right: "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
-  for x y :: "'a::field poly"
-  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]: "x div (- y) = - (x div y)"
-  for x y :: "'a::field poly"
-  using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq)
-
-lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y"
-  for x y :: "'a::field poly"
-  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: "x div (y * z) = (x div y) div z"
-  for x y z :: "'a::field poly"
-  by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
-
-lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y"
-  for x y z :: "'a::field poly"
-  by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
-
-lemma mod_pCons:
-  fixes a :: "'a::field"
-    and x y :: "'a::field poly"
-  assumes y: "y \<noteq> 0"
-  defines "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_def
-  by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl])
-
-
-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"
-  for 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) y)"
-  for a :: "'a::ring"
-  by(cases "a = 0") (simp_all add: minus_zero_does_nothing)
-
-lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b"
-  for a :: "'a::comm_semiring_1 list"
-  by (induct a) (auto simp: monom_0 monom_Suc)
-
-lemma minus_poly_rev_list: "length p \<ge> length q \<Longrightarrow>
-  Poly (rev (minus_poly_rev_list (rev p) (rev q))) =
-    Poly p - monom 1 (length p - length q) * Poly q"
-  for p q :: "'a :: comm_ring_1 list"
-proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct)
-  case (1 x xs y ys)
-  then have "length (rev q) \<le> length (rev p)"
-    by simp
-  from this[folded 1(2,3)] have ys_xs: "length ys \<le> length xs"
-    by simp
-  then have *: "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 from ys_xs have "\<dots> =
-    Poly (rev xs) + monom x (length xs) -
-      (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length 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 * diff_monom[symmetric] by simp
-  finally show ?case
-    by (simp add: 1(2,3)[symmetric] 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)) r) (map (op * (hd r)) (rev d))) = 0"
-  for d r :: "'a::comm_ring list"
-proof (induct r)
-  case Nil
-  then show ?case by simp
-next
-  case (Cons a rs)
-  then show ?case by (cases "rev d") (simp_all add: ac_simps)
-qed
-
-lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)"
-proof (induct p)
-  case Nil
-  then show ?case by simp
-next
-  case (Cons x xs)
-  then show ?case by (cases "Poly xs = 0") auto
-qed
-
-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 "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)
-  from Suc.prems have *: "\<not> Suc (length r) \<le> length d"
-    by simp
-  with \<open>d \<noteq> []\<close> have "r \<noteq> []"
-    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)"
-  from * Suc(3) have n: "n = (1 + length r - length d - 1)"
-    by simp
-  from * have rr_val:"(length ?rrr) = (length r - 1)"
-    by auto
-  with \<open>r \<noteq> []\<close> * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)"
-    by auto
-  from * have id: "Suc (length r) - length d = Suc (length r - length d)"
-    by auto
-  from Suc.prems *
-  have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')"
-    by (simp add: Let_def if_0_minus_poly_rev_list id)
-  with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')"
-    by auto
-  from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)"
-    using Suc_diff_le not_less_eq_eq by blast
-  from Suc(3) \<open>r \<noteq> []\<close> have n_ok : "n = 1 + (length ?rrr) - length d"
-    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 \<open>r \<noteq> []\<close>] 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
-      by (simp add: monom_Suc hd_rev[symmetric] 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 * \<open>d \<noteq> []\<close>, auto)
-      from * 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 g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))"
-    for f g :: "'a::comm_ring_1 poly"
-proof (cases "g = 0")
-  case False
-  then have "last (coeffs g) \<noteq> 0"
-    and "last (coeffs g) = lead_coeff g"
-    and "coeffs g \<noteq> []"
-    by (simp_all add: last_coeffs_eq_coeff_degree)
-  moreover 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
-  ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g
-    (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))"
-    by (subst pseudo_divmod_main_list_invar [symmetric]) auto
-  moreover have "pseudo_divmod_main_list
-    (hd (rev (coeffs g))) []
-    (rev (coeffs f)) (rev (coeffs g))
-    (1 + length (coeffs f) -
-    length (coeffs g)) = (q, r)"
-    using qr hd_rev [OF \<open>coeffs g \<noteq> []\<close>] by simp
-  ultimately show ?thesis
-    by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def)
-next
-  case True
-  then show ?thesis
-    by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def)
-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> (p div q, p mod q) = (r, s)"
-  by (metis eucl_rel_poly eucl_rel_poly_unique)
-
-lemma pdivmod_via_pseudo_divmod:
-  "(f div g, f mod 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 True
-  then show ?thesis by simp
-next
-  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"
-    by (auto simp: h_def lc_def)
-  then have 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)"
-    by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p)
-  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
-  from f r h0 have "eucl_rel_poly f h (q, r)"
-    by (auto simp: eucl_rel_poly_iff)
-  then have "(f div h, f mod h) = (q, r)"
-    by (simp add: pdivmod_pdivmodrel)
-  with lc have "(f div g, f mod g) = (smult lc q, r)"
-    by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc])
-  with id show ?thesis
-    by auto
-qed
-
-lemma pdivmod_via_pseudo_divmod_list:
-  "(f div g, f mod 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
-    with d show ?thesis by auto
-  next
-    case False
-    define ilc where "ilc = inverse (coeff g (degree g))"
-    from False have ilc: "ilc \<noteq> 0"
-      by (auto simp: ilc_def)
-    with False have id: "g = 0 \<longleftrightarrow> False" "coeffs g = [] \<longleftrightarrow> False"
-      "last (coeffs g) = coeff g (degree g)"
-      "coeffs (smult ilc g) = [] \<longleftrightarrow> 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)
-  have *: "map (op * 1) xs = xs" for xs :: "'a list"
-    by (induct xs) auto
-  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 = 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 "_ = ?rhs")
-proof -
-  have "snd (f div g, f mod g) = ?rhs"
-    unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil]
-    by (auto split: prod.splits)
-  then show ?thesis by simp
-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 "fst (f div g, f mod g) = div_field_poly_impl f g"
-    unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def
-    by (auto split: prod.splits)
-  then show "f div g =  div_field_poly_impl f g"
-    by simp
-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)
-  from Suc.prems have ifCond: "\<not> Suc (length r) \<le> length d"
-    by simp
-  with d have r: "r \<noteq> []"
-    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
-    with r d show ?thesis
-      unfolding Suc(2)[symmetric]
-      by (auto simp add: Let_def nth_default_append)
-  next
-    case True
-    with r d have id:
-      "?thesis \<longleftrightarrow>
-        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"
-      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)
-      then show ?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 by (auto simp: d True)
-  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 "lcg \<noteq> 0" by auto
-    from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = 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 \<open>lcg \<noteq> 0\<close>)
-        (auto simp: lcg ltp, simp add: degree_eq_length_coeffs)
-  qed
-qed
-
-no_notation cCons (infixr "##" 65)
-
-end
--- a/src/HOL/Library/Polynomial_FPS.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,313 +0,0 @@
-(*  Title:      HOL/Library/Polynomial_FPS.thy
-    Author:     Manuel Eberl, TU München
-  
-Converting polynomials to formal power series.
-*)
-
-section \<open>Converting polynomials to formal power series\<close>
-
-theory Polynomial_FPS
-  imports Polynomial Formal_Power_Series
-begin
-
-definition fps_of_poly where
-  "fps_of_poly p = Abs_fps (coeff p)"
-
-lemma fps_of_poly_eq_iff: "fps_of_poly p = fps_of_poly q \<longleftrightarrow> p = q"
-  by (simp add: fps_of_poly_def poly_eq_iff fps_eq_iff)
-
-lemma fps_of_poly_nth [simp]: "fps_of_poly p $ n = coeff p n"
-  by (simp add: fps_of_poly_def)
-  
-lemma fps_of_poly_const: "fps_of_poly [:c:] = fps_const c"
-proof (subst fps_eq_iff, clarify)
-  fix n :: nat show "fps_of_poly [:c:] $ n = fps_const c $ n"
-    by (cases n) (auto simp: fps_of_poly_def)
-qed
-
-lemma fps_of_poly_0 [simp]: "fps_of_poly 0 = 0"
-  by (subst fps_const_0_eq_0 [symmetric], subst fps_of_poly_const [symmetric]) simp
-
-lemma fps_of_poly_1 [simp]: "fps_of_poly 1 = 1"
-  by (subst fps_const_1_eq_1 [symmetric], subst fps_of_poly_const [symmetric])
-     (simp add: one_poly_def)
-
-lemma fps_of_poly_1' [simp]: "fps_of_poly [:1:] = 1"
-  by (subst fps_const_1_eq_1 [symmetric], subst fps_of_poly_const [symmetric])
-     (simp add: one_poly_def)
-
-lemma fps_of_poly_numeral [simp]: "fps_of_poly (numeral n) = numeral n"
-  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
-
-lemma fps_of_poly_numeral' [simp]: "fps_of_poly [:numeral n:] = numeral n"
-  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
-
-lemma fps_of_poly_X [simp]: "fps_of_poly [:0, 1:] = X"
-  by (auto simp add: fps_of_poly_def fps_eq_iff coeff_pCons split: nat.split)
-
-lemma fps_of_poly_add: "fps_of_poly (p + q) = fps_of_poly p + fps_of_poly q"
-  by (simp add: fps_of_poly_def plus_poly.rep_eq fps_plus_def)
-
-lemma fps_of_poly_diff: "fps_of_poly (p - q) = fps_of_poly p - fps_of_poly q"
-  by (simp add: fps_of_poly_def minus_poly.rep_eq fps_minus_def)
-
-lemma fps_of_poly_uminus: "fps_of_poly (-p) = -fps_of_poly p"
-  by (simp add: fps_of_poly_def uminus_poly.rep_eq fps_uminus_def)
-
-lemma fps_of_poly_mult: "fps_of_poly (p * q) = fps_of_poly p * fps_of_poly q"
-  by (simp add: fps_of_poly_def fps_times_def fps_eq_iff coeff_mult atLeast0AtMost)
-
-lemma fps_of_poly_smult: 
-  "fps_of_poly (smult c p) = fps_const c * fps_of_poly p"
-  using fps_of_poly_mult[of "[:c:]" p] by (simp add: fps_of_poly_mult fps_of_poly_const)
-  
-lemma fps_of_poly_sum: "fps_of_poly (sum f A) = sum (\<lambda>x. fps_of_poly (f x)) A"
-  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_add)
-
-lemma fps_of_poly_sum_list: "fps_of_poly (sum_list xs) = sum_list (map fps_of_poly xs)"
-  by (induction xs) (simp_all add: fps_of_poly_add)
-  
-lemma fps_of_poly_prod: "fps_of_poly (prod f A) = prod (\<lambda>x. fps_of_poly (f x)) A"
-  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_mult)
-  
-lemma fps_of_poly_prod_list: "fps_of_poly (prod_list xs) = prod_list (map fps_of_poly xs)"
-  by (induction xs) (simp_all add: fps_of_poly_mult)
-
-lemma fps_of_poly_pCons: 
-  "fps_of_poly (pCons (c :: 'a :: semiring_1) p) = fps_const c + fps_of_poly p * X"
-  by (subst fps_mult_X_commute [symmetric], intro fps_ext) 
-     (auto simp: fps_of_poly_def coeff_pCons split: nat.split)
-  
-lemma fps_of_poly_pderiv: "fps_of_poly (pderiv p) = fps_deriv (fps_of_poly p)"
-  by (intro fps_ext) (simp add: fps_of_poly_nth coeff_pderiv)
-
-lemma fps_of_poly_power: "fps_of_poly (p ^ n) = fps_of_poly p ^ n"
-  by (induction n) (simp_all add: fps_of_poly_mult)
-  
-lemma fps_of_poly_monom: "fps_of_poly (monom (c :: 'a :: comm_ring_1) n) = fps_const c * X ^ n"
-  by (intro fps_ext) simp_all
-
-lemma fps_of_poly_monom': "fps_of_poly (monom (1 :: 'a :: comm_ring_1) n) = X ^ n"
-  by (simp add: fps_of_poly_monom)
-
-lemma fps_of_poly_div:
-  assumes "(q :: 'a :: field poly) dvd p"
-  shows   "fps_of_poly (p div q) = fps_of_poly p / fps_of_poly q"
-proof (cases "q = 0")
-  case False
-  from False fps_of_poly_eq_iff[of q 0] have nz: "fps_of_poly q \<noteq> 0" by simp 
-  from assms have "p = (p div q) * q" by simp
-  also have "fps_of_poly \<dots> = fps_of_poly (p div q) * fps_of_poly q" 
-    by (simp add: fps_of_poly_mult)
-  also from nz have "\<dots> / fps_of_poly q = fps_of_poly (p div q)"
-    by (intro nonzero_mult_div_cancel_right) (auto simp: fps_of_poly_0)
-  finally show ?thesis ..
-qed simp
-
-lemma fps_of_poly_divide_numeral:
-  "fps_of_poly (smult (inverse (numeral c :: 'a :: field)) p) = fps_of_poly p / numeral c"
-proof -
-  have "smult (inverse (numeral c)) p = [:inverse (numeral c):] * p" by simp
-  also have "fps_of_poly \<dots> = fps_of_poly p / numeral c"
-    by (subst fps_of_poly_mult) (simp add: numeral_fps_const fps_of_poly_pCons)
-  finally show ?thesis by simp
-qed
-
-
-lemma subdegree_fps_of_poly:
-  assumes "p \<noteq> 0"
-  defines "n \<equiv> Polynomial.order 0 p"
-  shows   "subdegree (fps_of_poly p) = n"
-proof (rule subdegreeI)
-  from assms have "monom 1 n dvd p" by (simp add: monom_1_dvd_iff)
-  thus zero: "fps_of_poly p $ i = 0" if "i < n" for i
-    using that by (simp add: monom_1_dvd_iff')
-    
-  from assms have "\<not>monom 1 (Suc n) dvd p"
-    by (auto simp: monom_1_dvd_iff simp del: power_Suc)
-  then obtain k where k: "k \<le> n" "fps_of_poly p $ k \<noteq> 0" 
-    by (auto simp: monom_1_dvd_iff' less_Suc_eq_le)
-  with zero[of k] have "k = n" by linarith
-  with k show "fps_of_poly p $ n \<noteq> 0" by simp
-qed
-
-lemma fps_of_poly_dvd:
-  assumes "p dvd q"
-  shows   "fps_of_poly (p :: 'a :: field poly) dvd fps_of_poly q"
-proof (cases "p = 0 \<or> q = 0")
-  case False
-  with assms fps_of_poly_eq_iff[of p 0] fps_of_poly_eq_iff[of q 0] show ?thesis
-    by (auto simp: fps_dvd_iff subdegree_fps_of_poly dvd_imp_order_le)
-qed (insert assms, auto)
-
-
-lemmas fps_of_poly_simps =
-  fps_of_poly_0 fps_of_poly_1 fps_of_poly_numeral fps_of_poly_const fps_of_poly_X
-  fps_of_poly_add fps_of_poly_diff fps_of_poly_uminus fps_of_poly_mult fps_of_poly_smult
-  fps_of_poly_sum fps_of_poly_sum_list fps_of_poly_prod fps_of_poly_prod_list
-  fps_of_poly_pCons fps_of_poly_pderiv fps_of_poly_power fps_of_poly_monom
-  fps_of_poly_divide_numeral
-
-lemma fps_of_poly_pcompose:
-  assumes "coeff q 0 = (0 :: 'a :: idom)"
-  shows   "fps_of_poly (pcompose p q) = fps_compose (fps_of_poly p) (fps_of_poly q)"
-  using assms by (induction p rule: pCons_induct)
-                 (auto simp: pcompose_pCons fps_of_poly_simps fps_of_poly_pCons 
-                             fps_compose_add_distrib fps_compose_mult_distrib)
-  
-lemmas reify_fps_atom =
-  fps_of_poly_0 fps_of_poly_1' fps_of_poly_numeral' fps_of_poly_const fps_of_poly_X
-
-
-text \<open>
-  The following simproc can reduce the equality of two polynomial FPSs two equality of the
-  respective polynomials. A polynomial FPS is one that only has finitely many non-zero 
-  coefficients and can therefore be written as @{term "fps_of_poly p"} for some 
-  polynomial \<open>p\<close>.
-  
-  This may sound trivial, but it covers a number of annoying side conditions like 
-  @{term "1 + X \<noteq> 0"} that would otherwise not be solved automatically.
-\<close>
-
-ML \<open>
-
-(* TODO: Support for division *)
-signature POLY_FPS = sig
-
-val reify_conv : conv
-val eq_conv : conv
-val eq_simproc : cterm -> thm option
-
-end
-
-
-structure Poly_Fps = struct
-
-fun const_binop_conv s conv ct =
-  case Thm.term_of ct of
-    (Const (s', _) $ _ $ _) => 
-      if s = s' then 
-        Conv.binop_conv conv ct 
-      else 
-        raise CTERM ("const_binop_conv", [ct])
-  | _ => raise CTERM ("const_binop_conv", [ct])
-
-fun reify_conv ct = 
-  let
-    val rewr = Conv.rewrs_conv o map (fn thm => thm RS @{thm eq_reflection})
-    val un = Conv.arg_conv reify_conv
-    val bin = Conv.binop_conv reify_conv
-  in
-    case Thm.term_of ct of
-      (Const (@{const_name "fps_of_poly"}, _) $ _) => ct |> Conv.all_conv
-    | (Const (@{const_name "Groups.plus"}, _) $ _ $ _) => ct |> (
-        bin then_conv rewr @{thms fps_of_poly_add [symmetric]})
-    | (Const (@{const_name "Groups.uminus"}, _) $ _) => ct |> (
-        un then_conv rewr @{thms fps_of_poly_uminus [symmetric]})
-    | (Const (@{const_name "Groups.minus"}, _) $ _ $ _) => ct |> (
-        bin then_conv rewr @{thms fps_of_poly_diff [symmetric]})
-    | (Const (@{const_name "Groups.times"}, _) $ _ $ _) => ct |> (
-        bin then_conv rewr @{thms fps_of_poly_mult [symmetric]})
-    | (Const (@{const_name "Rings.divide"}, _) $ _ $ (Const (@{const_name "Num.numeral"}, _) $ _))
-        => ct |> (Conv.fun_conv (Conv.arg_conv reify_conv)
-             then_conv rewr @{thms fps_of_poly_divide_numeral [symmetric]})
-    | (Const (@{const_name "Power.power"}, _) $ Const (@{const_name "X"},_) $ _) => ct |> (
-        rewr @{thms fps_of_poly_monom' [symmetric]}) 
-    | (Const (@{const_name "Power.power"}, _) $ _ $ _) => ct |> (
-        Conv.fun_conv (Conv.arg_conv reify_conv) 
-        then_conv rewr @{thms fps_of_poly_power [symmetric]})
-    | _ => ct |> (
-        rewr @{thms reify_fps_atom [symmetric]})
-  end
-    
-
-fun eq_conv ct =
-  case Thm.term_of ct of
-    (Const (@{const_name "HOL.eq"}, _) $ _ $ _) => ct |> (
-      Conv.binop_conv reify_conv
-      then_conv Conv.rewr_conv @{thm fps_of_poly_eq_iff[THEN eq_reflection]})
-  | _ => raise CTERM ("poly_fps_eq_conv", [ct])
-
-val eq_simproc = try eq_conv
-
-end
-\<close> 
-
-simproc_setup poly_fps_eq ("(f :: 'a fps) = g") = \<open>K (K Poly_Fps.eq_simproc)\<close>
-
-lemma fps_of_poly_linear: "fps_of_poly [:a,1 :: 'a :: field:] = X + fps_const a"
-  by simp
-
-lemma fps_of_poly_linear': "fps_of_poly [:1,a :: 'a :: field:] = 1 + fps_const a * X"
-  by simp
-
-lemma fps_of_poly_cutoff [simp]: 
-  "fps_of_poly (poly_cutoff n p) = fps_cutoff n (fps_of_poly p)"
-  by (simp add: fps_eq_iff coeff_poly_cutoff)
-
-lemma fps_of_poly_shift [simp]: "fps_of_poly (poly_shift n p) = fps_shift n (fps_of_poly p)"
-  by (simp add: fps_eq_iff coeff_poly_shift)
-
-
-definition poly_subdegree :: "'a::zero poly \<Rightarrow> nat" where
-  "poly_subdegree p = subdegree (fps_of_poly p)"
-
-lemma coeff_less_poly_subdegree:
-  "k < poly_subdegree p \<Longrightarrow> coeff p k = 0"
-  unfolding poly_subdegree_def using nth_less_subdegree_zero[of k "fps_of_poly p"] by simp
-
-(* TODO: Move ? *)
-definition prefix_length :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "prefix_length P xs = length (takeWhile P xs)"
-
-primrec prefix_length_aux :: "('a \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "prefix_length_aux P acc [] = acc"
-| "prefix_length_aux P acc (x#xs) = (if P x then prefix_length_aux P (Suc acc) xs else acc)"
-
-lemma prefix_length_aux_correct: "prefix_length_aux P acc xs = prefix_length P xs + acc"
-  by (induction xs arbitrary: acc) (simp_all add: prefix_length_def)
-
-lemma prefix_length_code [code]: "prefix_length P xs = prefix_length_aux P 0 xs"
-  by (simp add: prefix_length_aux_correct)
-
-lemma prefix_length_le_length: "prefix_length P xs \<le> length xs"
-  by (induction xs) (simp_all add: prefix_length_def)
-  
-lemma prefix_length_less_length: "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> prefix_length P xs < length xs"
-  by (induction xs) (simp_all add: prefix_length_def)
-
-lemma nth_prefix_length:
-  "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> \<not>P (xs ! prefix_length P xs)"
-  by (induction xs) (simp_all add: prefix_length_def)
-  
-lemma nth_less_prefix_length:
-  "n < prefix_length P xs \<Longrightarrow> P (xs ! n)"
-  by (induction xs arbitrary: n) 
-     (auto simp: prefix_length_def nth_Cons split: if_splits nat.splits)
-(* END TODO *)
-  
-lemma poly_subdegree_code [code]: "poly_subdegree p = prefix_length (op = 0) (coeffs p)"
-proof (cases "p = 0")
-  case False
-  note [simp] = this
-  define n where "n = prefix_length (op = 0) (coeffs p)"
-  from False have "\<exists>k. coeff p k \<noteq> 0" by (auto simp: poly_eq_iff)
-  hence ex: "\<exists>x\<in>set (coeffs p). x \<noteq> 0" by (auto simp: coeffs_def)
-  hence n_less: "n < length (coeffs p)" and nonzero: "coeffs p ! n \<noteq> 0" 
-    unfolding n_def by (auto intro!: prefix_length_less_length nth_prefix_length)
-  show ?thesis unfolding poly_subdegree_def
-  proof (intro subdegreeI)
-    from n_less have "fps_of_poly p $ n = coeffs p ! n"
-      by (subst coeffs_nth) (simp_all add: degree_eq_length_coeffs)
-    with nonzero show "fps_of_poly p $ prefix_length (op = 0) (coeffs p) \<noteq> 0"
-      unfolding n_def by simp
-  next
-    fix k assume A: "k < prefix_length (op = 0) (coeffs p)"
-    also have "\<dots> \<le> length (coeffs p)" by (rule prefix_length_le_length)
-    finally show "fps_of_poly p $ k = 0"
-      using nth_less_prefix_length[OF A]
-      by (simp add: coeffs_nth degree_eq_length_coeffs)
-  qed
-qed (simp_all add: poly_subdegree_def prefix_length_def)
-
-end
--- a/src/HOL/Library/Polynomial_Factorial.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,992 +0,0 @@
-(*  Title:      HOL/Library/Polynomial_Factorial.thy
-    Author:     Brian Huffman
-    Author:     Clemens Ballarin
-    Author:     Amine Chaieb
-    Author:     Florian Haftmann
-    Author:     Manuel Eberl
-*)
-
-theory Polynomial_Factorial
-imports 
-  Complex_Main
-  Polynomial
-  Normalized_Fraction
-  Field_as_Ring
-begin
-
-subsection \<open>Various facts about polynomials\<close>
-
-lemma prod_mset_const_poly: " (\<Prod>x\<in>#A. [:f x:]) = [:prod_mset (image_mset f A):]"
-  by (induct A) (simp_all add: one_poly_def ac_simps)
-
-lemma irreducible_const_poly_iff:
-  fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
-  shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
-proof
-  assume A: "irreducible c"
-  show "irreducible [:c:]"
-  proof (rule irreducibleI)
-    fix a b assume ab: "[:c:] = a * b"
-    hence "degree [:c:] = degree (a * b)" by (simp only: )
-    also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
-    hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
-    finally have "degree a = 0" "degree b = 0" by auto
-    then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
-    from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
-    hence "c = a' * b'" by (simp add: ab' mult_ac)
-    from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
-    with ab' show "a dvd 1 \<or> b dvd 1" by (auto simp: one_poly_def)
-  qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
-next
-  assume A: "irreducible [:c:]"
-  show "irreducible c"
-  proof (rule irreducibleI)
-    fix a b assume ab: "c = a * b"
-    hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
-    from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
-    thus "a dvd 1 \<or> b dvd 1" by (simp add: one_poly_def)
-  qed (insert A, auto simp: irreducible_def one_poly_def)
-qed
-
-
-subsection \<open>Lifting elements into the field of fractions\<close>
-
-definition to_fract :: "'a :: idom \<Rightarrow> 'a fract" where "to_fract x = Fract x 1"
-  \<comment> \<open>FIXME: name \<open>of_idom\<close>, abbreviation\<close>
-
-lemma to_fract_0 [simp]: "to_fract 0 = 0"
-  by (simp add: to_fract_def eq_fract Zero_fract_def)
-
-lemma to_fract_1 [simp]: "to_fract 1 = 1"
-  by (simp add: to_fract_def eq_fract One_fract_def)
-
-lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y"
-  by (simp add: to_fract_def)
-
-lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y"
-  by (simp add: to_fract_def)
-  
-lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x"
-  by (simp add: to_fract_def)
-  
-lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y"
-  by (simp add: to_fract_def)
-
-lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \<longleftrightarrow> x = y"
-  by (simp add: to_fract_def eq_fract)
-  
-lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \<longleftrightarrow> x = 0"
-  by (simp add: to_fract_def Zero_fract_def eq_fract)
-
-lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \<noteq> 0"
-  by transfer simp
-
-lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x"
-  by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp)
-
-lemma to_fract_quot_of_fract:
-  assumes "snd (quot_of_fract x) = 1"
-  shows   "to_fract (fst (quot_of_fract x)) = x"
-proof -
-  have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp
-  also note assms
-  finally show ?thesis by (simp add: to_fract_def)
-qed
-
-lemma snd_quot_of_fract_Fract_whole:
-  assumes "y dvd x"
-  shows   "snd (quot_of_fract (Fract x y)) = 1"
-  using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd)
-  
-lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b"
-  by (simp add: to_fract_def)
-
-lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)"
-  unfolding to_fract_def by transfer (simp add: normalize_quot_def)
-
-lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \<longleftrightarrow> x = 0"
-  by transfer simp
- 
-lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1"
-  unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all
-
-lemma coprime_quot_of_fract:
-  "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))"
-  by transfer (simp add: coprime_normalize_quot)
-
-lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1"
-  using quot_of_fract_in_normalized_fracts[of x] 
-  by (simp add: normalized_fracts_def case_prod_unfold)  
-
-lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \<Longrightarrow> normalize x = x"
-  by (subst (2) normalize_mult_unit_factor [symmetric, of x])
-     (simp del: normalize_mult_unit_factor)
-  
-lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)"
-  by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract)
-
-
-subsection \<open>Lifting polynomial coefficients to the field of fractions\<close>
-
-abbreviation (input) fract_poly 
-  where "fract_poly \<equiv> map_poly to_fract"
-
-abbreviation (input) unfract_poly 
-  where "unfract_poly \<equiv> map_poly (fst \<circ> quot_of_fract)"
-  
-lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)"
-  by (simp add: smult_conv_map_poly map_poly_map_poly o_def)
-
-lemma fract_poly_0 [simp]: "fract_poly 0 = 0"
-  by (simp add: poly_eqI coeff_map_poly)
-
-lemma fract_poly_1 [simp]: "fract_poly 1 = 1"
-  by (simp add: one_poly_def map_poly_pCons)
-
-lemma fract_poly_add [simp]:
-  "fract_poly (p + q) = fract_poly p + fract_poly q"
-  by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-lemma fract_poly_diff [simp]:
-  "fract_poly (p - q) = fract_poly p - fract_poly q"
-  by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\<lambda>x. to_fract (f x)) A"
-  by (cases "finite A", induction A rule: finite_induct) simp_all 
-
-lemma fract_poly_mult [simp]:
-  "fract_poly (p * q) = fract_poly p * fract_poly q"
-  by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult)
-
-lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \<longleftrightarrow> p = q"
-  by (auto simp: poly_eq_iff coeff_map_poly)
-
-lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \<longleftrightarrow> p = 0"
-  using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff)
-
-lemma fract_poly_dvd: "p dvd q \<Longrightarrow> fract_poly p dvd fract_poly q"
-  by (auto elim!: dvdE)
-
-lemma prod_mset_fract_poly: 
-  "(\<Prod>x\<in>#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))"
-  by (induct A) (simp_all add: ac_simps)
-  
-lemma is_unit_fract_poly_iff:
-  "p dvd 1 \<longleftrightarrow> fract_poly p dvd 1 \<and> content p = 1"
-proof safe
-  assume A: "p dvd 1"
-  with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)"
-    by simp
-  from A show "content p = 1"
-    by (auto simp: is_unit_poly_iff normalize_1_iff)
-next
-  assume A: "fract_poly p dvd 1" and B: "content p = 1"
-  from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff)
-  {
-    fix n :: nat assume "n > 0"
-    have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly)
-    also note c
-    also from \<open>n > 0\<close> have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits)
-    finally have "coeff p n = 0" by simp
-  }
-  hence "degree p \<le> 0" by (intro degree_le) simp_all
-  with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE)
-qed
-  
-lemma fract_poly_is_unit: "p dvd 1 \<Longrightarrow> fract_poly p dvd 1"
-  using fract_poly_dvd[of p 1] by simp
-
-lemma fract_poly_smult_eqE:
-  fixes c :: "'a :: {idom_divide,ring_gcd} fract"
-  assumes "fract_poly p = smult c (fract_poly q)"
-  obtains a b 
-    where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a"
-proof -
-  define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)"
-  have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)"
-    by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms)
-  hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff)
-  hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff)
-  moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b"
-    by (simp_all add: a_def b_def coprime_quot_of_fract gcd.commute
-          normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric])
-  ultimately show ?thesis by (intro that[of a b])
-qed
-
-
-subsection \<open>Fractional content\<close>
-
-abbreviation (input) Lcm_coeff_denoms 
-    :: "'a :: {semiring_Gcd,idom_divide,ring_gcd} fract poly \<Rightarrow> 'a"
-  where "Lcm_coeff_denoms p \<equiv> Lcm (snd ` quot_of_fract ` set (coeffs p))"
-  
-definition fract_content :: 
-      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a fract" where
-  "fract_content p = 
-     (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" 
-
-definition primitive_part_fract :: 
-      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a poly" where
-  "primitive_part_fract p = 
-     primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))"
-
-lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0"
-  by (simp add: primitive_part_fract_def)
-
-lemma fract_content_eq_0_iff [simp]:
-  "fract_content p = 0 \<longleftrightarrow> p = 0"
-  unfolding fract_content_def Let_def Zero_fract_def
-  by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff)
-
-lemma content_primitive_part_fract [simp]: "p \<noteq> 0 \<Longrightarrow> content (primitive_part_fract p) = 1"
-  unfolding primitive_part_fract_def
-  by (rule content_primitive_part)
-     (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff)  
-
-lemma content_times_primitive_part_fract:
-  "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p"
-proof -
-  define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)"
-  have "fract_poly p' = 
-          map_poly (to_fract \<circ> fst \<circ> quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)"
-    unfolding primitive_part_fract_def p'_def 
-    by (subst map_poly_map_poly) (simp_all add: o_assoc)
-  also have "\<dots> = smult (to_fract (Lcm_coeff_denoms p)) p"
-  proof (intro map_poly_idI, unfold o_apply)
-    fix c assume "c \<in> set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))"
-    then obtain c' where c: "c' \<in> set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'"
-      by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits)
-    note c(2)
-    also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))"
-      by simp
-    also have "to_fract (Lcm_coeff_denoms p) * \<dots> = 
-                 Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))"
-      unfolding to_fract_def by (subst mult_fract) simp_all
-    also have "snd (quot_of_fract \<dots>) = 1"
-      by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto)
-    finally show "to_fract (fst (quot_of_fract c)) = c"
-      by (rule to_fract_quot_of_fract)
-  qed
-  also have "p' = smult (content p') (primitive_part p')" 
-    by (rule content_times_primitive_part [symmetric])
-  also have "primitive_part p' = primitive_part_fract p"
-    by (simp add: primitive_part_fract_def p'_def)
-  also have "fract_poly (smult (content p') (primitive_part_fract p)) = 
-               smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp
-  finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) =
-                      smult (to_fract (Lcm_coeff_denoms p)) p" .
-  thus ?thesis
-    by (subst (asm) smult_eq_iff)
-       (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def)
-qed
-
-lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)"
-proof -
-  have "Lcm_coeff_denoms (fract_poly p) = 1"
-    by (auto simp: set_coeffs_map_poly)
-  hence "fract_content (fract_poly p) = 
-           to_fract (content (map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p))"
-    by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff)
-  also have "map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p = p"
-    by (intro map_poly_idI) simp_all
-  finally show ?thesis .
-qed
-
-lemma content_decompose_fract:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly"
-  obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1"
-proof (cases "p = 0")
-  case True
-  hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all
-  thus ?thesis ..
-next
-  case False
-  thus ?thesis
-    by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract])
-qed
-
-
-subsection \<open>More properties of content and primitive part\<close>
-
-lemma lift_prime_elem_poly:
-  assumes "prime_elem (c :: 'a :: semidom)"
-  shows   "prime_elem [:c:]"
-proof (rule prime_elemI)
-  fix a b assume *: "[:c:] dvd a * b"
-  from * have dvd: "c dvd coeff (a * b) n" for n
-    by (subst (asm) const_poly_dvd_iff) blast
-  {
-    define m where "m = (GREATEST m. \<not>c dvd coeff b m)"
-    assume "\<not>[:c:] dvd b"
-    hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast
-    have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i < Suc (degree b)"
-      by (auto intro: le_degree simp: less_Suc_eq_le)
-    have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex[OF A B])
-    have "i \<le> m" if "\<not>c dvd coeff b i" for i
-      unfolding m_def by (rule Greatest_le[OF that B])
-    hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force
-
-    have "c dvd coeff a i" for i
-    proof (induction i rule: nat_descend_induct[of "degree a"])
-      case (base i)
-      thus ?case by (simp add: coeff_eq_0)
-    next
-      case (descend i)
-      let ?A = "{..i+m} - {i}"
-      have "c dvd coeff (a * b) (i + m)" by (rule dvd)
-      also have "coeff (a * b) (i + m) = (\<Sum>k\<le>i + m. coeff a k * coeff b (i + m - k))"
-        by (simp add: coeff_mult)
-      also have "{..i+m} = insert i ?A" by auto
-      also have "(\<Sum>k\<in>\<dots>. coeff a k * coeff b (i + m - k)) =
-                   coeff a i * coeff b m + (\<Sum>k\<in>?A. coeff a k * coeff b (i + m - k))"
-        (is "_ = _ + ?S")
-        by (subst sum.insert) simp_all
-      finally have eq: "c dvd coeff a i * coeff b m + ?S" .
-      moreover have "c dvd ?S"
-      proof (rule dvd_sum)
-        fix k assume k: "k \<in> {..i+m} - {i}"
-        show "c dvd coeff a k * coeff b (i + m - k)"
-        proof (cases "k < i")
-          case False
-          with k have "c dvd coeff a k" by (intro descend.IH) simp
-          thus ?thesis by simp
-        next
-          case True
-          hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp
-          thus ?thesis by simp
-        qed
-      qed
-      ultimately have "c dvd coeff a i * coeff b m"
-        by (simp add: dvd_add_left_iff)
-      with assms coeff_m show "c dvd coeff a i"
-        by (simp add: prime_elem_dvd_mult_iff)
-    qed
-    hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast
-  }
-  thus "[:c:] dvd a \<or> [:c:] dvd b" by blast
-qed (insert assms, simp_all add: prime_elem_def one_poly_def)
-
-lemma prime_elem_const_poly_iff:
-  fixes c :: "'a :: semidom"
-  shows   "prime_elem [:c:] \<longleftrightarrow> prime_elem c"
-proof
-  assume A: "prime_elem [:c:]"
-  show "prime_elem c"
-  proof (rule prime_elemI)
-    fix a b assume "c dvd a * b"
-    hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac)
-    from A and this have "[:c:] dvd [:a:] \<or> [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD)
-    thus "c dvd a \<or> c dvd b" by simp
-  qed (insert A, auto simp: prime_elem_def is_unit_poly_iff)
-qed (auto intro: lift_prime_elem_poly)
-
-context
-begin
-
-private lemma content_1_mult:
-  fixes f g :: "'a :: {semiring_Gcd,factorial_semiring} poly"
-  assumes "content f = 1" "content g = 1"
-  shows   "content (f * g) = 1"
-proof (cases "f * g = 0")
-  case False
-  from assms have "f \<noteq> 0" "g \<noteq> 0" by auto
-
-  hence "f * g \<noteq> 0" by auto
-  {
-    assume "\<not>is_unit (content (f * g))"
-    with False have "\<exists>p. p dvd content (f * g) \<and> prime p"
-      by (intro prime_divisor_exists) simp_all
-    then obtain p where "p dvd content (f * g)" "prime p" by blast
-    from \<open>p dvd content (f * g)\<close> have "[:p:] dvd f * g"
-      by (simp add: const_poly_dvd_iff_dvd_content)
-    moreover from \<open>prime p\<close> have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly)
-    ultimately have "[:p:] dvd f \<or> [:p:] dvd g"
-      by (simp add: prime_elem_dvd_mult_iff)
-    with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content)
-    with \<open>prime p\<close> have False by simp
-  }
-  hence "is_unit (content (f * g))" by blast
-  hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content)
-  thus ?thesis by simp
-qed (insert assms, auto)
-
-lemma content_mult:
-  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
-  shows "content (p * q) = content p * content q"
-proof -
-  from content_decompose[of p] guess p' . note p = this
-  from content_decompose[of q] guess q' . note q = this
-  have "content (p * q) = content p * content q * content (p' * q')"
-    by (subst p, subst q) (simp add: mult_ac normalize_mult)
-  also from p q have "content (p' * q') = 1" by (intro content_1_mult)
-  finally show ?thesis by simp
-qed
-
-lemma primitive_part_mult:
-  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
-  shows "primitive_part (p * q) = primitive_part p * primitive_part q"
-proof -
-  have "primitive_part (p * q) = p * q div [:content (p * q):]"
-    by (simp add: primitive_part_def div_const_poly_conv_map_poly)
-  also have "\<dots> = (p div [:content p:]) * (q div [:content q:])"
-    by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac)
-  also have "\<dots> = primitive_part p * primitive_part q"
-    by (simp add: primitive_part_def div_const_poly_conv_map_poly)
-  finally show ?thesis .
-qed
-
-lemma primitive_part_smult:
-  fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
-  shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)"
-proof -
-  have "smult a p = [:a:] * p" by simp
-  also have "primitive_part \<dots> = smult (unit_factor a) (primitive_part p)"
-    by (subst primitive_part_mult) simp_all
-  finally show ?thesis .
-qed  
-
-lemma primitive_part_dvd_primitive_partI [intro]:
-  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
-  shows "p dvd q \<Longrightarrow> primitive_part p dvd primitive_part q"
-  by (auto elim!: dvdE simp: primitive_part_mult)
-
-lemma content_prod_mset: 
-  fixes A :: "'a :: {factorial_semiring, semiring_Gcd} poly multiset"
-  shows "content (prod_mset A) = prod_mset (image_mset content A)"
-  by (induction A) (simp_all add: content_mult mult_ac)
-
-lemma fract_poly_dvdD:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
-  assumes "fract_poly p dvd fract_poly q" "content p = 1"
-  shows   "p dvd q"
-proof -
-  from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE)
-  from content_decompose_fract[of r] guess c r' . note r' = this
-  from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp  
-  from fract_poly_smult_eqE[OF this] guess a b . note ab = this
-  have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2))
-  hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4))
-  have "1 = gcd a (normalize b)" by (simp add: ab)
-  also note eq'
-  also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4))
-  finally have [simp]: "a = 1" by simp
-  from eq ab have "q = p * ([:b:] * r')" by simp
-  thus ?thesis by (rule dvdI)
-qed
-
-lemma content_prod_eq_1_iff: 
-  fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
-  shows "content (p * q) = 1 \<longleftrightarrow> content p = 1 \<and> content q = 1"
-proof safe
-  assume A: "content (p * q) = 1"
-  {
-    fix p q :: "'a poly" assume "content p * content q = 1"
-    hence "1 = content p * content q" by simp
-    hence "content p dvd 1" by (rule dvdI)
-    hence "content p = 1" by simp
-  } note B = this
-  from A B[of p q] B [of q p] show "content p = 1" "content q = 1" 
-    by (simp_all add: content_mult mult_ac)
-qed (auto simp: content_mult)
-
-end
-
-
-subsection \<open>Polynomials over a field are a Euclidean ring\<close>
-
-definition unit_factor_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
-  "unit_factor_field_poly p = [:lead_coeff p:]"
-
-definition normalize_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
-  "normalize_field_poly p = smult (inverse (lead_coeff p)) p"
-
-definition euclidean_size_field_poly :: "'a :: field poly \<Rightarrow> nat" where
-  "euclidean_size_field_poly p = (if p = 0 then 0 else 2 ^ degree p)" 
-
-lemma dvd_field_poly: "dvd.dvd (op * :: 'a :: field poly \<Rightarrow> _) = op dvd"
-  by (intro ext) (simp_all add: dvd.dvd_def dvd_def)
-
-interpretation field_poly: 
-  unique_euclidean_ring where zero = "0 :: 'a :: field poly"
-    and one = 1 and plus = plus and uminus = uminus and minus = minus
-    and times = times
-    and normalize = normalize_field_poly and unit_factor = unit_factor_field_poly
-    and euclidean_size = euclidean_size_field_poly
-    and uniqueness_constraint = top
-    and divide = divide and modulo = modulo
-proof (standard, unfold dvd_field_poly)
-  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)
-next
-  fix p :: "'a poly" assume "is_unit p"
-  then show "unit_factor_field_poly p = p"
-    by (elim is_unit_polyE) (auto simp: unit_factor_field_poly_def monom_0 one_poly_def field_simps)
-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 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"
-  ultimately show "euclidean_size_field_poly (p * s) < euclidean_size_field_poly (q * s)"
-    by (auto simp add: euclidean_size_field_poly_def degree_mult_eq)
-next
-  fix p q r :: "'a poly" assume "p \<noteq> 0"
-  moreover assume "euclidean_size_field_poly r < euclidean_size_field_poly p"
-  ultimately show "(q * p + r) div p = q"
-    by (cases "r = 0")
-      (auto simp add: unit_factor_field_poly_def euclidean_size_field_poly_def div_poly_less)
-qed (auto simp: unit_factor_field_poly_def normalize_field_poly_def lead_coeff_mult 
-       euclidean_size_field_poly_def Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le)
-
-lemma field_poly_irreducible_imp_prime:
-  assumes "irreducible (p :: 'a :: field poly)"
-  shows   "prime_elem p"
-proof -
-  have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
-  from field_poly.irreducible_imp_prime_elem[of p] assms
-    show ?thesis unfolding irreducible_def prime_elem_def dvd_field_poly
-      comm_semiring_1.irreducible_def[OF A] comm_semiring_1.prime_elem_def[OF A] by blast
-qed
-
-lemma field_poly_prod_mset_prime_factorization:
-  assumes "(x :: 'a :: field poly) \<noteq> 0"
-  shows   "prod_mset (field_poly.prime_factorization x) = normalize_field_poly x"
-proof -
-  have A: "class.comm_monoid_mult op * (1 :: 'a poly)" ..
-  have "comm_monoid_mult.prod_mset op * (1 :: 'a poly) = prod_mset"
-    by (intro ext) (simp add: comm_monoid_mult.prod_mset_def[OF A] prod_mset_def)
-  with field_poly.prod_mset_prime_factorization[OF assms] show ?thesis by simp
-qed
-
-lemma field_poly_in_prime_factorization_imp_prime:
-  assumes "(p :: 'a :: field poly) \<in># field_poly.prime_factorization x"
-  shows   "prime_elem p"
-proof -
-  have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
-  have B: "class.normalization_semidom op div op + op - (0 :: 'a poly) op * 1 
-             unit_factor_field_poly normalize_field_poly" ..
-  from field_poly.in_prime_factors_imp_prime [of p x] assms
-    show ?thesis unfolding prime_elem_def dvd_field_poly
-      comm_semiring_1.prime_elem_def[OF A] normalization_semidom.prime_def[OF B] by blast
-qed
-
-
-subsection \<open>Primality and irreducibility in polynomial rings\<close>
-
-lemma nonconst_poly_irreducible_iff:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
-  assumes "degree p \<noteq> 0"
-  shows   "irreducible p \<longleftrightarrow> irreducible (fract_poly p) \<and> content p = 1"
-proof safe
-  assume p: "irreducible p"
-
-  from content_decompose[of p] guess p' . note p' = this
-  hence "p = [:content p:] * p'" by simp
-  from p this have "[:content p:] dvd 1 \<or> p' dvd 1" by (rule irreducibleD)
-  moreover have "\<not>p' dvd 1"
-  proof
-    assume "p' dvd 1"
-    hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff)
-    with assms show False by contradiction
-  qed
-  ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff)
-  
-  show "irreducible (map_poly to_fract p)"
-  proof (rule irreducibleI)
-    have "fract_poly p = 0 \<longleftrightarrow> p = 0" by (intro map_poly_eq_0_iff) auto
-    with assms show "map_poly to_fract p \<noteq> 0" by auto
-  next
-    show "\<not>is_unit (fract_poly p)"
-    proof
-      assume "is_unit (map_poly to_fract p)"
-      hence "degree (map_poly to_fract p) = 0"
-        by (auto simp: is_unit_poly_iff)
-      hence "degree p = 0" by (simp add: degree_map_poly)
-      with assms show False by contradiction
-   qed
- next
-   fix q r assume qr: "fract_poly p = q * r"
-   from content_decompose_fract[of q] guess cg q' . note q = this
-   from content_decompose_fract[of r] guess cr r' . note r = this
-   from qr q r p have nz: "cg \<noteq> 0" "cr \<noteq> 0" by auto
-   from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))"
-     by (simp add: q r)
-   from fract_poly_smult_eqE[OF this] guess a b . note ab = this
-   hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:)
-   with ab(4) have a: "a = normalize b" by (simp add: content_mult q r)
-   hence "normalize b = gcd a b" by simp
-   also from ab(3) have "\<dots> = 1" .
-   finally have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff)
-   
-   note eq
-   also from ab(1) \<open>a = 1\<close> have "cr * cg = to_fract b" by simp
-   also have "smult \<dots> (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp
-   finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult)
-   from p and this have "([:b:] * q') dvd 1 \<or> r' dvd 1" by (rule irreducibleD)
-   hence "q' dvd 1 \<or> r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left)
-   hence "fract_poly q' dvd 1 \<or> fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit)
-   with q r show "is_unit q \<or> is_unit r"
-     by (auto simp add: is_unit_smult_iff dvd_field_iff nz)
- qed
-
-next
-
-  assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1"
-  show "irreducible p"
-  proof (rule irreducibleI)
-    from irred show "p \<noteq> 0" by auto
-  next
-    from irred show "\<not>p dvd 1"
-      by (auto simp: irreducible_def dest: fract_poly_is_unit)
-  next
-    fix q r assume qr: "p = q * r"
-    hence "fract_poly p = fract_poly q * fract_poly r" by simp
-    from irred and this have "fract_poly q dvd 1 \<or> fract_poly r dvd 1" 
-      by (rule irreducibleD)
-    with primitive qr show "q dvd 1 \<or> r dvd 1"
-      by (auto simp:  content_prod_eq_1_iff is_unit_fract_poly_iff)
-  qed
-qed
-
-context
-begin
-
-private lemma irreducible_imp_prime_poly:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
-  assumes "irreducible p"
-  shows   "prime_elem p"
-proof (cases "degree p = 0")
-  case True
-  with assms show ?thesis
-    by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff
-             intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE)
-next
-  case False
-  from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1"
-    by (simp_all add: nonconst_poly_irreducible_iff)
-  from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime)
-  show ?thesis
-  proof (rule prime_elemI)
-    fix q r assume "p dvd q * r"
-    hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd)
-    hence "fract_poly p dvd fract_poly q * fract_poly r" by simp
-    from prime and this have "fract_poly p dvd fract_poly q \<or> fract_poly p dvd fract_poly r"
-      by (rule prime_elem_dvd_multD)
-    with primitive show "p dvd q \<or> p dvd r" by (auto dest: fract_poly_dvdD)
-  qed (insert assms, auto simp: irreducible_def)
-qed
-
-
-lemma degree_primitive_part_fract [simp]:
-  "degree (primitive_part_fract p) = degree p"
-proof -
-  have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))"
-    by (simp add: content_times_primitive_part_fract)
-  also have "degree \<dots> = degree (primitive_part_fract p)"
-    by (auto simp: degree_map_poly)
-  finally show ?thesis ..
-qed
-
-lemma irreducible_primitive_part_fract:
-  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
-  assumes "irreducible p"
-  shows   "irreducible (primitive_part_fract p)"
-proof -
-  from assms have deg: "degree (primitive_part_fract p) \<noteq> 0"
-    by (intro notI) 
-       (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff)
-  hence [simp]: "p \<noteq> 0" by auto
-
-  note \<open>irreducible p\<close>
-  also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" 
-    by (simp add: content_times_primitive_part_fract)
-  also have "irreducible \<dots> \<longleftrightarrow> irreducible (fract_poly (primitive_part_fract p))"
-    by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff)
-  finally show ?thesis using deg
-    by (simp add: nonconst_poly_irreducible_iff)
-qed
-
-lemma prime_elem_primitive_part_fract:
-  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
-  shows "irreducible p \<Longrightarrow> prime_elem (primitive_part_fract p)"
-  by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract)
-
-lemma irreducible_linear_field_poly:
-  fixes a b :: "'a::field"
-  assumes "b \<noteq> 0"
-  shows "irreducible [:a,b:]"
-proof (rule irreducibleI)
-  fix p q assume pq: "[:a,b:] = p * q"
-  also from pq assms have "degree \<dots> = degree p + degree q" 
-    by (intro degree_mult_eq) auto
-  finally have "degree p = 0 \<or> degree q = 0" using assms by auto
-  with assms pq show "is_unit p \<or> is_unit q"
-    by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE)
-qed (insert assms, auto simp: is_unit_poly_iff)
-
-lemma prime_elem_linear_field_poly:
-  "(b :: 'a :: field) \<noteq> 0 \<Longrightarrow> prime_elem [:a,b:]"
-  by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly)
-
-lemma irreducible_linear_poly:
-  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
-  shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> irreducible [:a,b:]"
-  by (auto intro!: irreducible_linear_field_poly 
-           simp:   nonconst_poly_irreducible_iff content_def map_poly_pCons)
-
-lemma prime_elem_linear_poly:
-  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
-  shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> prime_elem [:a,b:]"
-  by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly)
-
-end
-
- 
-subsection \<open>Prime factorisation of polynomials\<close>   
-
-context
-begin 
-
-private lemma poly_prime_factorization_exists_content_1:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
-  assumes "p \<noteq> 0" "content p = 1"
-  shows   "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p"
-proof -
-  let ?P = "field_poly.prime_factorization (fract_poly p)"
-  define c where "c = prod_mset (image_mset fract_content ?P)"
-  define c' where "c' = c * to_fract (lead_coeff p)"
-  define e where "e = prod_mset (image_mset primitive_part_fract ?P)"
-  define A where "A = image_mset (normalize \<circ> primitive_part_fract) ?P"
-  have "content e = (\<Prod>x\<in>#field_poly.prime_factorization (map_poly to_fract p). 
-                      content (primitive_part_fract x))"
-    by (simp add: e_def content_prod_mset multiset.map_comp o_def)
-  also have "image_mset (\<lambda>x. content (primitive_part_fract x)) ?P = image_mset (\<lambda>_. 1) ?P"
-    by (intro image_mset_cong content_primitive_part_fract) auto
-  finally have content_e: "content e = 1"
-    by simp    
-  
-  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 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
-  also have "image_mset id ?P = 
-               image_mset (\<lambda>x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P"
-    by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract)
-  also have "prod_mset \<dots> = smult c (fract_poly e)"
-    by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def)
-  also have "[:to_fract (lead_coeff p):] * \<dots> = smult c' (fract_poly e)"
-    by (simp add: c'_def)
-  finally have eq: "fract_poly p = smult c' (fract_poly e)" .
-  also obtain b where b: "c' = to_fract b" "is_unit b"
-  proof -
-    from fract_poly_smult_eqE[OF eq] guess a b . note ab = this
-    from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: )
-    with assms content_e have "a = normalize b" by (simp add: ab(4))
-    with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff)
-    with ab ab' have "c' = to_fract b" by auto
-    from this and \<open>is_unit b\<close> show ?thesis by (rule that)
-  qed
-  hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp
-  finally have "p = smult b e" by (simp only: fract_poly_eq_iff)
-  hence "p = [:b:] * e" by simp
-  with b have "normalize p = normalize e" 
-    by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff)
-  also have "normalize e = prod_mset A"
-    by (simp add: multiset.map_comp e_def A_def normalize_prod_mset)
-  finally have "prod_mset A = normalize p" ..
-  
-  have "prime_elem p" if "p \<in># A" for p
-    using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible 
-                        dest!: field_poly_in_prime_factorization_imp_prime )
-  from this and \<open>prod_mset A = normalize p\<close> show ?thesis
-    by (intro exI[of _ A]) blast
-qed
-
-lemma poly_prime_factorization_exists:
-  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
-  assumes "p \<noteq> 0"
-  shows   "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p"
-proof -
-  define B where "B = image_mset (\<lambda>x. [:x:]) (prime_factorization (content p))"
-  have "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize (primitive_part p)"
-    by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all)
-  then guess A by (elim exE conjE) note A = this
-  moreover from assms have "prod_mset B = [:content p:]"
-    by (simp add: B_def prod_mset_const_poly prod_mset_prime_factorization)
-  moreover have "\<forall>p. p \<in># B \<longrightarrow> prime_elem p"
-    by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime)
-  ultimately show ?thesis by (intro exI[of _ "B + A"]) auto
-qed
-
-end
-
-
-subsection \<open>Typeclass instances\<close>
-
-instance poly :: (factorial_ring_gcd) factorial_semiring
-  by standard (rule poly_prime_factorization_exists)  
-
-instantiation poly :: (factorial_ring_gcd) factorial_ring_gcd
-begin
-
-definition gcd_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  [code del]: "gcd_poly = gcd_factorial"
-
-definition lcm_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  [code del]: "lcm_poly = lcm_factorial"
-  
-definition Gcd_poly :: "'a poly set \<Rightarrow> 'a poly" where
- [code del]: "Gcd_poly = Gcd_factorial"
-
-definition Lcm_poly :: "'a poly set \<Rightarrow> 'a poly" where
- [code del]: "Lcm_poly = Lcm_factorial"
- 
-instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def)
-
-end
-
-instantiation poly :: ("{field,factorial_ring_gcd}") unique_euclidean_ring
-begin
-
-definition euclidean_size_poly :: "'a poly \<Rightarrow> nat"
-  where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)"
-
-definition uniqueness_constraint_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
-  where [simp]: "uniqueness_constraint_poly = top"
-
-instance 
-  by standard
-   (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq intro!: degree_mod_less' degree_mult_right_le
-    split: if_splits)
-
-end
-
-instance poly :: ("{field,factorial_ring_gcd}") euclidean_ring_gcd
-  by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI)
-    standard
-
-  
-subsection \<open>Polynomial GCD\<close>
-
-lemma gcd_poly_decompose:
-  fixes p q :: "'a :: factorial_ring_gcd poly"
-  shows "gcd p q = 
-           smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"
-proof (rule sym, rule gcdI)
-  have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd
-          [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all
-  thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p"
-    by simp
-next
-  have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd
-          [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all
-  thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q"
-    by simp
-next
-  fix d assume "d dvd p" "d dvd q"
-  hence "[:content d:] * primitive_part d dvd 
-           [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)"
-    by (intro mult_dvd_mono) auto
-  thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"
-    by simp
-qed (auto simp: normalize_smult)
-  
-
-lemma gcd_poly_pseudo_mod:
-  fixes p q :: "'a :: factorial_ring_gcd poly"
-  assumes nz: "q \<noteq> 0" and prim: "content p = 1" "content q = 1"
-  shows   "gcd p q = gcd q (primitive_part (pseudo_mod p q))"
-proof -
-  define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)"
-  define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]"
-  have [simp]: "primitive_part a = unit_factor a"
-    by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0)
-  from nz have [simp]: "a \<noteq> 0" by (auto simp: a_def)
-  
-  have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def)
-  have "gcd (q * r + s) q = gcd q s"
-    using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac)
-  with pseudo_divmod(1)[OF nz rs]
-    have "gcd (p * a) q = gcd q s" by (simp add: a_def)
-  also from prim have "gcd (p * a) q = gcd p q"
-    by (subst gcd_poly_decompose)
-       (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim 
-             simp del: mult_pCons_right )
-  also from prim have "gcd q s = gcd q (primitive_part s)"
-    by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim)
-  also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def)
-  finally show ?thesis .
-qed
-
-lemma degree_pseudo_mod_less:
-  assumes "q \<noteq> 0" "pseudo_mod p q \<noteq> 0"
-  shows   "degree (pseudo_mod p q) < degree q"
-  using pseudo_mod(2)[of q p] assms by auto
-
-function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  "gcd_poly_code_aux p q = 
-     (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" 
-by auto
-termination
-  by (relation "measure ((\<lambda>p. if p = 0 then 0 else Suc (degree p)) \<circ> snd)")
-     (auto simp: degree_pseudo_mod_less)
-
-declare gcd_poly_code_aux.simps [simp del]
-
-lemma gcd_poly_code_aux_correct:
-  assumes "content p = 1" "q = 0 \<or> content q = 1"
-  shows   "gcd_poly_code_aux p q = gcd p q"
-  using assms
-proof (induction p q rule: gcd_poly_code_aux.induct)
-  case (1 p q)
-  show ?case
-  proof (cases "q = 0")
-    case True
-    thus ?thesis by (subst gcd_poly_code_aux.simps) auto
-  next
-    case False
-    hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))"
-      by (subst gcd_poly_code_aux.simps) simp_all
-    also from "1.prems" False 
-      have "primitive_part (pseudo_mod p q) = 0 \<or> 
-              content (primitive_part (pseudo_mod p q)) = 1"
-      by (cases "pseudo_mod p q = 0") auto
-    with "1.prems" False 
-      have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = 
-              gcd q (primitive_part (pseudo_mod p q))"
-      by (intro 1) simp_all
-    also from "1.prems" False 
-      have "\<dots> = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto
-    finally show ?thesis .
-  qed
-qed
-
-definition gcd_poly_code 
-    :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" 
-  where "gcd_poly_code p q = 
-           (if p = 0 then normalize q else if q = 0 then normalize p else
-              smult (gcd (content p) (content q)) 
-                (gcd_poly_code_aux (primitive_part p) (primitive_part q)))"
-
-lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q"
-  by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric])
-
-lemma lcm_poly_code [code]: 
-  fixes p q :: "'a :: factorial_ring_gcd poly"
-  shows "lcm p q = normalize (p * q) div gcd p q"
-  by (fact lcm_gcd)
-
-lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"]
-lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"]
-
-text \<open>Example:
-  @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval}
-\<close>
-  
-end
--- a/src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -7,7 +7,7 @@
 section \<open>The Nonstandard Primes as an Extension of the Prime Numbers\<close>
 
 theory NSPrimes
-  imports "~~/src/HOL/Number_Theory/Primes" "../Hyperreal"
+  imports "~~/src/HOL/Computational_Algebra/Primes" "../Hyperreal"
 begin
 
 text \<open>These can be used to derive an alternative proof of the infinitude of
--- a/src/HOL/Number_Theory/Cong.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Number_Theory/Cong.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -26,7 +26,7 @@
 section \<open>Congruence\<close>
 
 theory Cong
-imports Primes
+imports "~~/src/HOL/Computational_Algebra/Primes"
 begin
 
 subsection \<open>Turn off \<open>One_nat_def\<close>\<close>
--- a/src/HOL/Number_Theory/Eratosthenes.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Number_Theory/Eratosthenes.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -5,7 +5,7 @@
 section \<open>The sieve of Eratosthenes\<close>
 
 theory Eratosthenes
-imports Main Primes
+  imports Main "~~/src/HOL/Computational_Algebra/Primes"
 begin
 
 
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,632 +0,0 @@
-(*  Title:      HOL/Number_Theory/Euclidean_Algorithm.thy
-    Author:     Manuel Eberl, TU Muenchen
-*)
-
-section \<open>Abstract euclidean algorithm in euclidean (semi)rings\<close>
-
-theory Euclidean_Algorithm
-imports
-  "~~/src/HOL/Number_Theory/Factorial_Ring"
-begin
-
-subsection \<open>Generic construction of the (simple) euclidean algorithm\<close>
-  
-context euclidean_semiring
-begin
-
-context
-begin
-
-qualified function gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
-  where "gcd a b = (if b = 0 then normalize a else gcd b (a mod b))"
-  by pat_completeness simp
-termination
-  by (relation "measure (euclidean_size \<circ> snd)") (simp_all add: mod_size_less)
-
-declare gcd.simps [simp del]
-
-lemma eucl_induct [case_names zero mod]:
-  assumes H1: "\<And>b. P b 0"
-  and H2: "\<And>a b. b \<noteq> 0 \<Longrightarrow> P b (a mod b) \<Longrightarrow> P a b"
-  shows "P a b"
-proof (induct a b rule: gcd.induct)
-  case (1 a b)
-  show ?case
-  proof (cases "b = 0")
-    case True then show "P a b" by simp (rule H1)
-  next
-    case False
-    then have "P b (a mod b)"
-      by (rule "1.hyps")
-    with \<open>b \<noteq> 0\<close> show "P a b"
-      by (blast intro: H2)
-  qed
-qed
-  
-qualified lemma gcd_0:
-  "gcd a 0 = normalize a"
-  by (simp add: gcd.simps [of a 0])
-  
-qualified lemma gcd_mod:
-  "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd b a"
-  by (simp add: gcd.simps [of b 0] gcd.simps [of b a])
-
-qualified definition lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
-  where "lcm a b = normalize (a * b) div gcd a b"
-
-qualified definition Lcm :: "'a set \<Rightarrow> 'a" \<comment>
-    \<open>Somewhat complicated definition of Lcm that has the advantage of working
-    for infinite sets as well\<close>
-  where
-  [code del]: "Lcm A = (if \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) then
-     let l = SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l =
-       (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)
-       in normalize l 
-      else 0)"
-
-qualified definition Gcd :: "'a set \<Rightarrow> 'a"
-  where [code del]: "Gcd A = Lcm {d. \<forall>a\<in>A. d dvd a}"
-
-end    
-
-lemma semiring_gcd:
-  "class.semiring_gcd one zero times gcd lcm
-    divide plus minus unit_factor normalize"
-proof
-  show "gcd a b dvd a"
-    and "gcd a b dvd b" for a b
-    by (induct a b rule: eucl_induct)
-      (simp_all add: local.gcd_0 local.gcd_mod dvd_mod_iff)
-next
-  show "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> c dvd gcd a b" for a b c
-  proof (induct a b rule: eucl_induct)
-    case (zero a) from \<open>c dvd a\<close> show ?case
-      by (rule dvd_trans) (simp add: local.gcd_0)
-  next
-    case (mod a b)
-    then show ?case
-      by (simp add: local.gcd_mod dvd_mod_iff)
-  qed
-next
-  show "normalize (gcd a b) = gcd a b" for a b
-    by (induct a b rule: eucl_induct)
-      (simp_all add: local.gcd_0 local.gcd_mod)
-next
-  show "lcm a b = normalize (a * b) div gcd a b" for a b
-    by (fact local.lcm_def)
-qed
-
-interpretation semiring_gcd one zero times gcd lcm
-  divide plus minus unit_factor normalize
-  by (fact semiring_gcd)
-  
-lemma semiring_Gcd:
-  "class.semiring_Gcd one zero times gcd lcm Gcd Lcm
-    divide plus minus unit_factor normalize"
-proof -
-  show ?thesis
-  proof
-    have "(\<forall>a\<in>A. a dvd Lcm A) \<and> (\<forall>b. (\<forall>a\<in>A. a dvd b) \<longrightarrow> Lcm A dvd b)" for A
-    proof (cases "\<exists>l. l \<noteq>  0 \<and> (\<forall>a\<in>A. a dvd l)")
-      case False
-      then have "Lcm A = 0"
-        by (auto simp add: local.Lcm_def)
-      with False show ?thesis
-        by auto
-    next
-      case True
-      then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0" "\<forall>a\<in>A. a dvd l\<^sub>0" by blast
-      define n where "n = (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
-      define l where "l = (SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
-      have "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n"
-        apply (subst n_def)
-        apply (rule LeastI [of _ "euclidean_size l\<^sub>0"])
-        apply (rule exI [of _ l\<^sub>0])
-        apply (simp add: l\<^sub>0_props)
-        done
-      from someI_ex [OF this] have "l \<noteq> 0" and "\<forall>a\<in>A. a dvd l"
-        and "euclidean_size l = n" 
-        unfolding l_def by simp_all
-      {
-        fix l' assume "\<forall>a\<in>A. a dvd l'"
-        with \<open>\<forall>a\<in>A. a dvd l\<close> have "\<forall>a\<in>A. a dvd gcd l l'"
-          by (auto intro: gcd_greatest)
-        moreover from \<open>l \<noteq> 0\<close> have "gcd l l' \<noteq> 0"
-          by simp
-        ultimately have "\<exists>b. b \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd b) \<and> 
-          euclidean_size b = euclidean_size (gcd l l')"
-          by (intro exI [of _ "gcd l l'"], auto)
-        then have "euclidean_size (gcd l l') \<ge> n"
-          by (subst n_def) (rule Least_le)
-        moreover have "euclidean_size (gcd l l') \<le> n"
-        proof -
-          have "gcd l l' dvd l"
-            by simp
-          then obtain a where "l = gcd l l' * a" ..
-          with \<open>l \<noteq> 0\<close> have "a \<noteq> 0"
-            by auto
-          hence "euclidean_size (gcd l l') \<le> euclidean_size (gcd l l' * a)"
-            by (rule size_mult_mono)
-          also have "gcd l l' * a = l" using \<open>l = gcd l l' * a\<close> ..
-          also note \<open>euclidean_size l = n\<close>
-          finally show "euclidean_size (gcd l l') \<le> n" .
-        qed
-        ultimately have *: "euclidean_size l = euclidean_size (gcd l l')" 
-          by (intro le_antisym, simp_all add: \<open>euclidean_size l = n\<close>)
-        from \<open>l \<noteq> 0\<close> have "l dvd gcd l l'"
-          by (rule dvd_euclidean_size_eq_imp_dvd) (auto simp add: *)
-        hence "l dvd l'" by (rule dvd_trans [OF _ gcd_dvd2])
-      }
-      with \<open>\<forall>a\<in>A. a dvd l\<close> and \<open>l \<noteq> 0\<close>
-        have "(\<forall>a\<in>A. a dvd normalize l) \<and> 
-          (\<forall>l'. (\<forall>a\<in>A. a dvd l') \<longrightarrow> normalize l dvd l')"
-        by auto
-      also from True have "normalize l = Lcm A"
-        by (simp add: local.Lcm_def Let_def n_def l_def)
-      finally show ?thesis .
-    qed
-    then show dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A"
-      and Lcm_least: "(\<And>a. a \<in> A \<Longrightarrow> a dvd b) \<Longrightarrow> Lcm A dvd b" for A and a b
-      by auto
-    show "a \<in> A \<Longrightarrow> Gcd A dvd a" for A and a
-      by (auto simp add: local.Gcd_def intro: Lcm_least)
-    show "(\<And>a. a \<in> A \<Longrightarrow> b dvd a) \<Longrightarrow> b dvd Gcd A" for A and b
-      by (auto simp add: local.Gcd_def intro: dvd_Lcm)
-    show [simp]: "normalize (Lcm A) = Lcm A" for A
-      by (simp add: local.Lcm_def)
-    show "normalize (Gcd A) = Gcd A" for A
-      by (simp add: local.Gcd_def)
-  qed
-qed
-
-interpretation semiring_Gcd one zero times gcd lcm Gcd Lcm
-    divide plus minus unit_factor normalize
-  by (fact semiring_Gcd)
-
-subclass factorial_semiring
-proof -
-  show "class.factorial_semiring divide plus minus zero times one
-     unit_factor normalize"
-  proof (standard, rule factorial_semiring_altI_aux) \<comment> \<open>FIXME rule\<close>
-    fix x assume "x \<noteq> 0"
-    thus "finite {p. p dvd x \<and> normalize p = p}"
-    proof (induction "euclidean_size x" arbitrary: x rule: less_induct)
-      case (less x)
-      show ?case
-      proof (cases "\<exists>y. y dvd x \<and> \<not>x dvd y \<and> \<not>is_unit y")
-        case False
-        have "{p. p dvd x \<and> normalize p = p} \<subseteq> {1, normalize x}"
-        proof
-          fix p assume p: "p \<in> {p. p dvd x \<and> normalize p = p}"
-          with False have "is_unit p \<or> x dvd p" by blast
-          thus "p \<in> {1, normalize x}"
-          proof (elim disjE)
-            assume "is_unit p"
-            hence "normalize p = 1" by (simp add: is_unit_normalize)
-            with p show ?thesis by simp
-          next
-            assume "x dvd p"
-            with p have "normalize p = normalize x" by (intro associatedI) simp_all
-            with p show ?thesis by simp
-          qed
-        qed
-        moreover have "finite \<dots>" by simp
-        ultimately show ?thesis by (rule finite_subset)
-      next
-        case True
-        then obtain y where y: "y dvd x" "\<not>x dvd y" "\<not>is_unit y" by blast
-        define z where "z = x div y"
-        let ?fctrs = "\<lambda>x. {p. p dvd x \<and> normalize p = p}"
-        from y have x: "x = y * z" by (simp add: z_def)
-        with less.prems have "y \<noteq> 0" "z \<noteq> 0" by auto
-        have normalized_factors_product:
-          "{p. p dvd a * b \<and> normalize p = p} = 
-             (\<lambda>(x,y). x * y) ` ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})" for a b
-        proof safe
-          fix p assume p: "p dvd a * b" "normalize p = p"
-          from dvd_productE[OF p(1)] guess x y . note xy = this
-          define x' y' where "x' = normalize x" and "y' = normalize y"
-          have "p = x' * y'"
-            by (subst p(2) [symmetric]) (simp add: xy x'_def y'_def normalize_mult)
-          moreover from xy have "normalize x' = x'" "normalize y' = y'" "x' dvd a" "y' dvd b" 
-            by (simp_all add: x'_def y'_def)
-          ultimately show "p \<in> (\<lambda>(x, y). x * y) ` 
-            ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})"
-            by blast
-        qed (auto simp: normalize_mult mult_dvd_mono)
-        from x y have "\<not>is_unit z" by (auto simp: mult_unit_dvd_iff)
-        have "?fctrs x = (\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z)"
-          by (subst x) (rule normalized_factors_product)
-        also have "\<not>y * z dvd y * 1" "\<not>y * z dvd 1 * z"
-          by (subst dvd_times_left_cancel_iff dvd_times_right_cancel_iff; fact)+
-        hence "finite ((\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z))"
-          by (intro finite_imageI finite_cartesian_product less dvd_proper_imp_size_less)
-             (auto simp: x)
-        finally show ?thesis .
-      qed
-    qed
-  next
-    fix p
-    assume "irreducible p"
-    then show "prime_elem p"
-      by (rule irreducible_imp_prime_elem_gcd)
-  qed
-qed
-
-lemma Gcd_eucl_set [code]:
-  "Gcd (set xs) = fold gcd xs 0"
-  by (fact Gcd_set_eq_fold)
-
-lemma Lcm_eucl_set [code]:
-  "Lcm (set xs) = fold lcm xs 1"
-  by (fact Lcm_set_eq_fold)
- 
-end
-
-hide_const (open) gcd lcm Gcd Lcm
-
-lemma prime_elem_int_abs_iff [simp]:
-  fixes p :: int
-  shows "prime_elem \<bar>p\<bar> \<longleftrightarrow> prime_elem p"
-  using prime_elem_normalize_iff [of p] by simp
-  
-lemma prime_elem_int_minus_iff [simp]:
-  fixes p :: int
-  shows "prime_elem (- p) \<longleftrightarrow> prime_elem p"
-  using prime_elem_normalize_iff [of "- p"] by simp
-
-lemma prime_int_iff:
-  fixes p :: int
-  shows "prime p \<longleftrightarrow> p > 0 \<and> prime_elem p"
-  by (auto simp add: prime_def dest: prime_elem_not_zeroI)
-  
-  
-subsection \<open>The (simple) euclidean algorithm as gcd computation\<close>
-  
-class euclidean_semiring_gcd = euclidean_semiring + gcd + Gcd +
-  assumes gcd_eucl: "Euclidean_Algorithm.gcd = GCD.gcd"
-    and lcm_eucl: "Euclidean_Algorithm.lcm = GCD.lcm"
-  assumes Gcd_eucl: "Euclidean_Algorithm.Gcd = GCD.Gcd"
-    and Lcm_eucl: "Euclidean_Algorithm.Lcm = GCD.Lcm"
-begin
-
-subclass semiring_gcd
-  unfolding gcd_eucl [symmetric] lcm_eucl [symmetric]
-  by (fact semiring_gcd)
-
-subclass semiring_Gcd
-  unfolding  gcd_eucl [symmetric] lcm_eucl [symmetric]
-    Gcd_eucl [symmetric] Lcm_eucl [symmetric]
-  by (fact semiring_Gcd)
-
-subclass factorial_semiring_gcd
-proof
-  show "gcd a b = gcd_factorial a b" for a b
-    apply (rule sym)
-    apply (rule gcdI)
-       apply (fact gcd_lcm_factorial)+
-    done
-  then show "lcm a b = lcm_factorial a b" for a b
-    by (simp add: lcm_factorial_gcd_factorial lcm_gcd)
-  show "Gcd A = Gcd_factorial A" for A
-    apply (rule sym)
-    apply (rule GcdI)
-       apply (fact gcd_lcm_factorial)+
-    done
-  show "Lcm A = Lcm_factorial A" for A
-    apply (rule sym)
-    apply (rule LcmI)
-       apply (fact gcd_lcm_factorial)+
-    done
-qed
-
-lemma gcd_mod_right [simp]:
-  "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd a b"
-  unfolding gcd.commute [of a b]
-  by (simp add: gcd_eucl [symmetric] local.gcd_mod)
-
-lemma gcd_mod_left [simp]:
-  "b \<noteq> 0 \<Longrightarrow> gcd (a mod b) b = gcd a b"
-  by (drule gcd_mod_right [of _ a]) (simp add: gcd.commute)
-
-lemma euclidean_size_gcd_le1 [simp]:
-  assumes "a \<noteq> 0"
-  shows "euclidean_size (gcd a b) \<le> euclidean_size a"
-proof -
-  from gcd_dvd1 obtain c where A: "a = gcd a b * c" ..
-  with assms have "c \<noteq> 0"
-    by auto
-  moreover from this
-  have "euclidean_size (gcd a b) \<le> euclidean_size (gcd a b * c)"
-    by (rule size_mult_mono)
-  with A show ?thesis
-    by simp
-qed
-
-lemma euclidean_size_gcd_le2 [simp]:
-  "b \<noteq> 0 \<Longrightarrow> euclidean_size (gcd a b) \<le> euclidean_size b"
-  by (subst gcd.commute, rule euclidean_size_gcd_le1)
-
-lemma euclidean_size_gcd_less1:
-  assumes "a \<noteq> 0" and "\<not> a dvd b"
-  shows "euclidean_size (gcd a b) < euclidean_size a"
-proof (rule ccontr)
-  assume "\<not>euclidean_size (gcd a b) < euclidean_size a"
-  with \<open>a \<noteq> 0\<close> have A: "euclidean_size (gcd a b) = euclidean_size a"
-    by (intro le_antisym, simp_all)
-  have "a dvd gcd a b"
-    by (rule dvd_euclidean_size_eq_imp_dvd) (simp_all add: assms A)
-  hence "a dvd b" using dvd_gcdD2 by blast
-  with \<open>\<not> a dvd b\<close> show False by contradiction
-qed
-
-lemma euclidean_size_gcd_less2:
-  assumes "b \<noteq> 0" and "\<not> b dvd a"
-  shows "euclidean_size (gcd a b) < euclidean_size b"
-  using assms by (subst gcd.commute, rule euclidean_size_gcd_less1)
-
-lemma euclidean_size_lcm_le1: 
-  assumes "a \<noteq> 0" and "b \<noteq> 0"
-  shows "euclidean_size a \<le> euclidean_size (lcm a b)"
-proof -
-  have "a dvd lcm a b" by (rule dvd_lcm1)
-  then obtain c where A: "lcm a b = a * c" ..
-  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "c \<noteq> 0" by (auto simp: lcm_eq_0_iff)
-  then show ?thesis by (subst A, intro size_mult_mono)
-qed
-
-lemma euclidean_size_lcm_le2:
-  "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> euclidean_size b \<le> euclidean_size (lcm a b)"
-  using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps)
-
-lemma euclidean_size_lcm_less1:
-  assumes "b \<noteq> 0" and "\<not> b dvd a"
-  shows "euclidean_size a < euclidean_size (lcm a b)"
-proof (rule ccontr)
-  from assms have "a \<noteq> 0" by auto
-  assume "\<not>euclidean_size a < euclidean_size (lcm a b)"
-  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "euclidean_size (lcm a b) = euclidean_size a"
-    by (intro le_antisym, simp, intro euclidean_size_lcm_le1)
-  with assms have "lcm a b dvd a" 
-    by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_eq_0_iff)
-  hence "b dvd a" by (rule lcm_dvdD2)
-  with \<open>\<not>b dvd a\<close> show False by contradiction
-qed
-
-lemma euclidean_size_lcm_less2:
-  assumes "a \<noteq> 0" and "\<not> a dvd b"
-  shows "euclidean_size b < euclidean_size (lcm a b)"
-  using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps)
-
-end
-
-lemma factorial_euclidean_semiring_gcdI:
-  "OFCLASS('a::{factorial_semiring_gcd, euclidean_semiring}, euclidean_semiring_gcd_class)"
-proof
-  interpret semiring_Gcd 1 0 times
-    Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm
-    Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm
-    divide plus minus unit_factor normalize
-    rewrites "dvd.dvd op * = Rings.dvd"
-    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
-  show [simp]: "Euclidean_Algorithm.gcd = (gcd :: 'a \<Rightarrow> _)"
-  proof (rule ext)+
-    fix a b :: 'a
-    show "Euclidean_Algorithm.gcd a b = gcd a b"
-    proof (induct a b rule: eucl_induct)
-      case zero
-      then show ?case
-        by simp
-    next
-      case (mod a b)
-      moreover have "gcd b (a mod b) = gcd b a"
-        using GCD.gcd_add_mult [of b "a div b" "a mod b", symmetric]
-          by (simp add: div_mult_mod_eq)
-      ultimately show ?case
-        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
-    qed
-  qed
-  show [simp]: "Euclidean_Algorithm.Lcm = (Lcm :: 'a set \<Rightarrow> _)"
-    by (auto intro!: Lcm_eqI GCD.dvd_Lcm GCD.Lcm_least)
-  show "Euclidean_Algorithm.lcm = (lcm :: 'a \<Rightarrow> _)"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
-  show "Euclidean_Algorithm.Gcd = (Gcd :: 'a set \<Rightarrow> _)"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
-qed
-
-
-subsection \<open>The extended euclidean algorithm\<close>
-  
-class euclidean_ring_gcd = euclidean_semiring_gcd + idom
-begin
-
-subclass euclidean_ring ..
-subclass ring_gcd ..
-subclass factorial_ring_gcd ..
-
-function euclid_ext_aux :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
-  where "euclid_ext_aux s' s t' t r' r = (
-     if r = 0 then let c = 1 div unit_factor r' in ((s' * c, t' * c), normalize r')
-     else let q = r' div r
-          in euclid_ext_aux s (s' - q * s) t (t' - q * t) r (r' mod r))"
-  by auto
-termination
-  by (relation "measure (\<lambda>(_, _, _, _, _, b). euclidean_size b)")
-    (simp_all add: mod_size_less)
-
-abbreviation (input) euclid_ext :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
-  where "euclid_ext \<equiv> euclid_ext_aux 1 0 0 1"
-    
-lemma
-  assumes "gcd r' r = gcd a b"
-  assumes "s' * a + t' * b = r'"
-  assumes "s * a + t * b = r"
-  assumes "euclid_ext_aux s' s t' t r' r = ((x, y), c)"
-  shows euclid_ext_aux_eq_gcd: "c = gcd a b"
-    and euclid_ext_aux_bezout: "x * a + y * b = gcd a b"
-proof -
-  have "case euclid_ext_aux s' s t' t r' r of ((x, y), c) \<Rightarrow> 
-    x * a + y * b = c \<and> c = gcd a b" (is "?P (euclid_ext_aux s' s t' t r' r)")
-    using assms(1-3)
-  proof (induction s' s t' t r' r rule: euclid_ext_aux.induct)
-    case (1 s' s t' t r' r)
-    show ?case
-    proof (cases "r = 0")
-      case True
-      hence "euclid_ext_aux s' s t' t r' r = 
-               ((s' div unit_factor r', t' div unit_factor r'), normalize r')"
-        by (subst euclid_ext_aux.simps) (simp add: Let_def)
-      also have "?P \<dots>"
-      proof safe
-        have "s' div unit_factor r' * a + t' div unit_factor r' * b = 
-                (s' * a + t' * b) div unit_factor r'"
-          by (cases "r' = 0") (simp_all add: unit_div_commute)
-        also have "s' * a + t' * b = r'" by fact
-        also have "\<dots> div unit_factor r' = normalize r'" by simp
-        finally show "s' div unit_factor r' * a + t' div unit_factor r' * b = normalize r'" .
-      next
-        from "1.prems" True show "normalize r' = gcd a b"
-          by simp
-      qed
-      finally show ?thesis .
-    next
-      case False
-      hence "euclid_ext_aux s' s t' t r' r = 
-             euclid_ext_aux s (s' - r' div r * s) t (t' - r' div r * t) r (r' mod r)"
-        by (subst euclid_ext_aux.simps) (simp add: Let_def)
-      also from "1.prems" False have "?P \<dots>"
-      proof (intro "1.IH")
-        have "(s' - r' div r * s) * a + (t' - r' div r * t) * b =
-              (s' * a + t' * b) - r' div r * (s * a + t * b)" by (simp add: algebra_simps)
-        also have "s' * a + t' * b = r'" by fact
-        also have "s * a + t * b = r" by fact
-        also have "r' - r' div r * r = r' mod r" using div_mult_mod_eq [of r' r]
-          by (simp add: algebra_simps)
-        finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" .
-      qed (auto simp: gcd_mod_right algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute)
-      finally show ?thesis .
-    qed
-  qed
-  with assms(4) show "c = gcd a b" "x * a + y * b = gcd a b"
-    by simp_all
-qed
-
-declare euclid_ext_aux.simps [simp del]
-
-definition bezout_coefficients :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<times> 'a"
-  where [code]: "bezout_coefficients a b = fst (euclid_ext a b)"
-
-lemma bezout_coefficients_0: 
-  "bezout_coefficients a 0 = (1 div unit_factor a, 0)"
-  by (simp add: bezout_coefficients_def euclid_ext_aux.simps)
-
-lemma bezout_coefficients_left_0: 
-  "bezout_coefficients 0 a = (0, 1 div unit_factor a)"
-  by (simp add: bezout_coefficients_def euclid_ext_aux.simps)
-
-lemma bezout_coefficients:
-  assumes "bezout_coefficients a b = (x, y)"
-  shows "x * a + y * b = gcd a b"
-  using assms by (simp add: bezout_coefficients_def
-    euclid_ext_aux_bezout [of a b a b 1 0 0 1 x y] prod_eq_iff)
-
-lemma bezout_coefficients_fst_snd:
-  "fst (bezout_coefficients a b) * a + snd (bezout_coefficients a b) * b = gcd a b"
-  by (rule bezout_coefficients) simp
-
-lemma euclid_ext_eq [simp]:
-  "euclid_ext a b = (bezout_coefficients a b, gcd a b)" (is "?p = ?q")
-proof
-  show "fst ?p = fst ?q"
-    by (simp add: bezout_coefficients_def)
-  have "snd (euclid_ext_aux 1 0 0 1 a b) = gcd a b"
-    by (rule euclid_ext_aux_eq_gcd [of a b a b 1 0 0 1])
-      (simp_all add: prod_eq_iff)
-  then show "snd ?p = snd ?q"
-    by simp
-qed
-
-declare euclid_ext_eq [symmetric, code_unfold]
-
-end
-
-
-subsection \<open>Typical instances\<close>
-
-instance nat :: euclidean_semiring_gcd
-proof
-  interpret semiring_Gcd 1 0 times
-    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
-    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
-    divide plus minus unit_factor normalize
-    rewrites "dvd.dvd op * = Rings.dvd"
-    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
-  show [simp]: "(Euclidean_Algorithm.gcd :: nat \<Rightarrow> _) = gcd"
-  proof (rule ext)+
-    fix m n :: nat
-    show "Euclidean_Algorithm.gcd m n = gcd m n"
-    proof (induct m n rule: eucl_induct)
-      case zero
-      then show ?case
-        by simp
-    next
-      case (mod m n)
-      then have "gcd n (m mod n) = gcd n m"
-        using gcd_nat.simps [of m n] by (simp add: ac_simps)
-      with mod show ?case
-        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
-    qed
-  qed
-  show [simp]: "(Euclidean_Algorithm.Lcm :: nat set \<Rightarrow> _) = Lcm"
-    by (auto intro!: ext Lcm_eqI)
-  show "(Euclidean_Algorithm.lcm :: nat \<Rightarrow> _) = lcm"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
-  show "(Euclidean_Algorithm.Gcd :: nat set \<Rightarrow> _) = Gcd"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
-qed
-
-instance int :: euclidean_ring_gcd
-proof
-  interpret semiring_Gcd 1 0 times
-    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
-    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
-    divide plus minus unit_factor normalize
-    rewrites "dvd.dvd op * = Rings.dvd"
-    by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff)
-  show [simp]: "(Euclidean_Algorithm.gcd :: int \<Rightarrow> _) = gcd"
-  proof (rule ext)+
-    fix k l :: int
-    show "Euclidean_Algorithm.gcd k l = gcd k l"
-    proof (induct k l rule: eucl_induct)
-      case zero
-      then show ?case
-        by simp
-    next
-      case (mod k l)
-      have "gcd l (k mod l) = gcd l k"
-      proof (cases l "0::int" rule: linorder_cases)
-        case less
-        then show ?thesis
-          using gcd_non_0_int [of "- l" "- k"] by (simp add: ac_simps)
-      next
-        case equal
-        with mod show ?thesis
-          by simp
-      next
-        case greater
-        then show ?thesis
-          using gcd_non_0_int [of l k] by (simp add: ac_simps)
-      qed
-      with mod show ?case
-        by (simp add: Euclidean_Algorithm.gcd_mod ac_simps)
-    qed
-  qed
-  show [simp]: "(Euclidean_Algorithm.Lcm :: int set \<Rightarrow> _) = Lcm"
-    by (auto intro!: ext Lcm_eqI)
-  show "(Euclidean_Algorithm.lcm :: int \<Rightarrow> _) = lcm"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd)
-  show "(Euclidean_Algorithm.Gcd :: int set \<Rightarrow> _) = Gcd"
-    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
-qed
-
-end
--- a/src/HOL/Number_Theory/Factorial_Ring.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1818 +0,0 @@
-(*  Title:      HOL/Number_Theory/Factorial_Ring.thy
-    Author:     Manuel Eberl, TU Muenchen
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-section \<open>Factorial (semi)rings\<close>
-
-theory Factorial_Ring
-imports 
-  Main
-  "../GCD"
-  "~~/src/HOL/Library/Multiset"
-begin
-
-subsection \<open>Irreducible and prime elements\<close>
-
-context comm_semiring_1
-begin
-
-definition irreducible :: "'a \<Rightarrow> bool" where
-  "irreducible p \<longleftrightarrow> p \<noteq> 0 \<and> \<not>p dvd 1 \<and> (\<forall>a b. p = a * b \<longrightarrow> a dvd 1 \<or> b dvd 1)"
-
-lemma not_irreducible_zero [simp]: "\<not>irreducible 0"
-  by (simp add: irreducible_def)
-
-lemma irreducible_not_unit: "irreducible p \<Longrightarrow> \<not>p dvd 1"
-  by (simp add: irreducible_def)
-
-lemma not_irreducible_one [simp]: "\<not>irreducible 1"
-  by (simp add: irreducible_def)
-
-lemma irreducibleI:
-  "p \<noteq> 0 \<Longrightarrow> \<not>p dvd 1 \<Longrightarrow> (\<And>a b. p = a * b \<Longrightarrow> a dvd 1 \<or> b dvd 1) \<Longrightarrow> irreducible p"
-  by (simp add: irreducible_def)
-
-lemma irreducibleD: "irreducible p \<Longrightarrow> p = a * b \<Longrightarrow> a dvd 1 \<or> b dvd 1"
-  by (simp add: irreducible_def)
-
-definition prime_elem :: "'a \<Rightarrow> bool" where
-  "prime_elem p \<longleftrightarrow> p \<noteq> 0 \<and> \<not>p dvd 1 \<and> (\<forall>a b. p dvd (a * b) \<longrightarrow> p dvd a \<or> p dvd b)"
-
-lemma not_prime_elem_zero [simp]: "\<not>prime_elem 0"
-  by (simp add: prime_elem_def)
-
-lemma prime_elem_not_unit: "prime_elem p \<Longrightarrow> \<not>p dvd 1"
-  by (simp add: prime_elem_def)
-
-lemma prime_elemI:
-    "p \<noteq> 0 \<Longrightarrow> \<not>p dvd 1 \<Longrightarrow> (\<And>a b. p dvd (a * b) \<Longrightarrow> p dvd a \<or> p dvd b) \<Longrightarrow> prime_elem p"
-  by (simp add: prime_elem_def)
-
-lemma prime_elem_dvd_multD:
-    "prime_elem p \<Longrightarrow> p dvd (a * b) \<Longrightarrow> p dvd a \<or> p dvd b"
-  by (simp add: prime_elem_def)
-
-lemma prime_elem_dvd_mult_iff:
-  "prime_elem p \<Longrightarrow> p dvd (a * b) \<longleftrightarrow> p dvd a \<or> p dvd b"
-  by (auto simp: prime_elem_def)
-
-lemma not_prime_elem_one [simp]:
-  "\<not> prime_elem 1"
-  by (auto dest: prime_elem_not_unit)
-
-lemma prime_elem_not_zeroI:
-  assumes "prime_elem p"
-  shows "p \<noteq> 0"
-  using assms by (auto intro: ccontr)
-
-lemma prime_elem_dvd_power: 
-  "prime_elem p \<Longrightarrow> p dvd x ^ n \<Longrightarrow> p dvd x"
-  by (induction n) (auto dest: prime_elem_dvd_multD intro: dvd_trans[of _ 1])
-
-lemma prime_elem_dvd_power_iff:
-  "prime_elem p \<Longrightarrow> n > 0 \<Longrightarrow> p dvd x ^ n \<longleftrightarrow> p dvd x"
-  by (auto dest: prime_elem_dvd_power intro: dvd_trans)
-
-lemma prime_elem_imp_nonzero [simp]:
-  "ASSUMPTION (prime_elem x) \<Longrightarrow> x \<noteq> 0"
-  unfolding ASSUMPTION_def by (rule prime_elem_not_zeroI)
-
-lemma prime_elem_imp_not_one [simp]:
-  "ASSUMPTION (prime_elem x) \<Longrightarrow> x \<noteq> 1"
-  unfolding ASSUMPTION_def by auto
-
-end
-
-context algebraic_semidom
-begin
-
-lemma prime_elem_imp_irreducible:
-  assumes "prime_elem p"
-  shows   "irreducible p"
-proof (rule irreducibleI)
-  fix a b
-  assume p_eq: "p = a * b"
-  with assms have nz: "a \<noteq> 0" "b \<noteq> 0" by auto
-  from p_eq have "p dvd a * b" by simp
-  with \<open>prime_elem p\<close> have "p dvd a \<or> p dvd b" by (rule prime_elem_dvd_multD)
-  with \<open>p = a * b\<close> have "a * b dvd 1 * b \<or> a * b dvd a * 1" by auto
-  thus "a dvd 1 \<or> b dvd 1"
-    by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)])
-qed (insert assms, simp_all add: prime_elem_def)
-
-lemma (in algebraic_semidom) unit_imp_no_irreducible_divisors:
-  assumes "is_unit x" "irreducible p"
-  shows   "\<not>p dvd x"
-proof (rule notI)
-  assume "p dvd x"
-  with \<open>is_unit x\<close> have "is_unit p"
-    by (auto intro: dvd_trans)
-  with \<open>irreducible p\<close> show False
-    by (simp add: irreducible_not_unit)
-qed
-   
-lemma unit_imp_no_prime_divisors:
-  assumes "is_unit x" "prime_elem p"
-  shows   "\<not>p dvd x"
-  using unit_imp_no_irreducible_divisors[OF assms(1) prime_elem_imp_irreducible[OF assms(2)]] .
-
-lemma prime_elem_mono:
-  assumes "prime_elem p" "\<not>q dvd 1" "q dvd p"
-  shows   "prime_elem q"
-proof -
-  from \<open>q dvd p\<close> obtain r where r: "p = q * r" by (elim dvdE)
-  hence "p dvd q * r" by simp
-  with \<open>prime_elem p\<close> have "p dvd q \<or> p dvd r" by (rule prime_elem_dvd_multD)
-  hence "p dvd q"
-  proof
-    assume "p dvd r"
-    then obtain s where s: "r = p * s" by (elim dvdE)
-    from r have "p * 1 = p * (q * s)" by (subst (asm) s) (simp add: mult_ac)
-    with \<open>prime_elem p\<close> have "q dvd 1"
-      by (subst (asm) mult_cancel_left) auto
-    with \<open>\<not>q dvd 1\<close> show ?thesis by contradiction
-  qed
-
-  show ?thesis
-  proof (rule prime_elemI)
-    fix a b assume "q dvd (a * b)"
-    with \<open>p dvd q\<close> have "p dvd (a * b)" by (rule dvd_trans)
-    with \<open>prime_elem p\<close> have "p dvd a \<or> p dvd b" by (rule prime_elem_dvd_multD)
-    with \<open>q dvd p\<close> show "q dvd a \<or> q dvd b" by (blast intro: dvd_trans)
-  qed (insert assms, auto)
-qed
-
-lemma irreducibleD':
-  assumes "irreducible a" "b dvd a"
-  shows   "a dvd b \<or> is_unit b"
-proof -
-  from assms obtain c where c: "a = b * c" by (elim dvdE)
-  from irreducibleD[OF assms(1) this] have "is_unit b \<or> is_unit c" .
-  thus ?thesis by (auto simp: c mult_unit_dvd_iff)
-qed
-
-lemma irreducibleI':
-  assumes "a \<noteq> 0" "\<not>is_unit a" "\<And>b. b dvd a \<Longrightarrow> a dvd b \<or> is_unit b"
-  shows   "irreducible a"
-proof (rule irreducibleI)
-  fix b c assume a_eq: "a = b * c"
-  hence "a dvd b \<or> is_unit b" by (intro assms) simp_all
-  thus "is_unit b \<or> is_unit c"
-  proof
-    assume "a dvd b"
-    hence "b * c dvd b * 1" by (simp add: a_eq)
-    moreover from \<open>a \<noteq> 0\<close> a_eq have "b \<noteq> 0" by auto
-    ultimately show ?thesis by (subst (asm) dvd_times_left_cancel_iff) auto
-  qed blast
-qed (simp_all add: assms(1,2))
-
-lemma irreducible_altdef:
-  "irreducible x \<longleftrightarrow> x \<noteq> 0 \<and> \<not>is_unit x \<and> (\<forall>b. b dvd x \<longrightarrow> x dvd b \<or> is_unit b)"
-  using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto
-
-lemma prime_elem_multD:
-  assumes "prime_elem (a * b)"
-  shows "is_unit a \<or> is_unit b"
-proof -
-  from assms have "a \<noteq> 0" "b \<noteq> 0" by (auto dest!: prime_elem_not_zeroI)
-  moreover from assms prime_elem_dvd_multD [of "a * b"] have "a * b dvd a \<or> a * b dvd b"
-    by auto
-  ultimately show ?thesis
-    using dvd_times_left_cancel_iff [of a b 1]
-      dvd_times_right_cancel_iff [of b a 1]
-    by auto
-qed
-
-lemma prime_elemD2:
-  assumes "prime_elem p" and "a dvd p" and "\<not> is_unit a"
-  shows "p dvd a"
-proof -
-  from \<open>a dvd p\<close> obtain b where "p = a * b" ..
-  with \<open>prime_elem p\<close> prime_elem_multD \<open>\<not> is_unit a\<close> have "is_unit b" by auto
-  with \<open>p = a * b\<close> show ?thesis
-    by (auto simp add: mult_unit_dvd_iff)
-qed
-
-lemma prime_elem_dvd_prod_msetE:
-  assumes "prime_elem p"
-  assumes dvd: "p dvd prod_mset A"
-  obtains a where "a \<in># A" and "p dvd a"
-proof -
-  from dvd have "\<exists>a. a \<in># A \<and> p dvd a"
-  proof (induct A)
-    case empty then show ?case
-    using \<open>prime_elem p\<close> by (simp add: prime_elem_not_unit)
-  next
-    case (add a A)
-    then have "p dvd a * prod_mset A" by simp
-    with \<open>prime_elem p\<close> consider (A) "p dvd prod_mset A" | (B) "p dvd a"
-      by (blast dest: prime_elem_dvd_multD)
-    then show ?case proof cases
-      case B then show ?thesis by auto
-    next
-      case A
-      with add.hyps obtain b where "b \<in># A" "p dvd b"
-        by auto
-      then show ?thesis by auto
-    qed
-  qed
-  with that show thesis by blast
-qed
-
-context
-begin
-
-private lemma prime_elem_powerD:
-  assumes "prime_elem (p ^ n)"
-  shows   "prime_elem p \<and> n = 1"
-proof (cases n)
-  case (Suc m)
-  note assms
-  also from Suc have "p ^ n = p * p^m" by simp
-  finally have "is_unit p \<or> is_unit (p^m)" by (rule prime_elem_multD)
-  moreover from assms have "\<not>is_unit p" by (simp add: prime_elem_def is_unit_power_iff)
-  ultimately have "is_unit (p ^ m)" by simp
-  with \<open>\<not>is_unit p\<close> have "m = 0" by (simp add: is_unit_power_iff)
-  with Suc assms show ?thesis by simp
-qed (insert assms, simp_all)
-
-lemma prime_elem_power_iff:
-  "prime_elem (p ^ n) \<longleftrightarrow> prime_elem p \<and> n = 1"
-  by (auto dest: prime_elem_powerD)
-
-end
-
-lemma irreducible_mult_unit_left:
-  "is_unit a \<Longrightarrow> irreducible (a * p) \<longleftrightarrow> irreducible p"
-  by (auto simp: irreducible_altdef mult.commute[of a] is_unit_mult_iff
-        mult_unit_dvd_iff dvd_mult_unit_iff)
-
-lemma prime_elem_mult_unit_left:
-  "is_unit a \<Longrightarrow> prime_elem (a * p) \<longleftrightarrow> prime_elem p"
-  by (auto simp: prime_elem_def mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff)
-
-lemma prime_elem_dvd_cases:
-  assumes pk: "p*k dvd m*n" and p: "prime_elem p"
-  shows "(\<exists>x. k dvd x*n \<and> m = p*x) \<or> (\<exists>y. k dvd m*y \<and> n = p*y)"
-proof -
-  have "p dvd m*n" using dvd_mult_left pk by blast
-  then consider "p dvd m" | "p dvd n"
-    using p prime_elem_dvd_mult_iff by blast
-  then show ?thesis
-  proof cases
-    case 1 then obtain a where "m = p * a" by (metis dvd_mult_div_cancel) 
-      then have "\<exists>x. k dvd x * n \<and> m = p * x"
-        using p pk by (auto simp: mult.assoc)
-    then show ?thesis ..
-  next
-    case 2 then obtain b where "n = p * b" by (metis dvd_mult_div_cancel) 
-    with p pk have "\<exists>y. k dvd m*y \<and> n = p*y" 
-      by (metis dvd_mult_right dvd_times_left_cancel_iff mult.left_commute mult_zero_left)
-    then show ?thesis ..
-  qed
-qed
-
-lemma prime_elem_power_dvd_prod:
-  assumes pc: "p^c dvd m*n" and p: "prime_elem p"
-  shows "\<exists>a b. a+b = c \<and> p^a dvd m \<and> p^b dvd n"
-using pc
-proof (induct c arbitrary: m n)
-  case 0 show ?case by simp
-next
-  case (Suc c)
-  consider x where "p^c dvd x*n" "m = p*x" | y where "p^c dvd m*y" "n = p*y"
-    using prime_elem_dvd_cases [of _ "p^c", OF _ p] Suc.prems by force
-  then show ?case
-  proof cases
-    case (1 x) 
-    with Suc.hyps[of x n] obtain a b where "a + b = c \<and> p ^ a dvd x \<and> p ^ b dvd n" by blast
-    with 1 have "Suc a + b = Suc c \<and> p ^ Suc a dvd m \<and> p ^ b dvd n"
-      by (auto intro: mult_dvd_mono)
-    thus ?thesis by blast
-  next
-    case (2 y) 
-    with Suc.hyps[of m y] obtain a b where "a + b = c \<and> p ^ a dvd m \<and> p ^ b dvd y" by blast
-    with 2 have "a + Suc b = Suc c \<and> p ^ a dvd m \<and> p ^ Suc b dvd n"
-      by (auto intro: mult_dvd_mono)
-    with Suc.hyps [of m y] show "\<exists>a b. a + b = Suc c \<and> p ^ a dvd m \<and> p ^ b dvd n"
-      by force
-  qed
-qed
-
-lemma prime_elem_power_dvd_cases:
-  assumes "p ^ c dvd m * n" and "a + b = Suc c" and "prime_elem p"
-  shows "p ^ a dvd m \<or> p ^ b dvd n"
-proof -
-  from assms obtain r s
-    where "r + s = c \<and> p ^ r dvd m \<and> p ^ s dvd n"
-    by (blast dest: prime_elem_power_dvd_prod)
-  moreover with assms have
-    "a \<le> r \<or> b \<le> s" by arith
-  ultimately show ?thesis by (auto intro: power_le_dvd)
-qed
-
-lemma prime_elem_not_unit' [simp]:
-  "ASSUMPTION (prime_elem x) \<Longrightarrow> \<not>is_unit x"
-  unfolding ASSUMPTION_def by (rule prime_elem_not_unit)
-
-lemma prime_elem_dvd_power_iff:
-  assumes "prime_elem p"
-  shows "p dvd a ^ n \<longleftrightarrow> p dvd a \<and> n > 0"
-  using assms by (induct n) (auto dest: prime_elem_not_unit prime_elem_dvd_multD)
-
-lemma prime_power_dvd_multD:
-  assumes "prime_elem p"
-  assumes "p ^ n dvd a * b" and "n > 0" and "\<not> p dvd a"
-  shows "p ^ n dvd b"
-  using \<open>p ^ n dvd a * b\<close> and \<open>n > 0\<close> 
-proof (induct n arbitrary: b)
-  case 0 then show ?case by simp
-next
-  case (Suc n) show ?case
-  proof (cases "n = 0")
-    case True with Suc \<open>prime_elem p\<close> \<open>\<not> p dvd a\<close> show ?thesis
-      by (simp add: prime_elem_dvd_mult_iff)
-  next
-    case False then have "n > 0" by simp
-    from \<open>prime_elem p\<close> have "p \<noteq> 0" by auto
-    from Suc.prems have *: "p * p ^ n dvd a * b"
-      by simp
-    then have "p dvd a * b"
-      by (rule dvd_mult_left)
-    with Suc \<open>prime_elem p\<close> \<open>\<not> p dvd a\<close> have "p dvd b"
-      by (simp add: prime_elem_dvd_mult_iff)
-    moreover define c where "c = b div p"
-    ultimately have b: "b = p * c" by simp
-    with * have "p * p ^ n dvd p * (a * c)"
-      by (simp add: ac_simps)
-    with \<open>p \<noteq> 0\<close> have "p ^ n dvd a * c"
-      by simp
-    with Suc.hyps \<open>n > 0\<close> have "p ^ n dvd c"
-      by blast
-    with \<open>p \<noteq> 0\<close> show ?thesis
-      by (simp add: b)
-  qed
-qed
-
-end
-
-
-subsection \<open>Generalized primes: normalized prime elements\<close>
-
-context normalization_semidom
-begin
-
-lemma irreducible_normalized_divisors:
-  assumes "irreducible x" "y dvd x" "normalize y = y"
-  shows   "y = 1 \<or> y = normalize x"
-proof -
-  from assms have "is_unit y \<or> x dvd y" by (auto simp: irreducible_altdef)
-  thus ?thesis
-  proof (elim disjE)
-    assume "is_unit y"
-    hence "normalize y = 1" by (simp add: is_unit_normalize)
-    with assms show ?thesis by simp
-  next
-    assume "x dvd y"
-    with \<open>y dvd x\<close> have "normalize y = normalize x" by (rule associatedI)
-    with assms show ?thesis by simp
-  qed
-qed
-
-lemma irreducible_normalize_iff [simp]: "irreducible (normalize x) = irreducible x"
-  using irreducible_mult_unit_left[of "1 div unit_factor x" x]
-  by (cases "x = 0") (simp_all add: unit_div_commute)
-
-lemma prime_elem_normalize_iff [simp]: "prime_elem (normalize x) = prime_elem x"
-  using prime_elem_mult_unit_left[of "1 div unit_factor x" x]
-  by (cases "x = 0") (simp_all add: unit_div_commute)
-
-lemma prime_elem_associated:
-  assumes "prime_elem p" and "prime_elem q" and "q dvd p"
-  shows "normalize q = normalize p"
-using \<open>q dvd p\<close> proof (rule associatedI)
-  from \<open>prime_elem q\<close> have "\<not> is_unit q"
-    by (auto simp add: prime_elem_not_unit)
-  with \<open>prime_elem p\<close> \<open>q dvd p\<close> show "p dvd q"
-    by (blast intro: prime_elemD2)
-qed
-
-definition prime :: "'a \<Rightarrow> bool" where
-  "prime p \<longleftrightarrow> prime_elem p \<and> normalize p = p"
-
-lemma not_prime_0 [simp]: "\<not>prime 0" by (simp add: prime_def)
-
-lemma not_prime_unit: "is_unit x \<Longrightarrow> \<not>prime x"
-  using prime_elem_not_unit[of x] by (auto simp add: prime_def)
-
-lemma not_prime_1 [simp]: "\<not>prime 1" by (simp add: not_prime_unit)
-
-lemma primeI: "prime_elem x \<Longrightarrow> normalize x = x \<Longrightarrow> prime x"
-  by (simp add: prime_def)
-
-lemma prime_imp_prime_elem [dest]: "prime p \<Longrightarrow> prime_elem p"
-  by (simp add: prime_def)
-
-lemma normalize_prime: "prime p \<Longrightarrow> normalize p = p"
-  by (simp add: prime_def)
-
-lemma prime_normalize_iff [simp]: "prime (normalize p) \<longleftrightarrow> prime_elem p"
-  by (auto simp add: prime_def)
-
-lemma prime_power_iff:
-  "prime (p ^ n) \<longleftrightarrow> prime p \<and> n = 1"
-  by (auto simp: prime_def prime_elem_power_iff)
-
-lemma prime_imp_nonzero [simp]:
-  "ASSUMPTION (prime x) \<Longrightarrow> x \<noteq> 0"
-  unfolding ASSUMPTION_def prime_def by auto
-
-lemma prime_imp_not_one [simp]:
-  "ASSUMPTION (prime x) \<Longrightarrow> x \<noteq> 1"
-  unfolding ASSUMPTION_def by auto
-
-lemma prime_not_unit' [simp]:
-  "ASSUMPTION (prime x) \<Longrightarrow> \<not>is_unit x"
-  unfolding ASSUMPTION_def prime_def by auto
-
-lemma prime_normalize' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> normalize x = x"
-  unfolding ASSUMPTION_def prime_def by simp
-
-lemma unit_factor_prime: "prime x \<Longrightarrow> unit_factor x = 1"
-  using unit_factor_normalize[of x] unfolding prime_def by auto
-
-lemma unit_factor_prime' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> unit_factor x = 1"
-  unfolding ASSUMPTION_def by (rule unit_factor_prime)
-
-lemma prime_imp_prime_elem' [simp]: "ASSUMPTION (prime x) \<Longrightarrow> prime_elem x"
-  by (simp add: prime_def ASSUMPTION_def)
-
-lemma prime_dvd_multD: "prime p \<Longrightarrow> p dvd a * b \<Longrightarrow> p dvd a \<or> p dvd b"
-  by (intro prime_elem_dvd_multD) simp_all
-
-lemma prime_dvd_mult_iff: "prime p \<Longrightarrow> p dvd a * b \<longleftrightarrow> p dvd a \<or> p dvd b"
-  by (auto dest: prime_dvd_multD)
-
-lemma prime_dvd_power: 
-  "prime p \<Longrightarrow> p dvd x ^ n \<Longrightarrow> p dvd x"
-  by (auto dest!: prime_elem_dvd_power simp: prime_def)
-
-lemma prime_dvd_power_iff:
-  "prime p \<Longrightarrow> n > 0 \<Longrightarrow> p dvd x ^ n \<longleftrightarrow> p dvd x"
-  by (subst prime_elem_dvd_power_iff) simp_all
-
-lemma prime_dvd_prod_mset_iff: "prime p \<Longrightarrow> p dvd prod_mset A \<longleftrightarrow> (\<exists>x. x \<in># A \<and> p dvd x)"
-  by (induction A) (simp_all add: prime_elem_dvd_mult_iff prime_imp_prime_elem, blast+)
-
-lemma primes_dvd_imp_eq:
-  assumes "prime p" "prime q" "p dvd q"
-  shows   "p = q"
-proof -
-  from assms have "irreducible q" by (simp add: prime_elem_imp_irreducible prime_def)
-  from irreducibleD'[OF this \<open>p dvd q\<close>] assms have "q dvd p" by simp
-  with \<open>p dvd q\<close> have "normalize p = normalize q" by (rule associatedI)
-  with assms show "p = q" by simp
-qed
-
-lemma prime_dvd_prod_mset_primes_iff:
-  assumes "prime p" "\<And>q. q \<in># A \<Longrightarrow> prime q"
-  shows   "p dvd prod_mset A \<longleftrightarrow> p \<in># A"
-proof -
-  from assms(1) have "p dvd prod_mset A \<longleftrightarrow> (\<exists>x. x \<in># A \<and> p dvd x)" by (rule prime_dvd_prod_mset_iff)
-  also from assms have "\<dots> \<longleftrightarrow> p \<in># A" by (auto dest: primes_dvd_imp_eq)
-  finally show ?thesis .
-qed
-
-lemma prod_mset_primes_dvd_imp_subset:
-  assumes "prod_mset A dvd prod_mset B" "\<And>p. p \<in># A \<Longrightarrow> prime p" "\<And>p. p \<in># B \<Longrightarrow> prime p"
-  shows   "A \<subseteq># B"
-using assms
-proof (induction A arbitrary: B)
-  case empty
-  thus ?case by simp
-next
-  case (add p A B)
-  hence p: "prime p" by simp
-  define B' where "B' = B - {#p#}"
-  from add.prems have "p dvd prod_mset B" by (simp add: dvd_mult_left)
-  with add.prems have "p \<in># B"
-    by (subst (asm) (2) prime_dvd_prod_mset_primes_iff) simp_all
-  hence B: "B = B' + {#p#}" by (simp add: B'_def)
-  from add.prems p have "A \<subseteq># B'" by (intro add.IH) (simp_all add: B)
-  thus ?case by (simp add: B)
-qed
-
-lemma normalize_prod_mset_primes:
-  "(\<And>p. p \<in># A \<Longrightarrow> prime p) \<Longrightarrow> normalize (prod_mset A) = prod_mset A"
-proof (induction A)
-  case (add p A)
-  hence "prime p" by simp
-  hence "normalize p = p" by simp
-  with add show ?case by (simp add: normalize_mult)
-qed simp_all
-
-lemma prod_mset_dvd_prod_mset_primes_iff:
-  assumes "\<And>x. x \<in># A \<Longrightarrow> prime x" "\<And>x. x \<in># B \<Longrightarrow> prime x"
-  shows   "prod_mset A dvd prod_mset B \<longleftrightarrow> A \<subseteq># B"
-  using assms by (auto intro: prod_mset_subset_imp_dvd prod_mset_primes_dvd_imp_subset)
-
-lemma is_unit_prod_mset_primes_iff:
-  assumes "\<And>x. x \<in># A \<Longrightarrow> prime x"
-  shows   "is_unit (prod_mset A) \<longleftrightarrow> A = {#}"
-  by (auto simp add: is_unit_prod_mset_iff)
-    (meson all_not_in_conv assms not_prime_unit set_mset_eq_empty_iff)
-
-lemma prod_mset_primes_irreducible_imp_prime:
-  assumes irred: "irreducible (prod_mset A)"
-  assumes A: "\<And>x. x \<in># A \<Longrightarrow> prime x"
-  assumes B: "\<And>x. x \<in># B \<Longrightarrow> prime x"
-  assumes C: "\<And>x. x \<in># C \<Longrightarrow> prime x"
-  assumes dvd: "prod_mset A dvd prod_mset B * prod_mset C"
-  shows   "prod_mset A dvd prod_mset B \<or> prod_mset A dvd prod_mset C"
-proof -
-  from dvd have "prod_mset A dvd prod_mset (B + C)"
-    by simp
-  with A B C have subset: "A \<subseteq># B + C"
-    by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto
-  define A1 and A2 where "A1 = A \<inter># B" and "A2 = A - A1"
-  have "A = A1 + A2" unfolding A1_def A2_def
-    by (rule sym, intro subset_mset.add_diff_inverse) simp_all
-  from subset have "A1 \<subseteq># B" "A2 \<subseteq># C"
-    by (auto simp: A1_def A2_def Multiset.subset_eq_diff_conv Multiset.union_commute)
-  from \<open>A = A1 + A2\<close> have "prod_mset A = prod_mset A1 * prod_mset A2" by simp
-  from irred and this have "is_unit (prod_mset A1) \<or> is_unit (prod_mset A2)"
-    by (rule irreducibleD)
-  with A have "A1 = {#} \<or> A2 = {#}" unfolding A1_def A2_def
-    by (subst (asm) (1 2) is_unit_prod_mset_primes_iff) (auto dest: Multiset.in_diffD)
-  with dvd \<open>A = A1 + A2\<close> \<open>A1 \<subseteq># B\<close> \<open>A2 \<subseteq># C\<close> show ?thesis
-    by (auto intro: prod_mset_subset_imp_dvd)
-qed
-
-lemma prod_mset_primes_finite_divisor_powers:
-  assumes A: "\<And>x. x \<in># A \<Longrightarrow> prime x"
-  assumes B: "\<And>x. x \<in># B \<Longrightarrow> prime x"
-  assumes "A \<noteq> {#}"
-  shows   "finite {n. prod_mset A ^ n dvd prod_mset B}"
-proof -
-  from \<open>A \<noteq> {#}\<close> obtain x where x: "x \<in># A" by blast
-  define m where "m = count B x"
-  have "{n. prod_mset A ^ n dvd prod_mset B} \<subseteq> {..m}"
-  proof safe
-    fix n assume dvd: "prod_mset A ^ n dvd prod_mset B"
-    from x have "x ^ n dvd prod_mset A ^ n" by (intro dvd_power_same dvd_prod_mset)
-    also note dvd
-    also have "x ^ n = prod_mset (replicate_mset n x)" by simp
-    finally have "replicate_mset n x \<subseteq># B"
-      by (rule prod_mset_primes_dvd_imp_subset) (insert A B x, simp_all split: if_splits)
-    thus "n \<le> m" by (simp add: count_le_replicate_mset_subset_eq m_def)
-  qed
-  moreover have "finite {..m}" by simp
-  ultimately show ?thesis by (rule finite_subset)
-qed
-
-end
-
-
-subsection \<open>In a semiring with GCD, each irreducible element is a prime elements\<close>
-
-context semiring_gcd
-begin
-
-lemma irreducible_imp_prime_elem_gcd:
-  assumes "irreducible x"
-  shows   "prime_elem x"
-proof (rule prime_elemI)
-  fix a b assume "x dvd a * b"
-  from dvd_productE[OF this] obtain y z where yz: "x = y * z" "y dvd a" "z dvd b" .
-  from \<open>irreducible x\<close> and \<open>x = y * z\<close> have "is_unit y \<or> is_unit z" by (rule irreducibleD)
-  with yz show "x dvd a \<or> x dvd b"
-    by (auto simp: mult_unit_dvd_iff mult_unit_dvd_iff')
-qed (insert assms, auto simp: irreducible_not_unit)
-
-lemma prime_elem_imp_coprime:
-  assumes "prime_elem p" "\<not>p dvd n"
-  shows   "coprime p n"
-proof (rule coprimeI)
-  fix d assume "d dvd p" "d dvd n"
-  show "is_unit d"
-  proof (rule ccontr)
-    assume "\<not>is_unit d"
-    from \<open>prime_elem p\<close> and \<open>d dvd p\<close> and this have "p dvd d"
-      by (rule prime_elemD2)
-    from this and \<open>d dvd n\<close> have "p dvd n" by (rule dvd_trans)
-    with \<open>\<not>p dvd n\<close> show False by contradiction
-  qed
-qed
-
-lemma prime_imp_coprime:
-  assumes "prime p" "\<not>p dvd n"
-  shows   "coprime p n"
-  using assms by (simp add: prime_elem_imp_coprime)
-
-lemma prime_elem_imp_power_coprime: 
-  "prime_elem p \<Longrightarrow> \<not>p dvd a \<Longrightarrow> coprime a (p ^ m)"
-  by (auto intro!: coprime_exp dest: prime_elem_imp_coprime simp: gcd.commute)
-
-lemma prime_imp_power_coprime: 
-  "prime p \<Longrightarrow> \<not>p dvd a \<Longrightarrow> coprime a (p ^ m)"
-  by (simp add: prime_elem_imp_power_coprime)
-
-lemma prime_elem_divprod_pow:
-  assumes p: "prime_elem p" and ab: "coprime a b" and pab: "p^n dvd a * b"
-  shows   "p^n dvd a \<or> p^n dvd b"
-  using assms
-proof -
-  from ab p have "\<not>p dvd a \<or> \<not>p dvd b"
-    by (auto simp: coprime prime_elem_def)
-  with p have "coprime (p^n) a \<or> coprime (p^n) b" 
-    by (auto intro: prime_elem_imp_coprime coprime_exp_left)
-  with pab show ?thesis by (auto intro: coprime_dvd_mult simp: mult_ac)
-qed
-
-lemma primes_coprime: 
-  "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
-  using prime_imp_coprime primes_dvd_imp_eq by blast
-
-end
-
-
-subsection \<open>Factorial semirings: algebraic structures with unique prime factorizations\<close>
-
-class factorial_semiring = normalization_semidom +
-  assumes prime_factorization_exists:
-    "x \<noteq> 0 \<Longrightarrow> \<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize x"
-
-text \<open>Alternative characterization\<close>
-  
-lemma (in normalization_semidom) factorial_semiring_altI_aux:
-  assumes finite_divisors: "\<And>x. x \<noteq> 0 \<Longrightarrow> finite {y. y dvd x \<and> normalize y = y}"
-  assumes irreducible_imp_prime_elem: "\<And>x. irreducible x \<Longrightarrow> prime_elem x"
-  assumes "x \<noteq> 0"
-  shows   "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize x"
-using \<open>x \<noteq> 0\<close>
-proof (induction "card {b. b dvd x \<and> normalize b = b}" arbitrary: x rule: less_induct)
-  case (less a)
-  let ?fctrs = "\<lambda>a. {b. b dvd a \<and> normalize b = b}"
-  show ?case
-  proof (cases "is_unit a")
-    case True
-    thus ?thesis by (intro exI[of _ "{#}"]) (auto simp: is_unit_normalize)
-  next
-    case False
-    show ?thesis
-    proof (cases "\<exists>b. b dvd a \<and> \<not>is_unit b \<and> \<not>a dvd b")
-      case False
-      with \<open>\<not>is_unit a\<close> less.prems have "irreducible a" by (auto simp: irreducible_altdef)
-      hence "prime_elem a" by (rule irreducible_imp_prime_elem)
-      thus ?thesis by (intro exI[of _ "{#normalize a#}"]) auto
-    next
-      case True
-      then guess b by (elim exE conjE) note b = this
-
-      from b have "?fctrs b \<subseteq> ?fctrs a" by (auto intro: dvd_trans)
-      moreover from b have "normalize a \<notin> ?fctrs b" "normalize a \<in> ?fctrs a" by simp_all
-      hence "?fctrs b \<noteq> ?fctrs a" by blast
-      ultimately have "?fctrs b \<subset> ?fctrs a" by (subst subset_not_subset_eq) blast
-      with finite_divisors[OF \<open>a \<noteq> 0\<close>] have "card (?fctrs b) < card (?fctrs a)"
-        by (rule psubset_card_mono)
-      moreover from \<open>a \<noteq> 0\<close> b have "b \<noteq> 0" by auto
-      ultimately have "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize b"
-        by (intro less) auto
-      then guess A .. note A = this
-
-      define c where "c = a div b"
-      from b have c: "a = b * c" by (simp add: c_def)
-      from less.prems c have "c \<noteq> 0" by auto
-      from b c have "?fctrs c \<subseteq> ?fctrs a" by (auto intro: dvd_trans)
-      moreover have "normalize a \<notin> ?fctrs c"
-      proof safe
-        assume "normalize a dvd c"
-        hence "b * c dvd 1 * c" by (simp add: c)
-        hence "b dvd 1" by (subst (asm) dvd_times_right_cancel_iff) fact+
-        with b show False by simp
-      qed
-      with \<open>normalize a \<in> ?fctrs a\<close> have "?fctrs a \<noteq> ?fctrs c" by blast
-      ultimately have "?fctrs c \<subset> ?fctrs a" by (subst subset_not_subset_eq) blast
-      with finite_divisors[OF \<open>a \<noteq> 0\<close>] have "card (?fctrs c) < card (?fctrs a)"
-        by (rule psubset_card_mono)
-      with \<open>c \<noteq> 0\<close> have "\<exists>A. (\<forall>x. x \<in># A \<longrightarrow> prime_elem x) \<and> prod_mset A = normalize c"
-        by (intro less) auto
-      then guess B .. note B = this
-
-      from A B show ?thesis by (intro exI[of _ "A + B"]) (auto simp: c normalize_mult)
-    qed
-  qed
-qed 
-
-lemma factorial_semiring_altI:
-  assumes finite_divisors: "\<And>x::'a. x \<noteq> 0 \<Longrightarrow> finite {y. y dvd x \<and> normalize y = y}"
-  assumes irreducible_imp_prime: "\<And>x::'a. irreducible x \<Longrightarrow> prime_elem x"
-  shows   "OFCLASS('a :: normalization_semidom, factorial_semiring_class)"
-  by intro_classes (rule factorial_semiring_altI_aux[OF assms])
-  
-text \<open>Properties\<close>
-
-context factorial_semiring
-begin
-
-lemma prime_factorization_exists':
-  assumes "x \<noteq> 0"
-  obtains A where "\<And>x. x \<in># A \<Longrightarrow> prime x" "prod_mset A = normalize x"
-proof -
-  from prime_factorization_exists[OF assms] obtain A
-    where A: "\<And>x. x \<in># A \<Longrightarrow> prime_elem x" "prod_mset A = normalize x" by blast
-  define A' where "A' = image_mset normalize A"
-  have "prod_mset A' = normalize (prod_mset A)"
-    by (simp add: A'_def normalize_prod_mset)
-  also note A(2)
-  finally have "prod_mset A' = normalize x" by simp
-  moreover from A(1) have "\<forall>x. x \<in># A' \<longrightarrow> prime x" by (auto simp: prime_def A'_def)
-  ultimately show ?thesis by (intro that[of A']) blast
-qed
-
-lemma irreducible_imp_prime_elem:
-  assumes "irreducible x"
-  shows   "prime_elem x"
-proof (rule prime_elemI)
-  fix a b assume dvd: "x dvd a * b"
-  from assms have "x \<noteq> 0" by auto
-  show "x dvd a \<or> x dvd b"
-  proof (cases "a = 0 \<or> b = 0")
-    case False
-    hence "a \<noteq> 0" "b \<noteq> 0" by blast+
-    note nz = \<open>x \<noteq> 0\<close> this
-    from nz[THEN prime_factorization_exists'] guess A B C . note ABC = this
-    from assms ABC have "irreducible (prod_mset A)" by simp
-    from dvd prod_mset_primes_irreducible_imp_prime[of A B C, OF this ABC(1,3,5)] ABC(2,4,6)
-      show ?thesis by (simp add: normalize_mult [symmetric])
-  qed auto
-qed (insert assms, simp_all add: irreducible_def)
-
-lemma finite_divisor_powers:
-  assumes "y \<noteq> 0" "\<not>is_unit x"
-  shows   "finite {n. x ^ n dvd y}"
-proof (cases "x = 0")
-  case True
-  with assms have "{n. x ^ n dvd y} = {0}" by (auto simp: power_0_left)
-  thus ?thesis by simp
-next
-  case False
-  note nz = this \<open>y \<noteq> 0\<close>
-  from nz[THEN prime_factorization_exists'] guess A B . note AB = this
-  from AB assms have "A \<noteq> {#}" by (auto simp: normalize_1_iff)
-  from AB(2,4) prod_mset_primes_finite_divisor_powers [of A B, OF AB(1,3) this]
-    show ?thesis by (simp add: normalize_power [symmetric])
-qed
-
-lemma finite_prime_divisors:
-  assumes "x \<noteq> 0"
-  shows   "finite {p. prime p \<and> p dvd x}"
-proof -
-  from prime_factorization_exists'[OF assms] guess A . note A = this
-  have "{p. prime p \<and> p dvd x} \<subseteq> set_mset A"
-  proof safe
-    fix p assume p: "prime p" and dvd: "p dvd x"
-    from dvd have "p dvd normalize x" by simp
-    also from A have "normalize x = prod_mset A" by simp
-    finally show "p \<in># A" using p A by (subst (asm) prime_dvd_prod_mset_primes_iff)
-  qed
-  moreover have "finite (set_mset A)" by simp
-  ultimately show ?thesis by (rule finite_subset)
-qed
-
-lemma prime_elem_iff_irreducible: "prime_elem x \<longleftrightarrow> irreducible x"
-  by (blast intro: irreducible_imp_prime_elem prime_elem_imp_irreducible)
-
-lemma prime_divisor_exists:
-  assumes "a \<noteq> 0" "\<not>is_unit a"
-  shows   "\<exists>b. b dvd a \<and> prime b"
-proof -
-  from prime_factorization_exists'[OF assms(1)] guess A . note A = this
-  moreover from A and assms have "A \<noteq> {#}" by auto
-  then obtain x where "x \<in># A" by blast
-  with A(1) have *: "x dvd prod_mset A" "prime x" by (auto simp: dvd_prod_mset)
-  with A have "x dvd a" by simp
-  with * show ?thesis by blast
-qed
-
-lemma prime_divisors_induct [case_names zero unit factor]:
-  assumes "P 0" "\<And>x. is_unit x \<Longrightarrow> P x" "\<And>p x. prime p \<Longrightarrow> P x \<Longrightarrow> P (p * x)"
-  shows   "P x"
-proof (cases "x = 0")
-  case False
-  from prime_factorization_exists'[OF this] guess A . note A = this
-  from A(1) have "P (unit_factor x * prod_mset A)"
-  proof (induction A)
-    case (add p A)
-    from add.prems have "prime p" by simp
-    moreover from add.prems have "P (unit_factor x * prod_mset A)" by (intro add.IH) simp_all
-    ultimately have "P (p * (unit_factor x * prod_mset A))" by (rule assms(3))
-    thus ?case by (simp add: mult_ac)
-  qed (simp_all add: assms False)
-  with A show ?thesis by simp
-qed (simp_all add: assms(1))
-
-lemma no_prime_divisors_imp_unit:
-  assumes "a \<noteq> 0" "\<And>b. b dvd a \<Longrightarrow> normalize b = b \<Longrightarrow> \<not> prime_elem b"
-  shows "is_unit a"
-proof (rule ccontr)
-  assume "\<not>is_unit a"
-  from prime_divisor_exists[OF assms(1) this] guess b by (elim exE conjE)
-  with assms(2)[of b] show False by (simp add: prime_def)
-qed
-
-lemma prime_divisorE:
-  assumes "a \<noteq> 0" and "\<not> is_unit a"
-  obtains p where "prime p" and "p dvd a"
-  using assms no_prime_divisors_imp_unit unfolding prime_def by blast
-
-definition multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
-  "multiplicity p x = (if finite {n. p ^ n dvd x} then Max {n. p ^ n dvd x} else 0)"
-
-lemma multiplicity_dvd: "p ^ multiplicity p x dvd x"
-proof (cases "finite {n. p ^ n dvd x}")
-  case True
-  hence "multiplicity p x = Max {n. p ^ n dvd x}"
-    by (simp add: multiplicity_def)
-  also have "\<dots> \<in> {n. p ^ n dvd x}"
-    by (rule Max_in) (auto intro!: True exI[of _ "0::nat"])
-  finally show ?thesis by simp
-qed (simp add: multiplicity_def)
-
-lemma multiplicity_dvd': "n \<le> multiplicity p x \<Longrightarrow> p ^ n dvd x"
-  by (rule dvd_trans[OF le_imp_power_dvd multiplicity_dvd])
-
-context
-  fixes x p :: 'a
-  assumes xp: "x \<noteq> 0" "\<not>is_unit p"
-begin
-
-lemma multiplicity_eq_Max: "multiplicity p x = Max {n. p ^ n dvd x}"
-  using finite_divisor_powers[OF xp] by (simp add: multiplicity_def)
-
-lemma multiplicity_geI:
-  assumes "p ^ n dvd x"
-  shows   "multiplicity p x \<ge> n"
-proof -
-  from assms have "n \<le> Max {n. p ^ n dvd x}"
-    by (intro Max_ge finite_divisor_powers xp) simp_all
-  thus ?thesis by (subst multiplicity_eq_Max)
-qed
-
-lemma multiplicity_lessI:
-  assumes "\<not>p ^ n dvd x"
-  shows   "multiplicity p x < n"
-proof (rule ccontr)
-  assume "\<not>(n > multiplicity p x)"
-  hence "p ^ n dvd x" by (intro multiplicity_dvd') simp
-  with assms show False by contradiction
-qed
-
-lemma power_dvd_iff_le_multiplicity:
-  "p ^ n dvd x \<longleftrightarrow> n \<le> multiplicity p x"
-  using multiplicity_geI[of n] multiplicity_lessI[of n] by (cases "p ^ n dvd x") auto
-
-lemma multiplicity_eq_zero_iff:
-  shows   "multiplicity p x = 0 \<longleftrightarrow> \<not>p dvd x"
-  using power_dvd_iff_le_multiplicity[of 1] by auto
-
-lemma multiplicity_gt_zero_iff:
-  shows   "multiplicity p x > 0 \<longleftrightarrow> p dvd x"
-  using power_dvd_iff_le_multiplicity[of 1] by auto
-
-lemma multiplicity_decompose:
-  "\<not>p dvd (x div p ^ multiplicity p x)"
-proof
-  assume *: "p dvd x div p ^ multiplicity p x"
-  have "x = x div p ^ multiplicity p x * (p ^ multiplicity p x)"
-    using multiplicity_dvd[of p x] by simp
-  also from * have "x div p ^ multiplicity p x = (x div p ^ multiplicity p x div p) * p" by simp
-  also have "x div p ^ multiplicity p x div p * p * p ^ multiplicity p x =
-               x div p ^ multiplicity p x div p * p ^ Suc (multiplicity p x)"
-    by (simp add: mult_assoc)
-  also have "p ^ Suc (multiplicity p x) dvd \<dots>" by (rule dvd_triv_right)
-  finally show False by (subst (asm) power_dvd_iff_le_multiplicity) simp
-qed
-
-lemma multiplicity_decompose':
-  obtains y where "x = p ^ multiplicity p x * y" "\<not>p dvd y"
-  using that[of "x div p ^ multiplicity p x"]
-  by (simp add: multiplicity_decompose multiplicity_dvd)
-
-end
-
-lemma multiplicity_zero [simp]: "multiplicity p 0 = 0"
-  by (simp add: multiplicity_def)
-
-lemma prime_elem_multiplicity_eq_zero_iff:
-  "prime_elem p \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> multiplicity p x = 0 \<longleftrightarrow> \<not>p dvd x"
-  by (rule multiplicity_eq_zero_iff) simp_all
-
-lemma prime_multiplicity_other:
-  assumes "prime p" "prime q" "p \<noteq> q"
-  shows   "multiplicity p q = 0"
-  using assms by (subst prime_elem_multiplicity_eq_zero_iff) (auto dest: primes_dvd_imp_eq)  
-
-lemma prime_multiplicity_gt_zero_iff:
-  "prime_elem p \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> multiplicity p x > 0 \<longleftrightarrow> p dvd x"
-  by (rule multiplicity_gt_zero_iff) simp_all
-
-lemma multiplicity_unit_left: "is_unit p \<Longrightarrow> multiplicity p x = 0"
-  by (simp add: multiplicity_def is_unit_power_iff unit_imp_dvd)
-
-lemma multiplicity_unit_right:
-  assumes "is_unit x"
-  shows   "multiplicity p x = 0"
-proof (cases "is_unit p \<or> x = 0")
-  case False
-  with multiplicity_lessI[of x p 1] this assms
-    show ?thesis by (auto dest: dvd_unit_imp_unit)
-qed (auto simp: multiplicity_unit_left)
-
-lemma multiplicity_one [simp]: "multiplicity p 1 = 0"
-  by (rule multiplicity_unit_right) simp_all
-
-lemma multiplicity_eqI:
-  assumes "p ^ n dvd x" "\<not>p ^ Suc n dvd x"
-  shows   "multiplicity p x = n"
-proof -
-  consider "x = 0" | "is_unit p" | "x \<noteq> 0" "\<not>is_unit p" by blast
-  thus ?thesis
-  proof cases
-    assume xp: "x \<noteq> 0" "\<not>is_unit p"
-    from xp assms(1) have "multiplicity p x \<ge> n" by (intro multiplicity_geI)
-    moreover from assms(2) xp have "multiplicity p x < Suc n" by (intro multiplicity_lessI)
-    ultimately show ?thesis by simp
-  next
-    assume "is_unit p"
-    hence "is_unit (p ^ Suc n)" by (simp add: is_unit_power_iff del: power_Suc)
-    hence "p ^ Suc n dvd x" by (rule unit_imp_dvd)
-    with \<open>\<not>p ^ Suc n dvd x\<close> show ?thesis by contradiction
-  qed (insert assms, simp_all)
-qed
-
-
-context
-  fixes x p :: 'a
-  assumes xp: "x \<noteq> 0" "\<not>is_unit p"
-begin
-
-lemma multiplicity_times_same:
-  assumes "p \<noteq> 0"
-  shows   "multiplicity p (p * x) = Suc (multiplicity p x)"
-proof (rule multiplicity_eqI)
-  show "p ^ Suc (multiplicity p x) dvd p * x"
-    by (auto intro!: mult_dvd_mono multiplicity_dvd)
-  from xp assms show "\<not> p ^ Suc (Suc (multiplicity p x)) dvd p * x"
-    using power_dvd_iff_le_multiplicity[OF xp, of "Suc (multiplicity p x)"] by simp
-qed
-
-end
-
-lemma multiplicity_same_power': "multiplicity p (p ^ n) = (if p = 0 \<or> is_unit p then 0 else n)"
-proof -
-  consider "p = 0" | "is_unit p" |"p \<noteq> 0" "\<not>is_unit p" by blast
-  thus ?thesis
-  proof cases
-    assume "p \<noteq> 0" "\<not>is_unit p"
-    thus ?thesis by (induction n) (simp_all add: multiplicity_times_same)
-  qed (simp_all add: power_0_left multiplicity_unit_left)
-qed
-
-lemma multiplicity_same_power:
-  "p \<noteq> 0 \<Longrightarrow> \<not>is_unit p \<Longrightarrow> multiplicity p (p ^ n) = n"
-  by (simp add: multiplicity_same_power')
-
-lemma multiplicity_prime_elem_times_other:
-  assumes "prime_elem p" "\<not>p dvd q"
-  shows   "multiplicity p (q * x) = multiplicity p x"
-proof (cases "x = 0")
-  case False
-  show ?thesis
-  proof (rule multiplicity_eqI)
-    have "1 * p ^ multiplicity p x dvd q * x"
-      by (intro mult_dvd_mono multiplicity_dvd) simp_all
-    thus "p ^ multiplicity p x dvd q * x" by simp
-  next
-    define n where "n = multiplicity p x"
-    from assms have "\<not>is_unit p" by simp
-    from multiplicity_decompose'[OF False this] guess y . note y = this [folded n_def]
-    from y have "p ^ Suc n dvd q * x \<longleftrightarrow> p ^ n * p dvd p ^ n * (q * y)" by (simp add: mult_ac)
-    also from assms have "\<dots> \<longleftrightarrow> p dvd q * y" by simp
-    also have "\<dots> \<longleftrightarrow> p dvd q \<or> p dvd y" by (rule prime_elem_dvd_mult_iff) fact+
-    also from assms y have "\<dots> \<longleftrightarrow> False" by simp
-    finally show "\<not>(p ^ Suc n dvd q * x)" by blast
-  qed
-qed simp_all
-
-lemma multiplicity_self:
-  assumes "p \<noteq> 0" "\<not>is_unit p"
-  shows   "multiplicity p p = 1"
-proof -
-  from assms have "multiplicity p p = Max {n. p ^ n dvd p}"
-    by (simp add: multiplicity_eq_Max)
-  also from assms have "p ^ n dvd p \<longleftrightarrow> n \<le> 1" for n
-    using dvd_power_iff[of p n 1] by auto
-  hence "{n. p ^ n dvd p} = {..1}" by auto
-  also have "\<dots> = {0,1}" by auto
-  finally show ?thesis by simp
-qed
-
-lemma multiplicity_times_unit_left:
-  assumes "is_unit c"
-  shows   "multiplicity (c * p) x = multiplicity p x"
-proof -
-  from assms have "{n. (c * p) ^ n dvd x} = {n. p ^ n dvd x}"
-    by (subst mult.commute) (simp add: mult_unit_dvd_iff power_mult_distrib is_unit_power_iff)
-  thus ?thesis by (simp add: multiplicity_def)
-qed
-
-lemma multiplicity_times_unit_right:
-  assumes "is_unit c"
-  shows   "multiplicity p (c * x) = multiplicity p x"
-proof -
-  from assms have "{n. p ^ n dvd c * x} = {n. p ^ n dvd x}"
-    by (subst mult.commute) (simp add: dvd_mult_unit_iff)
-  thus ?thesis by (simp add: multiplicity_def)
-qed
-
-lemma multiplicity_normalize_left [simp]:
-  "multiplicity (normalize p) x = multiplicity p x"
-proof (cases "p = 0")
-  case [simp]: False
-  have "normalize p = (1 div unit_factor p) * p"
-    by (simp add: unit_div_commute is_unit_unit_factor)
-  also have "multiplicity \<dots> x = multiplicity p x"
-    by (rule multiplicity_times_unit_left) (simp add: is_unit_unit_factor)
-  finally show ?thesis .
-qed simp_all
-
-lemma multiplicity_normalize_right [simp]:
-  "multiplicity p (normalize x) = multiplicity p x"
-proof (cases "x = 0")
-  case [simp]: False
-  have "normalize x = (1 div unit_factor x) * x"
-    by (simp add: unit_div_commute is_unit_unit_factor)
-  also have "multiplicity p \<dots> = multiplicity p x"
-    by (rule multiplicity_times_unit_right) (simp add: is_unit_unit_factor)
-  finally show ?thesis .
-qed simp_all   
-
-lemma multiplicity_prime [simp]: "prime_elem p \<Longrightarrow> multiplicity p p = 1"
-  by (rule multiplicity_self) auto
-
-lemma multiplicity_prime_power [simp]: "prime_elem p \<Longrightarrow> multiplicity p (p ^ n) = n"
-  by (subst multiplicity_same_power') auto
-
-lift_definition prime_factorization :: "'a \<Rightarrow> 'a multiset" is
-  "\<lambda>x p. if prime p then multiplicity p x else 0"
-  unfolding multiset_def
-proof clarify
-  fix x :: 'a
-  show "finite {p. 0 < (if prime p then multiplicity p x else 0)}" (is "finite ?A")
-  proof (cases "x = 0")
-    case False
-    from False have "?A \<subseteq> {p. prime p \<and> p dvd x}"
-      by (auto simp: multiplicity_gt_zero_iff)
-    moreover from False have "finite {p. prime p \<and> p dvd x}"
-      by (rule finite_prime_divisors)
-    ultimately show ?thesis by (rule finite_subset)
-  qed simp_all
-qed
-
-abbreviation prime_factors :: "'a \<Rightarrow> 'a set" where
-  "prime_factors a \<equiv> set_mset (prime_factorization a)"
-
-lemma count_prime_factorization_nonprime:
-  "\<not>prime p \<Longrightarrow> count (prime_factorization x) p = 0"
-  by transfer simp
-
-lemma count_prime_factorization_prime:
-  "prime p \<Longrightarrow> count (prime_factorization x) p = multiplicity p x"
-  by transfer simp
-
-lemma count_prime_factorization:
-  "count (prime_factorization x) p = (if prime p then multiplicity p x else 0)"
-  by transfer simp
-
-lemma dvd_imp_multiplicity_le:
-  assumes "a dvd b" "b \<noteq> 0"
-  shows   "multiplicity p a \<le> multiplicity p b"
-proof (cases "is_unit p")
-  case False
-  with assms show ?thesis
-    by (intro multiplicity_geI ) (auto intro: dvd_trans[OF multiplicity_dvd' assms(1)])
-qed (insert assms, auto simp: multiplicity_unit_left)
-
-lemma prime_factorization_0 [simp]: "prime_factorization 0 = {#}"
-  by (simp add: multiset_eq_iff count_prime_factorization)
-
-lemma prime_factorization_empty_iff:
-  "prime_factorization x = {#} \<longleftrightarrow> x = 0 \<or> is_unit x"
-proof
-  assume *: "prime_factorization x = {#}"
-  {
-    assume x: "x \<noteq> 0" "\<not>is_unit x"
-    {
-      fix p assume p: "prime p"
-      have "count (prime_factorization x) p = 0" by (simp add: *)
-      also from p have "count (prime_factorization x) p = multiplicity p x"
-        by (rule count_prime_factorization_prime)
-      also from x p have "\<dots> = 0 \<longleftrightarrow> \<not>p dvd x" by (simp add: multiplicity_eq_zero_iff)
-      finally have "\<not>p dvd x" .
-    }
-    with prime_divisor_exists[OF x] have False by blast
-  }
-  thus "x = 0 \<or> is_unit x" by blast
-next
-  assume "x = 0 \<or> is_unit x"
-  thus "prime_factorization x = {#}"
-  proof
-    assume x: "is_unit x"
-    {
-      fix p assume p: "prime p"
-      from p x have "multiplicity p x = 0"
-        by (subst multiplicity_eq_zero_iff)
-           (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors)
-    }
-    thus ?thesis by (simp add: multiset_eq_iff count_prime_factorization)
-  qed simp_all
-qed
-
-lemma prime_factorization_unit:
-  assumes "is_unit x"
-  shows   "prime_factorization x = {#}"
-proof (rule multiset_eqI)
-  fix p :: 'a
-  show "count (prime_factorization x) p = count {#} p"
-  proof (cases "prime p")
-    case True
-    with assms have "multiplicity p x = 0"
-      by (subst multiplicity_eq_zero_iff)
-         (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors)
-    with True show ?thesis by (simp add: count_prime_factorization_prime)
-  qed (simp_all add: count_prime_factorization_nonprime)
-qed
-
-lemma prime_factorization_1 [simp]: "prime_factorization 1 = {#}"
-  by (simp add: prime_factorization_unit)
-
-lemma prime_factorization_times_prime:
-  assumes "x \<noteq> 0" "prime p"
-  shows   "prime_factorization (p * x) = {#p#} + prime_factorization x"
-proof (rule multiset_eqI)
-  fix q :: 'a
-  consider "\<not>prime q" | "p = q" | "prime q" "p \<noteq> q" by blast
-  thus "count (prime_factorization (p * x)) q = count ({#p#} + prime_factorization x) q"
-  proof cases
-    assume q: "prime q" "p \<noteq> q"
-    with assms primes_dvd_imp_eq[of q p] have "\<not>q dvd p" by auto
-    with q assms show ?thesis
-      by (simp add: multiplicity_prime_elem_times_other count_prime_factorization)
-  qed (insert assms, auto simp: count_prime_factorization multiplicity_times_same)
-qed
-
-lemma prod_mset_prime_factorization:
-  assumes "x \<noteq> 0"
-  shows   "prod_mset (prime_factorization x) = normalize x"
-  using assms
-  by (induction x rule: prime_divisors_induct)
-     (simp_all add: prime_factorization_unit prime_factorization_times_prime
-                    is_unit_normalize normalize_mult)
-
-lemma in_prime_factors_iff:
-  "p \<in> prime_factors x \<longleftrightarrow> x \<noteq> 0 \<and> p dvd x \<and> prime p"
-proof -
-  have "p \<in> prime_factors x \<longleftrightarrow> count (prime_factorization x) p > 0" by simp
-  also have "\<dots> \<longleftrightarrow> x \<noteq> 0 \<and> p dvd x \<and> prime p"
-   by (subst count_prime_factorization, cases "x = 0")
-      (auto simp: multiplicity_eq_zero_iff multiplicity_gt_zero_iff)
-  finally show ?thesis .
-qed
-
-lemma in_prime_factors_imp_prime [intro]:
-  "p \<in> prime_factors x \<Longrightarrow> prime p"
-  by (simp add: in_prime_factors_iff)
-
-lemma in_prime_factors_imp_dvd [dest]:
-  "p \<in> prime_factors x \<Longrightarrow> p dvd x"
-  by (simp add: in_prime_factors_iff)
-
-lemma prime_factorsI:
-  "x \<noteq> 0 \<Longrightarrow> prime p \<Longrightarrow> p dvd x \<Longrightarrow> p \<in> prime_factors x"
-  by (auto simp: in_prime_factors_iff)
-
-lemma prime_factors_dvd:
-  "x \<noteq> 0 \<Longrightarrow> prime_factors x = {p. prime p \<and> p dvd x}"
-  by (auto intro: prime_factorsI)
-
-lemma prime_factors_multiplicity:
-  "prime_factors n = {p. prime p \<and> multiplicity p n > 0}"
-  by (cases "n = 0") (auto simp add: prime_factors_dvd prime_multiplicity_gt_zero_iff)
-
-lemma prime_factorization_prime:
-  assumes "prime p"
-  shows   "prime_factorization p = {#p#}"
-proof (rule multiset_eqI)
-  fix q :: 'a
-  consider "\<not>prime q" | "q = p" | "prime q" "q \<noteq> p" by blast
-  thus "count (prime_factorization p) q = count {#p#} q"
-    by cases (insert assms, auto dest: primes_dvd_imp_eq
-                simp: count_prime_factorization multiplicity_self multiplicity_eq_zero_iff)
-qed
-
-lemma prime_factorization_prod_mset_primes:
-  assumes "\<And>p. p \<in># A \<Longrightarrow> prime p"
-  shows   "prime_factorization (prod_mset A) = A"
-  using assms
-proof (induction A)
-  case (add p A)
-  from add.prems[of 0] have "0 \<notin># A" by auto
-  hence "prod_mset A \<noteq> 0" by auto
-  with add show ?case
-    by (simp_all add: mult_ac prime_factorization_times_prime Multiset.union_commute)
-qed simp_all
-
-lemma prime_factorization_cong:
-  "normalize x = normalize y \<Longrightarrow> prime_factorization x = prime_factorization y"
-  by (simp add: multiset_eq_iff count_prime_factorization
-                multiplicity_normalize_right [of _ x, symmetric]
-                multiplicity_normalize_right [of _ y, symmetric]
-           del:  multiplicity_normalize_right)
-
-lemma prime_factorization_unique:
-  assumes "x \<noteq> 0" "y \<noteq> 0"
-  shows   "prime_factorization x = prime_factorization y \<longleftrightarrow> normalize x = normalize y"
-proof
-  assume "prime_factorization x = prime_factorization y"
-  hence "prod_mset (prime_factorization x) = prod_mset (prime_factorization y)" by simp
-  with assms show "normalize x = normalize y" by (simp add: prod_mset_prime_factorization)
-qed (rule prime_factorization_cong)
-
-lemma prime_factorization_mult:
-  assumes "x \<noteq> 0" "y \<noteq> 0"
-  shows   "prime_factorization (x * y) = prime_factorization x + prime_factorization y"
-proof -
-  have "prime_factorization (prod_mset (prime_factorization (x * y))) =
-          prime_factorization (prod_mset (prime_factorization x + prime_factorization y))"
-    by (simp add: prod_mset_prime_factorization assms normalize_mult)
-  also have "prime_factorization (prod_mset (prime_factorization (x * y))) =
-               prime_factorization (x * y)"
-    by (rule prime_factorization_prod_mset_primes) (simp_all add: in_prime_factors_imp_prime)
-  also have "prime_factorization (prod_mset (prime_factorization x + prime_factorization y)) =
-               prime_factorization x + prime_factorization y"
-    by (rule prime_factorization_prod_mset_primes) (auto simp: in_prime_factors_imp_prime)
-  finally show ?thesis .
-qed
-
-lemma prime_elem_multiplicity_mult_distrib:
-  assumes "prime_elem p" "x \<noteq> 0" "y \<noteq> 0"
-  shows   "multiplicity p (x * y) = multiplicity p x + multiplicity p y"
-proof -
-  have "multiplicity p (x * y) = count (prime_factorization (x * y)) (normalize p)"
-    by (subst count_prime_factorization_prime) (simp_all add: assms)
-  also from assms 
-    have "prime_factorization (x * y) = prime_factorization x + prime_factorization y"
-      by (intro prime_factorization_mult)
-  also have "count \<dots> (normalize p) = 
-    count (prime_factorization x) (normalize p) + count (prime_factorization y) (normalize p)"
-    by simp
-  also have "\<dots> = multiplicity p x + multiplicity p y" 
-    by (subst (1 2) count_prime_factorization_prime) (simp_all add: assms)
-  finally show ?thesis .
-qed
-
-lemma prime_elem_multiplicity_prod_mset_distrib:
-  assumes "prime_elem p" "0 \<notin># A"
-  shows   "multiplicity p (prod_mset A) = sum_mset (image_mset (multiplicity p) A)"
-  using assms by (induction A) (auto simp: prime_elem_multiplicity_mult_distrib)
-
-lemma prime_elem_multiplicity_power_distrib:
-  assumes "prime_elem p" "x \<noteq> 0"
-  shows   "multiplicity p (x ^ n) = n * multiplicity p x"
-  using assms prime_elem_multiplicity_prod_mset_distrib [of p "replicate_mset n x"]
-  by simp
-
-lemma prime_elem_multiplicity_prod_distrib:
-  assumes "prime_elem p" "0 \<notin> f ` A" "finite A"
-  shows   "multiplicity p (prod f A) = (\<Sum>x\<in>A. multiplicity p (f x))"
-proof -
-  have "multiplicity p (prod f A) = (\<Sum>x\<in>#mset_set A. multiplicity p (f x))"
-    using assms by (subst prod_unfold_prod_mset)
-                   (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset 
-                      multiset.map_comp o_def)
-  also from \<open>finite A\<close> have "\<dots> = (\<Sum>x\<in>A. multiplicity p (f x))"
-    by (induction A rule: finite_induct) simp_all
-  finally show ?thesis .
-qed
-
-lemma multiplicity_distinct_prime_power:
-  "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> multiplicity p (q ^ n) = 0"
-  by (subst prime_elem_multiplicity_power_distrib) (auto simp: prime_multiplicity_other)
-
-lemma prime_factorization_prime_power:
-  "prime p \<Longrightarrow> prime_factorization (p ^ n) = replicate_mset n p"
-  by (induction n)
-     (simp_all add: prime_factorization_mult prime_factorization_prime Multiset.union_commute)
-
-lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x"
-  by (cases "x = 0") (simp_all add: prod_mset_prime_factorization)
-
-lemma prime_factorization_subset_iff_dvd:
-  assumes [simp]: "x \<noteq> 0" "y \<noteq> 0"
-  shows   "prime_factorization x \<subseteq># prime_factorization y \<longleftrightarrow> x dvd y"
-proof -
-  have "x dvd y \<longleftrightarrow> prod_mset (prime_factorization x) dvd prod_mset (prime_factorization y)"
-    by (simp add: prod_mset_prime_factorization)
-  also have "\<dots> \<longleftrightarrow> prime_factorization x \<subseteq># prime_factorization y"
-    by (auto intro!: prod_mset_primes_dvd_imp_subset prod_mset_subset_imp_dvd)
-  finally show ?thesis ..
-qed
-
-lemma prime_factorization_subset_imp_dvd: 
-  "x \<noteq> 0 \<Longrightarrow> (prime_factorization x \<subseteq># prime_factorization y) \<Longrightarrow> x dvd y"
-  by (cases "y = 0") (simp_all add: prime_factorization_subset_iff_dvd)
-
-lemma prime_factorization_divide:
-  assumes "b dvd a"
-  shows   "prime_factorization (a div b) = prime_factorization a - prime_factorization b"
-proof (cases "a = 0")
-  case [simp]: False
-  from assms have [simp]: "b \<noteq> 0" by auto
-  have "prime_factorization ((a div b) * b) = prime_factorization (a div b) + prime_factorization b"
-    by (intro prime_factorization_mult) (insert assms, auto elim!: dvdE)
-  with assms show ?thesis by simp
-qed simp_all
-
-lemma zero_not_in_prime_factors [simp]: "0 \<notin> prime_factors x"
-  by (auto dest: in_prime_factors_imp_prime)
-
-lemma prime_prime_factors:
-  "prime p \<Longrightarrow> prime_factors p = {p}"
-  by (drule prime_factorization_prime) simp
-
-lemma prod_prime_factors:
-  assumes "x \<noteq> 0"
-  shows   "(\<Prod>p \<in> prime_factors x. p ^ multiplicity p x) = normalize x"
-proof -
-  have "normalize x = prod_mset (prime_factorization x)"
-    by (simp add: prod_mset_prime_factorization assms)
-  also have "\<dots> = (\<Prod>p \<in> prime_factors x. p ^ count (prime_factorization x) p)"
-    by (subst prod_mset_multiplicity) simp_all
-  also have "\<dots> = (\<Prod>p \<in> prime_factors x. p ^ multiplicity p x)"
-    by (intro prod.cong) 
-      (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime)
-  finally show ?thesis ..
-qed
-
-lemma prime_factorization_unique'':
-  assumes S_eq: "S = {p. 0 < f p}"
-    and "finite S"
-    and S: "\<forall>p\<in>S. prime p" "normalize n = (\<Prod>p\<in>S. p ^ f p)"
-  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
-proof
-  define A where "A = Abs_multiset f"
-  from \<open>finite S\<close> S(1) have "(\<Prod>p\<in>S. p ^ f p) \<noteq> 0" by auto
-  with S(2) have nz: "n \<noteq> 0" by auto
-  from S_eq \<open>finite S\<close> have count_A: "count A x = f x" for x
-    unfolding A_def by (subst multiset.Abs_multiset_inverse) (simp_all add: multiset_def)
-  from S_eq count_A have set_mset_A: "set_mset A = S"
-    by (simp only: set_mset_def)
-  from S(2) have "normalize n = (\<Prod>p\<in>S. p ^ f p)" .
-  also have "\<dots> = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A)
-  also from nz have "normalize n = prod_mset (prime_factorization n)" 
-    by (simp add: prod_mset_prime_factorization)
-  finally have "prime_factorization (prod_mset A) = 
-                  prime_factorization (prod_mset (prime_factorization n))" by simp
-  also from S(1) have "prime_factorization (prod_mset A) = A"
-    by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A)
-  also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n"
-    by (intro prime_factorization_prod_mset_primes) auto
-  finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric])
-  
-  show "(\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
-  proof safe
-    fix p :: 'a assume p: "prime p"
-    have "multiplicity p n = multiplicity p (normalize n)" by simp
-    also have "normalize n = prod_mset A" 
-      by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S)
-    also from p set_mset_A S(1) 
-    have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
-      by (intro prime_elem_multiplicity_prod_mset_distrib) auto
-    also from S(1) p
-    have "image_mset (multiplicity p) A = image_mset (\<lambda>q. if p = q then 1 else 0) A"
-      by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other)
-    also have "sum_mset \<dots> = f p" by (simp add: sum_mset_delta' count_A)
-    finally show "f p = multiplicity p n" ..
-  qed
-qed
-
-lemma prime_factors_product: 
-  "x \<noteq> 0 \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> prime_factors (x * y) = prime_factors x \<union> prime_factors y"
-  by (simp add: prime_factorization_mult)
-
-lemma dvd_prime_factors [intro]:
-  "y \<noteq> 0 \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x \<subseteq> prime_factors y"
-  by (intro set_mset_mono, subst prime_factorization_subset_iff_dvd) auto
-
-(* RENAMED multiplicity_dvd *)
-lemma multiplicity_le_imp_dvd:
-  assumes "x \<noteq> 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x \<le> multiplicity p y"
-  shows   "x dvd y"
-proof (cases "y = 0")
-  case False
-  from assms this have "prime_factorization x \<subseteq># prime_factorization y"
-    by (intro mset_subset_eqI) (auto simp: count_prime_factorization)
-  with assms False show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd)
-qed auto
-
-lemma dvd_multiplicity_eq:
-  "x \<noteq> 0 \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> x dvd y \<longleftrightarrow> (\<forall>p. multiplicity p x \<le> multiplicity p y)"
-  by (auto intro: dvd_imp_multiplicity_le multiplicity_le_imp_dvd)
-
-lemma multiplicity_eq_imp_eq:
-  assumes "x \<noteq> 0" "y \<noteq> 0"
-  assumes "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
-  shows   "normalize x = normalize y"
-  using assms by (intro associatedI multiplicity_le_imp_dvd) simp_all
-
-lemma prime_factorization_unique':
-  assumes "\<forall>p \<in># M. prime p" "\<forall>p \<in># N. prime p" "(\<Prod>i \<in># M. i) = (\<Prod>i \<in># N. i)"
-  shows   "M = N"
-proof -
-  have "prime_factorization (\<Prod>i \<in># M. i) = prime_factorization (\<Prod>i \<in># N. i)"
-    by (simp only: assms)
-  also from assms have "prime_factorization (\<Prod>i \<in># M. i) = M"
-    by (subst prime_factorization_prod_mset_primes) simp_all
-  also from assms have "prime_factorization (\<Prod>i \<in># N. i) = N"
-    by (subst prime_factorization_prod_mset_primes) simp_all
-  finally show ?thesis .
-qed
-
-lemma multiplicity_cong:
-  "(\<And>r. p ^ r dvd a \<longleftrightarrow> p ^ r dvd b) \<Longrightarrow> multiplicity p a = multiplicity p b"
-  by (simp add: multiplicity_def)
-
-lemma not_dvd_imp_multiplicity_0: 
-  assumes "\<not>p dvd x"
-  shows   "multiplicity p x = 0"
-proof -
-  from assms have "multiplicity p x < 1"
-    by (intro multiplicity_lessI) auto
-  thus ?thesis by simp
-qed
-
-
-subsection \<open>GCD and LCM computation with unique factorizations\<close>
-
-definition "gcd_factorial a b = (if a = 0 then normalize b
-     else if b = 0 then normalize a
-     else prod_mset (prime_factorization a \<inter># prime_factorization b))"
-
-definition "lcm_factorial a b = (if a = 0 \<or> b = 0 then 0
-     else prod_mset (prime_factorization a \<union># prime_factorization b))"
-
-definition "Gcd_factorial A =
-  (if A \<subseteq> {0} then 0 else prod_mset (Inf (prime_factorization ` (A - {0}))))"
-
-definition "Lcm_factorial A =
-  (if A = {} then 1
-   else if 0 \<notin> A \<and> subset_mset.bdd_above (prime_factorization ` (A - {0})) then
-     prod_mset (Sup (prime_factorization ` A))
-   else
-     0)"
-
-lemma prime_factorization_gcd_factorial:
-  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (gcd_factorial a b) = prime_factorization a \<inter># prime_factorization b"
-proof -
-  have "prime_factorization (gcd_factorial a b) =
-          prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
-    by (simp add: gcd_factorial_def)
-  also have "\<dots> = prime_factorization a \<inter># prime_factorization b"
-    by (subst prime_factorization_prod_mset_primes) auto
-  finally show ?thesis .
-qed
-
-lemma prime_factorization_lcm_factorial:
-  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (lcm_factorial a b) = prime_factorization a \<union># prime_factorization b"
-proof -
-  have "prime_factorization (lcm_factorial a b) =
-          prime_factorization (prod_mset (prime_factorization a \<union># prime_factorization b))"
-    by (simp add: lcm_factorial_def)
-  also have "\<dots> = prime_factorization a \<union># prime_factorization b"
-    by (subst prime_factorization_prod_mset_primes) auto
-  finally show ?thesis .
-qed
-
-lemma prime_factorization_Gcd_factorial:
-  assumes "\<not>A \<subseteq> {0}"
-  shows   "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))"
-proof -
-  from assms obtain x where x: "x \<in> A - {0}" by auto
-  hence "Inf (prime_factorization ` (A - {0})) \<subseteq># prime_factorization x"
-    by (intro subset_mset.cInf_lower) simp_all
-  hence "\<forall>y. y \<in># Inf (prime_factorization ` (A - {0})) \<longrightarrow> y \<in> prime_factors x"
-    by (auto dest: mset_subset_eqD)
-  with in_prime_factors_imp_prime[of _ x]
-    have "\<forall>p. p \<in># Inf (prime_factorization ` (A - {0})) \<longrightarrow> prime p" by blast
-  with assms show ?thesis
-    by (simp add: Gcd_factorial_def prime_factorization_prod_mset_primes)
-qed
-
-lemma prime_factorization_Lcm_factorial:
-  assumes "0 \<notin> A" "subset_mset.bdd_above (prime_factorization ` A)"
-  shows   "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)"
-proof (cases "A = {}")
-  case True
-  hence "prime_factorization ` A = {}" by auto
-  also have "Sup \<dots> = {#}" by (simp add: Sup_multiset_empty)
-  finally show ?thesis by (simp add: Lcm_factorial_def)
-next
-  case False
-  have "\<forall>y. y \<in># Sup (prime_factorization ` A) \<longrightarrow> prime y"
-    by (auto simp: in_Sup_multiset_iff assms)
-  with assms False show ?thesis
-    by (simp add: Lcm_factorial_def prime_factorization_prod_mset_primes)
-qed
-
-lemma gcd_factorial_commute: "gcd_factorial a b = gcd_factorial b a"
-  by (simp add: gcd_factorial_def multiset_inter_commute)
-
-lemma gcd_factorial_dvd1: "gcd_factorial a b dvd a"
-proof (cases "a = 0 \<or> b = 0")
-  case False
-  hence "gcd_factorial a b \<noteq> 0" by (auto simp: gcd_factorial_def)
-  with False show ?thesis
-    by (subst prime_factorization_subset_iff_dvd [symmetric])
-       (auto simp: prime_factorization_gcd_factorial)
-qed (auto simp: gcd_factorial_def)
-
-lemma gcd_factorial_dvd2: "gcd_factorial a b dvd b"
-  by (subst gcd_factorial_commute) (rule gcd_factorial_dvd1)
-
-lemma normalize_gcd_factorial: "normalize (gcd_factorial a b) = gcd_factorial a b"
-proof -
-  have "normalize (prod_mset (prime_factorization a \<inter># prime_factorization b)) =
-          prod_mset (prime_factorization a \<inter># prime_factorization b)"
-    by (intro normalize_prod_mset_primes) auto
-  thus ?thesis by (simp add: gcd_factorial_def)
-qed
-
-lemma gcd_factorial_greatest: "c dvd gcd_factorial a b" if "c dvd a" "c dvd b" for a b c
-proof (cases "a = 0 \<or> b = 0")
-  case False
-  with that have [simp]: "c \<noteq> 0" by auto
-  let ?p = "prime_factorization"
-  from that False have "?p c \<subseteq># ?p a" "?p c \<subseteq># ?p b"
-    by (simp_all add: prime_factorization_subset_iff_dvd)
-  hence "prime_factorization c \<subseteq>#
-           prime_factorization (prod_mset (prime_factorization a \<inter># prime_factorization b))"
-    using False by (subst prime_factorization_prod_mset_primes) auto
-  with False show ?thesis
-    by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric])
-qed (auto simp: gcd_factorial_def that)
-
-lemma lcm_factorial_gcd_factorial:
-  "lcm_factorial a b = normalize (a * b) div gcd_factorial a b" for a b
-proof (cases "a = 0 \<or> b = 0")
-  case False
-  let ?p = "prime_factorization"
-  from False have "prod_mset (?p (a * b)) div gcd_factorial a b =
-                     prod_mset (?p a + ?p b - ?p a \<inter># ?p b)"
-    by (subst prod_mset_diff) (auto simp: lcm_factorial_def gcd_factorial_def
-                                prime_factorization_mult subset_mset.le_infI1)
-  also from False have "prod_mset (?p (a * b)) = normalize (a * b)"
-    by (intro prod_mset_prime_factorization) simp_all
-  also from False have "prod_mset (?p a + ?p b - ?p a \<inter># ?p b) = lcm_factorial a b"
-    by (simp add: union_diff_inter_eq_sup lcm_factorial_def)
-  finally show ?thesis ..
-qed (auto simp: lcm_factorial_def)
-
-lemma normalize_Gcd_factorial:
-  "normalize (Gcd_factorial A) = Gcd_factorial A"
-proof (cases "A \<subseteq> {0}")
-  case False
-  then obtain x where "x \<in> A" "x \<noteq> 0" by blast
-  hence "Inf (prime_factorization ` (A - {0})) \<subseteq># prime_factorization x"
-    by (intro subset_mset.cInf_lower) auto
-  hence "prime p" if "p \<in># Inf (prime_factorization ` (A - {0}))" for p
-    using that by (auto dest: mset_subset_eqD)
-  with False show ?thesis
-    by (auto simp add: Gcd_factorial_def normalize_prod_mset_primes)
-qed (simp_all add: Gcd_factorial_def)
-
-lemma Gcd_factorial_eq_0_iff:
-  "Gcd_factorial A = 0 \<longleftrightarrow> A \<subseteq> {0}"
-  by (auto simp: Gcd_factorial_def in_Inf_multiset_iff split: if_splits)
-
-lemma Gcd_factorial_dvd:
-  assumes "x \<in> A"
-  shows   "Gcd_factorial A dvd x"
-proof (cases "x = 0")
-  case False
-  with assms have "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))"
-    by (intro prime_factorization_Gcd_factorial) auto
-  also from False assms have "\<dots> \<subseteq># prime_factorization x"
-    by (intro subset_mset.cInf_lower) auto
-  finally show ?thesis
-    by (subst (asm) prime_factorization_subset_iff_dvd)
-       (insert assms False, auto simp: Gcd_factorial_eq_0_iff)
-qed simp_all
-
-lemma Gcd_factorial_greatest:
-  assumes "\<And>y. y \<in> A \<Longrightarrow> x dvd y"
-  shows   "x dvd Gcd_factorial A"
-proof (cases "A \<subseteq> {0}")
-  case False
-  from False obtain y where "y \<in> A" "y \<noteq> 0" by auto
-  with assms[of y] have nz: "x \<noteq> 0" by auto
-  from nz assms have "prime_factorization x \<subseteq># prime_factorization y" if "y \<in> A - {0}" for y
-    using that by (subst prime_factorization_subset_iff_dvd) auto
-  with False have "prime_factorization x \<subseteq># Inf (prime_factorization ` (A - {0}))"
-    by (intro subset_mset.cInf_greatest) auto
-  also from False have "\<dots> = prime_factorization (Gcd_factorial A)"
-    by (rule prime_factorization_Gcd_factorial [symmetric])
-  finally show ?thesis
-    by (subst (asm) prime_factorization_subset_iff_dvd)
-       (insert nz False, auto simp: Gcd_factorial_eq_0_iff)
-qed (simp_all add: Gcd_factorial_def)
-
-lemma normalize_Lcm_factorial:
-  "normalize (Lcm_factorial A) = Lcm_factorial A"
-proof (cases "subset_mset.bdd_above (prime_factorization ` A)")
-  case True
-  hence "normalize (prod_mset (Sup (prime_factorization ` A))) =
-           prod_mset (Sup (prime_factorization ` A))"
-    by (intro normalize_prod_mset_primes)
-       (auto simp: in_Sup_multiset_iff)
-  with True show ?thesis by (simp add: Lcm_factorial_def)
-qed (auto simp: Lcm_factorial_def)
-
-lemma Lcm_factorial_eq_0_iff:
-  "Lcm_factorial A = 0 \<longleftrightarrow> 0 \<in> A \<or> \<not>subset_mset.bdd_above (prime_factorization ` A)"
-  by (auto simp: Lcm_factorial_def in_Sup_multiset_iff)
-
-lemma dvd_Lcm_factorial:
-  assumes "x \<in> A"
-  shows   "x dvd Lcm_factorial A"
-proof (cases "0 \<notin> A \<and> subset_mset.bdd_above (prime_factorization ` A)")
-  case True
-  with assms have [simp]: "0 \<notin> A" "x \<noteq> 0" "A \<noteq> {}" by auto
-  from assms True have "prime_factorization x \<subseteq># Sup (prime_factorization ` A)"
-    by (intro subset_mset.cSup_upper) auto
-  also have "\<dots> = prime_factorization (Lcm_factorial A)"
-    by (rule prime_factorization_Lcm_factorial [symmetric]) (insert True, simp_all)
-  finally show ?thesis
-    by (subst (asm) prime_factorization_subset_iff_dvd)
-       (insert True, auto simp: Lcm_factorial_eq_0_iff)
-qed (insert assms, auto simp: Lcm_factorial_def)
-
-lemma Lcm_factorial_least:
-  assumes "\<And>y. y \<in> A \<Longrightarrow> y dvd x"
-  shows   "Lcm_factorial A dvd x"
-proof -
-  consider "A = {}" | "0 \<in> A" | "x = 0" | "A \<noteq> {}" "0 \<notin> A" "x \<noteq> 0" by blast
-  thus ?thesis
-  proof cases
-    assume *: "A \<noteq> {}" "0 \<notin> A" "x \<noteq> 0"
-    hence nz: "x \<noteq> 0" if "x \<in> A" for x using that by auto
-    from * have bdd: "subset_mset.bdd_above (prime_factorization ` A)"
-      by (intro subset_mset.bdd_aboveI[of _ "prime_factorization x"])
-         (auto simp: prime_factorization_subset_iff_dvd nz dest: assms)
-    have "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)"
-      by (rule prime_factorization_Lcm_factorial) fact+
-    also from * have "\<dots> \<subseteq># prime_factorization x"
-      by (intro subset_mset.cSup_least)
-         (auto simp: prime_factorization_subset_iff_dvd nz dest: assms)
-    finally show ?thesis
-      by (subst (asm) prime_factorization_subset_iff_dvd)
-         (insert * bdd, auto simp: Lcm_factorial_eq_0_iff)
-  qed (auto simp: Lcm_factorial_def dest: assms)
-qed
-
-lemmas gcd_lcm_factorial =
-  gcd_factorial_dvd1 gcd_factorial_dvd2 gcd_factorial_greatest
-  normalize_gcd_factorial lcm_factorial_gcd_factorial
-  normalize_Gcd_factorial Gcd_factorial_dvd Gcd_factorial_greatest
-  normalize_Lcm_factorial dvd_Lcm_factorial Lcm_factorial_least
-
-end
-
-class factorial_semiring_gcd = factorial_semiring + gcd + Gcd +
-  assumes gcd_eq_gcd_factorial: "gcd a b = gcd_factorial a b"
-  and     lcm_eq_lcm_factorial: "lcm a b = lcm_factorial a b"
-  and     Gcd_eq_Gcd_factorial: "Gcd A = Gcd_factorial A"
-  and     Lcm_eq_Lcm_factorial: "Lcm A = Lcm_factorial A"
-begin
-
-lemma prime_factorization_gcd:
-  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (gcd a b) = prime_factorization a \<inter># prime_factorization b"
-  by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial)
-
-lemma prime_factorization_lcm:
-  assumes [simp]: "a \<noteq> 0" "b \<noteq> 0"
-  shows   "prime_factorization (lcm a b) = prime_factorization a \<union># prime_factorization b"
-  by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial)
-
-lemma prime_factorization_Gcd:
-  assumes "Gcd A \<noteq> 0"
-  shows   "prime_factorization (Gcd A) = Inf (prime_factorization ` (A - {0}))"
-  using assms
-  by (simp add: prime_factorization_Gcd_factorial Gcd_eq_Gcd_factorial Gcd_factorial_eq_0_iff)
-
-lemma prime_factorization_Lcm:
-  assumes "Lcm A \<noteq> 0"
-  shows   "prime_factorization (Lcm A) = Sup (prime_factorization ` A)"
-  using assms
-  by (simp add: prime_factorization_Lcm_factorial Lcm_eq_Lcm_factorial Lcm_factorial_eq_0_iff)
-
-subclass semiring_gcd
-  by (standard, unfold gcd_eq_gcd_factorial lcm_eq_lcm_factorial)
-     (rule gcd_lcm_factorial; assumption)+
-
-subclass semiring_Gcd
-  by (standard, unfold Gcd_eq_Gcd_factorial Lcm_eq_Lcm_factorial)
-     (rule gcd_lcm_factorial; assumption)+
-
-lemma
-  assumes "x \<noteq> 0" "y \<noteq> 0"
-  shows gcd_eq_factorial': 
-          "gcd x y = (\<Prod>p \<in> prime_factors x \<inter> prime_factors y. 
-                          p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs1")
-    and lcm_eq_factorial':
-          "lcm x y = (\<Prod>p \<in> prime_factors x \<union> prime_factors y. 
-                          p ^ max (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs2")
-proof -
-  have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial)
-  also have "\<dots> = ?rhs1"
-    by (auto simp: gcd_factorial_def assms prod_mset_multiplicity
-          count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong)
-  finally show "gcd x y = ?rhs1" .
-  have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial)
-  also have "\<dots> = ?rhs2"
-    by (auto simp: lcm_factorial_def assms prod_mset_multiplicity
-          count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong)
-  finally show "lcm x y = ?rhs2" .
-qed
-
-lemma
-  assumes "x \<noteq> 0" "y \<noteq> 0" "prime p"
-  shows   multiplicity_gcd: "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)"
-    and   multiplicity_lcm: "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)"
-proof -
-  have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial)
-  also from assms have "multiplicity p \<dots> = min (multiplicity p x) (multiplicity p y)"
-    by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_gcd_factorial)
-  finally show "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" .
-  have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial)
-  also from assms have "multiplicity p \<dots> = max (multiplicity p x) (multiplicity p y)"
-    by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_lcm_factorial)
-  finally show "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" .
-qed
-
-lemma gcd_lcm_distrib:
-  "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)"
-proof (cases "x = 0 \<or> y = 0 \<or> z = 0")
-  case True
-  thus ?thesis
-    by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd)
-next
-  case False
-  hence "normalize (gcd x (lcm y z)) = normalize (lcm (gcd x y) (gcd x z))"
-    by (intro associatedI prime_factorization_subset_imp_dvd)
-       (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm 
-          subset_mset.inf_sup_distrib1)
-  thus ?thesis by simp
-qed
-
-lemma lcm_gcd_distrib:
-  "lcm x (gcd y z) = gcd (lcm x y) (lcm x z)"
-proof (cases "x = 0 \<or> y = 0 \<or> z = 0")
-  case True
-  thus ?thesis
-    by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd)
-next
-  case False
-  hence "normalize (lcm x (gcd y z)) = normalize (gcd (lcm x y) (lcm x z))"
-    by (intro associatedI prime_factorization_subset_imp_dvd)
-       (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm 
-          subset_mset.sup_inf_distrib1)
-  thus ?thesis by simp
-qed
-
-end
-
-class factorial_ring_gcd = factorial_semiring_gcd + idom
-begin
-
-subclass ring_gcd ..
-
-subclass idom_divide ..
-
-end
-
-end
--- a/src/HOL/Number_Theory/Primes.thy	Thu Apr 06 08:33:37 2017 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,725 +0,0 @@
-(*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
-                Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow, 
-                Manuel Eberl
-
-
-This file deals with properties of primes. Definitions and lemmas are
-proved uniformly for the natural numbers and integers.
-
-This file combines and revises a number of prior developments.
-
-The original theories "GCD" and "Primes" were by Christophe Tabacznyj
-and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
-gcd, lcm, and prime for the natural numbers.
-
-The original theory "IntPrimes" was by Thomas M. Rasmussen, and
-extended gcd, lcm, primes to the integers. Amine Chaieb provided
-another extension of the notions to the integers, and added a number
-of results to "Primes" and "GCD". IntPrimes also defined and developed
-the congruence relations on the integers. The notion was extended to
-the natural numbers by Chaieb.
-
-Jeremy Avigad combined all of these, made everything uniform for the
-natural numbers and the integers, and added a number of new theorems.
-
-Tobias Nipkow cleaned up a lot.
-
-Florian Haftmann and Manuel Eberl put primality and prime factorisation
-onto an algebraic foundation and thus generalised these concepts to 
-other rings, such as polynomials. (see also the Factorial_Ring theory).
-
-There were also previous formalisations of unique factorisation by 
-Thomas Marthedal Rasmussen, Jeremy Avigad, and David Gray.
-*)
-
-
-section \<open>Primes\<close>
-
-theory Primes
-imports "~~/src/HOL/Binomial" Euclidean_Algorithm
-begin
-
-(* As a simp or intro rule,
-
-     prime p \<Longrightarrow> p > 0
-
-   wreaks havoc here. When the premise includes \<forall>x \<in># M. prime x, it
-   leads to the backchaining
-
-     x > 0
-     prime x
-     x \<in># M   which is, unfortunately,
-     count M x > 0
-*)
-
-declare [[coercion int]]
-declare [[coercion_enabled]]
-
-lemma prime_elem_nat_iff:
-  "prime_elem (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
-proof safe
-  assume *: "prime_elem n"
-  from * have "n \<noteq> 0" "n \<noteq> 1" by (intro notI, simp del: One_nat_def)+
-  thus "n > 1" by (cases n) simp_all
-  fix m assume m: "m dvd n" "m \<noteq> n"
-  from * \<open>m dvd n\<close> have "n dvd m \<or> is_unit m"
-    by (intro irreducibleD' prime_elem_imp_irreducible)
-  with m show "m = 1" by (auto dest: dvd_antisym)
-next
-  assume "n > 1" "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n"
-  thus "prime_elem n"
-    by (auto simp: prime_elem_iff_irreducible irreducible_altdef)
-qed
-
-lemma prime_nat_iff_prime_elem: "prime (n :: nat) \<longleftrightarrow> prime_elem n"
-  by (simp add: prime_def)
-
-lemma prime_nat_iff:
-  "prime (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
-  by (simp add: prime_nat_iff_prime_elem prime_elem_nat_iff)
-
-lemma prime_elem_int_nat_transfer: "prime_elem (n::int) \<longleftrightarrow> prime_elem (nat (abs n))"
-proof
-  assume "prime_elem n"
-  thus "prime_elem (nat (abs n))" by (auto simp: prime_elem_def nat_dvd_iff)
-next
-  assume "prime_elem (nat (abs n))"
-  thus "prime_elem n"
-    by (auto simp: dvd_int_unfold_dvd_nat prime_elem_def abs_mult nat_mult_distrib)
-qed
-
-lemma prime_elem_nat_int_transfer [simp]: "prime_elem (int n) \<longleftrightarrow> prime_elem n"
-  by (auto simp: prime_elem_int_nat_transfer)
-
-lemma prime_nat_int_transfer [simp]: "prime (int n) \<longleftrightarrow> prime n"
-  by (auto simp: prime_elem_int_nat_transfer prime_def)
-
-lemma prime_int_nat_transfer: "prime (n::int) \<longleftrightarrow> n \<ge> 0 \<and> prime (nat n)"
-  by (auto simp: prime_elem_int_nat_transfer prime_def)
-
-lemma prime_int_iff:
-  "prime (n::int) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n))"
-proof (intro iffI conjI allI impI; (elim conjE)?)
-  assume *: "prime n"
-  hence irred: "irreducible n" by (simp add: prime_elem_imp_irreducible)
-  from * have "n \<ge> 0" "n \<noteq> 0" "n \<noteq> 1" 
-    by (auto simp: prime_def zabs_def not_less split: if_splits)
-  thus "n > 1" by presburger
-  fix m assume "m dvd n" \<open>m \<ge> 0\<close>
-  with irred have "m dvd 1 \<or> n dvd m" by (auto simp: irreducible_altdef)
-  with \<open>m dvd n\<close> \<open>m \<ge> 0\<close> \<open>n > 1\<close> show "m = 1 \<or> m = n"
-    using associated_iff_dvd[of m n] by auto
-next
-  assume n: "1 < n" "\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n"
-  hence "nat n > 1" by simp
-  moreover have "\<forall>m. m dvd nat n \<longrightarrow> m = 1 \<or> m = nat n"
-  proof (intro allI impI)
-    fix m assume "m dvd nat n"
-    with \<open>n > 1\<close> have "int m dvd n" by (auto simp: int_dvd_iff)
-    with n(2) have "int m = 1 \<or> int m = n" by auto
-    thus "m = 1 \<or> m = nat n" by auto
-  qed
-  ultimately show "prime n" 
-    unfolding prime_int_nat_transfer prime_nat_iff by auto
-qed
-
-
-lemma prime_nat_not_dvd:
-  assumes "prime p" "p > n" "n \<noteq> (1::nat)"
-  shows   "\<not>n dvd p"
-proof
-  assume "n dvd p"
-  from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
-  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
-    by (cases "n = 0") (auto dest!: dvd_imp_le)
-qed
-
-lemma prime_int_not_dvd:
-  assumes "prime p" "p > n" "n > (1::int)"
-  shows   "\<not>n dvd p"
-proof
-  assume "n dvd p"
-  from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
-  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
-    by (auto dest!: zdvd_imp_le)
-qed
-
-lemma prime_odd_nat: "prime p \<Longrightarrow> p > (2::nat) \<Longrightarrow> odd p"
-  by (intro prime_nat_not_dvd) auto
-
-lemma prime_odd_int: "prime p \<Longrightarrow> p > (2::int) \<Longrightarrow> odd p"
-  by (intro prime_int_not_dvd) auto
-
-lemma prime_ge_0_int: "prime p \<Longrightarrow> p \<ge> (0::int)"
-  unfolding prime_int_iff by auto
-
-lemma prime_gt_0_nat: "prime p \<Longrightarrow> p > (0::nat)"
-  unfolding prime_nat_iff by auto
-
-lemma prime_gt_0_int: "prime p \<Longrightarrow> p > (0::int)"
-  unfolding prime_int_iff by auto
-
-lemma prime_ge_1_nat: "prime p \<Longrightarrow> p \<ge> (1::nat)"
-  unfolding prime_nat_iff by auto
-
-lemma prime_ge_Suc_0_nat: "prime p \<Longrightarrow> p \<ge> Suc 0"
-  unfolding prime_nat_iff by auto
-
-lemma prime_ge_1_int: "prime p \<Longrightarrow> p \<ge> (1::int)"
-  unfolding prime_int_iff by auto
-
-lemma prime_gt_1_nat: "prime p \<Longrightarrow> p > (1::nat)"
-  unfolding prime_nat_iff by auto
-
-lemma prime_gt_Suc_0_nat: "prime p \<Longrightarrow> p > Suc 0"
-  unfolding prime_nat_iff by auto
-
-lemma prime_gt_1_int: "prime p \<Longrightarrow> p > (1::int)"
-  unfolding prime_int_iff by auto
-
-lemma prime_ge_2_nat: "prime p \<Longrightarrow> p \<ge> (2::nat)"
-  unfolding prime_nat_iff by auto
-
-lemma prime_ge_2_int: "prime p \<Longrightarrow> p \<ge> (2::int)"
-  unfolding prime_int_iff by auto
-
-lemma prime_int_altdef:
-  "prime p = (1 < p \<and> (\<forall>m::int. m \<ge> 0 \<longrightarrow> m dvd p \<longrightarrow>
-    m = 1 \<or> m = p))"
-  unfolding prime_int_iff by blast
-
-lemma not_prime_eq_prod_nat:
-  assumes "m > 1" "\<not>prime (m::nat)"
-  shows   "\<exists>n k. n = m * k \<and> 1 < m \<and> m < n \<and> 1 < k \<and> k < n"
-  using assms irreducible_altdef[of m]
-  by (auto simp: prime_elem_iff_irreducible prime_def irreducible_altdef)
-
-    
-subsection\<open>Largest exponent of a prime factor\<close>
-text\<open>Possibly duplicates other material, but avoid the complexities of multisets.\<close>
-  
-lemma prime_power_cancel_less:
-  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and less: "k < k'" and "\<not> p dvd m"
-  shows False
-proof -
-  obtain l where l: "k' = k + l" and "l > 0"
-    using less less_imp_add_positive by auto
-  have "m = m * (p ^ k) div (p ^ k)"
-    using \<open>prime p\<close> by simp
-  also have "\<dots> = m' * (p ^ k') div (p ^ k)"
-    using eq by simp
-  also have "\<dots> = m' * (p ^ l) * (p ^ k) div (p ^ k)"
-    by (simp add: l mult.commute mult.left_commute power_add)
-  also have "... = m' * (p ^ l)"
-    using \<open>prime p\<close> by simp
-  finally have "p dvd m"
-    using \<open>l > 0\<close> by simp
-  with assms show False
-    by simp
-qed
-
-lemma prime_power_cancel:
-  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and "\<not> p dvd m" "\<not> p dvd m'"
-  shows "k = k'"
-  using prime_power_cancel_less [OF \<open>prime p\<close>] assms
-  by (metis linorder_neqE_nat)
-
-lemma prime_power_cancel2:
-  assumes "prime p" "m * (p ^ k) = m' * (p ^ k')" "\<not> p dvd m" "\<not> p dvd m'"
-  obtains "m = m'" "k = k'"
-  using prime_power_cancel [OF assms] assms by auto
-
-lemma prime_power_canonical:
-  fixes m::nat
-  assumes "prime p" "m > 0"
-  shows "\<exists>k n. \<not> p dvd n \<and> m = n * p^k"
-using \<open>m > 0\<close>
-proof (induction m rule: less_induct)
-  case (less m)
-  show ?case
-  proof (cases "p dvd m")
-    case True
-    then obtain m' where m': "m = p * m'"
-      using dvdE by blast
-    with \<open>prime p\<close> have "0 < m'" "m' < m"
-      using less.prems prime_nat_iff by auto
-    with m' less show ?thesis
-      by (metis power_Suc mult.left_commute)
-  next
-    case False
-    then show ?thesis
-      by (metis mult.right_neutral power_0)
-  qed
-qed
-
-
-subsubsection \<open>Make prime naively executable\<close>
-
-lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
-  unfolding One_nat_def [symmetric] by (rule not_prime_1)
-
-lemma prime_nat_iff':
-  "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)"
-proof safe
-  assume "p > 1" and *: "\<forall>n\<in>{2..<p}. \<not>n dvd p"
-  show "prime p" unfolding prime_nat_iff
-  proof (intro conjI allI impI)
-    fix m assume "m dvd p"
-    with \<open>p > 1\<close> have "m \<noteq> 0" by (intro notI) auto
-    hence "m \<ge> 1" by simp
-    moreover from \<open>m dvd p\<close> and * have "m \<notin> {2..<p}" by blast
-    with \<open>m dvd p\<close> and \<open>p > 1\<close> have "m \<le> 1 \<or> m = p" by (auto dest: dvd_imp_le)
-    ultimately show "m = 1 \<or> m = p" by simp
-  qed fact+
-qed (auto simp: prime_nat_iff)
-
-lemma prime_int_iff':
-  "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)" (is "?lhs = ?rhs")
-proof
-  assume "?lhs"
-  thus "?rhs"
-      by (auto simp: prime_int_nat_transfer dvd_int_unfold_dvd_nat prime_nat_iff')
-next
-  assume "?rhs"
-  thus "?lhs"
-    by (auto simp: prime_int_nat_transfer zdvd_int prime_nat_iff')
-qed
-
-lemma prime_int_numeral_eq [simp]:
-  "prime (numeral m :: int) \<longleftrightarrow> prime (numeral m :: nat)"
-  by (simp add: prime_int_nat_transfer)
-
-lemma two_is_prime_nat [simp]: "prime (2::nat)"
-  by (simp add: prime_nat_iff')
-
-lemma prime_nat_numeral_eq [simp]:
-  "prime (numeral m :: nat) \<longleftrightarrow>
-    (1::nat) < numeral m \<and>
-    (\<forall>n::nat \<in> set [2..<numeral m]. \<not> n dvd numeral m)"
-  by (simp only: prime_nat_iff' set_upt)  \<comment> \<open>TODO Sieve Of Erathosthenes might speed this up\<close>
-
-
-text\<open>A bit of regression testing:\<close>
-
-lemma "prime(97::nat)" by simp
-lemma "prime(97::int)" by simp
-
-lemma prime_factor_nat: 
-  "n \<noteq> (1::nat) \<Longrightarrow> \<exists>p. prime p \<and> p dvd n"
-  using prime_divisor_exists[of n]
-  by (cases "n = 0") (auto intro: exI[of _ "2::nat"])
-
-
-subsection \<open>Infinitely many primes\<close>
-
-lemma next_prime_bound: "\<exists>p::nat. prime p \<and> n < p \<and> p \<le> fact n + 1"
-proof-
-  have f1: "fact n + 1 \<noteq> (1::nat)" using fact_ge_1 [of n, where 'a=nat] by arith
-  from prime_factor_nat [OF f1]
-  obtain p :: nat where "prime p" and "p dvd fact n + 1" by auto
-  then have "p \<le> fact n + 1" apply (intro dvd_imp_le) apply auto done
-  { assume "p \<le> n"
-    from \<open>prime p\<close> have "p \<ge> 1"
-      by (cases p, simp_all)
-    with \<open>p <= n\<close> have "p dvd fact n"
-      by (intro dvd_fact)
-    with \<open>p dvd fact n + 1\<close> have "p dvd fact n + 1 - fact n"
-      by (rule dvd_diff_nat)
-    then have "p dvd 1" by simp
-    then have "p <= 1" by auto
-    moreover from \<open>prime p\<close> have "p > 1"
-      using prime_nat_iff by blast
-    ultimately have False by auto}
-  then have "n < p" by presburger
-  with \<open>prime p\<close> and \<open>p <= fact n + 1\<close> show ?thesis by auto
-qed
-
-lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
-  using next_prime_bound by auto
-
-lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
-proof
-  assume "finite {(p::nat). prime p}"
-  with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
-    by auto
-  then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
-    by auto
-  with bigger_prime [of b] show False
-    by auto
-qed
-
-subsection\<open>Powers of Primes\<close>
-
-text\<open>Versions for type nat only\<close>
-
-lemma prime_product:
-  fixes p::nat
-  assumes "prime (p * q)"
-  shows "p = 1 \<or> q = 1"
-proof -
-  from assms have
-    "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
-    unfolding prime_nat_iff by auto
-  from \<open>1 < p * q\<close> have "p \<noteq> 0" by (cases p) auto
-  then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
-  have "p dvd p * q" by simp
-  then have "p = 1 \<or> p = p * q" by (rule P)
-  then show ?thesis by (simp add: Q)
-qed
-
-(* TODO: Generalise? *)
-lemma prime_power_mult_nat:
-  fixes p::nat
-  assumes p: "prime p" and xy: "x * y = p ^ k"
-  shows "\<exists>i j. x = p ^i \<and> y = p^ j"
-using xy
-proof(induct k arbitrary: x y)
-  case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
-next
-  case (Suc k x y)
-  from Suc.prems have pxy: "p dvd x*y" by auto
-  from prime_dvd_multD [OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
-  from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)
-  {assume px: "p dvd x"
-    then obtain d where d: "x = p*d" unfolding dvd_def by blast
-    from Suc.prems d  have "p*d*y = p^Suc k" by simp
-    hence th: "d*y = p^k" using p0 by simp
-    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
-    with d have "x = p^Suc i" by simp
-    with ij(2) have ?case by blast}
-  moreover
-  {assume px: "p dvd y"
-    then obtain d where d: "y = p*d" unfolding dvd_def by blast
-    from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult.commute)
-    hence th: "d*x = p^k" using p0 by simp
-    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
-    with d have "y = p^Suc i" by simp
-    with ij(2) have ?case by blast}
-  ultimately show ?case  using pxyc by blast
-qed
-
-lemma prime_power_exp_nat:
-  fixes p::nat
-  assumes p: "prime p" and n: "n \<noteq> 0"
-    and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
-  using n xn
-proof(induct n arbitrary: k)
-  case 0 thus ?case by simp
-next
-  case (Suc n k) hence th: "x*x^n = p^k" by simp
-  {assume "n = 0" with Suc have ?case by simp (rule exI[where x="k"], simp)}
-  moreover
-  {assume n: "n \<noteq> 0"
-    from prime_power_mult_nat[OF p th]
-    obtain i j where ij: "x = p^i" "x^n = p^j"by blast
-    from Suc.hyps[OF n ij(2)] have ?case .}
-  ultimately show ?case by blast
-qed
-
-lemma divides_primepow_nat:
-  fixes p::nat
-  assumes p: "prime p"
-  shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
-proof
-  assume H: "d dvd p^k" then obtain e where e: "d*e = p^k"
-    unfolding dvd_def  apply (auto simp add: mult.commute) by blast
-  from prime_power_mult_nat[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
-  from e ij have "p^(i + j) = p^k" by (simp add: power_add)
-  hence "i + j = k" using p prime_gt_1_nat power_inject_exp[of p "i+j" k] by simp
-  hence "i \<le> k" by arith
-  with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
-next
-  {fix i assume H: "i \<le> k" "d = p^i"
-    then obtain j where j: "k = i + j"
-      by (metis le_add_diff_inverse)
-    hence "p^k = p^j*d" using H(2) by (simp add: power_add)
-    hence "d dvd p^k" unfolding dvd_def by auto}
-  thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
-qed
-
-
-subsection \<open>Chinese Remainder Theorem Variants\<close>
-
-lemma bezout_gcd_nat:
-  fixes a::nat shows "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
-  using bezout_nat[of a b]
-by (metis bezout_nat diff_add_inverse gcd_add_mult gcd.commute
-  gcd_nat.right_neutral mult_0)
-
-lemma gcd_bezout_sum_nat:
-  fixes a::nat
-  assumes "a * x + b * y = d"
-  shows "gcd a b dvd d"
-proof-
-  let ?g = "gcd a b"
-    have dv: "?g dvd a*x" "?g dvd b * y"
-      by simp_all
-    from dvd_add[OF dv] assms
-    show ?thesis by auto
-qed
-
-
-text \<open>A binary form of the Chinese Remainder Theorem.\<close>
-
-(* TODO: Generalise? *)
-lemma chinese_remainder:
-  fixes a::nat  assumes ab: "coprime a b" and a: "a \<noteq> 0" and b: "b \<noteq> 0"
-  shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
-proof-
-  from bezout_add_strong_nat[OF a, of b] bezout_add_strong_nat[OF b, of a]
-  obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"
-    and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
-  then have d12: "d1 = 1" "d2 =1"
-    by (metis ab coprime_nat)+
-  let ?x = "v * a * x1 + u * b * x2"
-  let ?q1 = "v * x1 + u * y2"
-  let ?q2 = "v * y1 + u * x2"
-  from dxy2(3)[simplified d12] dxy1(3)[simplified d12]
-  have "?x = u + ?q1 * a" "?x = v + ?q2 * b"
-    by algebra+
-  thus ?thesis by blast
-qed
-
-text \<open>Primality\<close>
-
-lemma coprime_bezout_strong:
-  fixes a::nat assumes "coprime a b"  "b \<noteq> 1"
-  shows "\<exists>x y. a * x = b * y + 1"
-by (metis assms bezout_nat gcd_nat.left_neutral)
-
-lemma bezout_prime:
-  assumes p: "prime p" and pa: "\<not> p dvd a"
-  shows "\<exists>x y. a*x = Suc (p*y)"
-proof -
-  have ap: "coprime a p"
-    by (metis gcd.commute p pa prime_imp_coprime)
-  moreover from p have "p \<noteq> 1" by auto
-  ultimately have "\<exists>x y. a * x = p * y + 1"
-    by (rule coprime_bezout_strong)
-  then show ?thesis by simp    
-qed
-(* END TODO *)
-
-
-
-subsection \<open>Multiplicity and primality for natural numbers and integers\<close>
-
-lemma prime_factors_gt_0_nat:
-  "p \<in> prime_factors x \<Longrightarrow> p > (0::nat)"
-  by (simp add: in_prime_factors_imp_prime prime_gt_0_nat)
-
-lemma prime_factors_gt_0_int:
-  "p \<in> prime_factors x \<Longrightarrow> p > (0::int)"
-  by (simp add: in_prime_factors_imp_prime prime_gt_0_int)
-
-lemma prime_factors_ge_0_int [elim]: (* FIXME !? *)
-  fixes n :: int
-  shows "p \<in> prime_factors n \<Longrightarrow> p \<ge> 0"
-  by (drule prime_factors_gt_0_int) simp
-  
-lemma prod_mset_prime_factorization_int:
-  fixes n :: int
-  assumes "n > 0"
-  shows   "prod_mset (prime_factorization n) = n"
-  using assms by (simp add: prod_mset_prime_factorization)
-
-lemma prime_factorization_exists_nat:
-  "n > 0 \<Longrightarrow> (\<exists>M. (\<forall>p::nat \<in> set_mset M. prime p) \<and> n = (\<Prod>i \<in># M. i))"
-  using prime_factorization_exists[of n] by (auto simp: prime_def)
-
-lemma prod_mset_prime_factorization_nat [simp]: 
-  "(n::nat) > 0 \<Longrightarrow> prod_mset (prime_factorization n) = n"
-  by (subst prod_mset_prime_factorization) simp_all
-
-lemma prime_factorization_nat:
-    "n > (0::nat) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
-  by (simp add: prod_prime_factors)
-
-lemma prime_factorization_int:
-    "n > (0::int) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
-  by (simp add: prod_prime_factors)
-
-lemma prime_factorization_unique_nat:
-  fixes f :: "nat \<Rightarrow> _"
-  assumes S_eq: "S = {p. 0 < f p}"
-    and "finite S"
-    and S: "\<forall>p\<in>S. prime p" "n = (\<Prod>p\<in>S. p ^ f p)"
-  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
-  using assms by (intro prime_factorization_unique'') auto
-
-lemma prime_factorization_unique_int:
-  fixes f :: "int \<Rightarrow> _"
-  assumes S_eq: "S = {p. 0 < f p}"
-    and "finite S"
-    and S: "\<forall>p\<in>S. prime p" "abs n = (\<Prod>p\<in>S. p ^ f p)"
-  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
-  using assms by (intro prime_factorization_unique'') auto
-
-lemma prime_factors_characterization_nat:
-  "S = {p. 0 < f (p::nat)} \<Longrightarrow>
-    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
-  by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric])
-
-lemma prime_factors_characterization'_nat:
-  "finite {p. 0 < f (p::nat)} \<Longrightarrow>
-    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
-      prime_factors (\<Prod>p | 0 < f p. p ^ f p) = {p. 0 < f p}"
-  by (rule prime_factors_characterization_nat) auto
-
-lemma prime_factors_characterization_int:
-  "S = {p. 0 < f (p::int)} \<Longrightarrow> finite S \<Longrightarrow>
-    \<forall>p\<in>S. prime p \<Longrightarrow> abs n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
-  by (rule prime_factorization_unique_int [THEN conjunct1, symmetric])
-
-(* TODO Move *)
-lemma abs_prod: "abs (prod f A :: 'a :: linordered_idom) = prod (\<lambda>x. abs (f x)) A"
-  by (cases "finite A", induction A rule: finite_induct) (simp_all add: abs_mult)
-
-lemma primes_characterization'_int [rule_format]:
-  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow> \<forall>p. 0 < f p \<longrightarrow> prime p \<Longrightarrow>
-      prime_factors (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = {p. p \<ge> 0 \<and> 0 < f p}"
-  by (rule prime_factors_characterization_int) (auto simp: abs_prod prime_ge_0_int)
-
-lemma multiplicity_characterization_nat:
-  "S = {p. 0 < f (p::nat)} \<Longrightarrow> finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow>
-    n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
-  by (frule prime_factorization_unique_nat [of S f n, THEN conjunct2, rule_format, symmetric]) auto
-
-lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
-    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<longrightarrow> prime p \<longrightarrow>
-      multiplicity p (\<Prod>p | 0 < f p. p ^ f p) = f p"
-  by (intro impI, rule multiplicity_characterization_nat) auto
-
-lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
-    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
-  by (frule prime_factorization_unique_int [of S f n, THEN conjunct2, rule_format, symmetric]) 
-     (auto simp: abs_prod power_abs prime_ge_0_int intro!: prod.cong)
-
-lemma multiplicity_characterization'_int [rule_format]:
-  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow>
-    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> prime p \<Longrightarrow>
-      multiplicity p (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = f p"
-  by (rule multiplicity_characterization_int) (auto simp: prime_ge_0_int)
-
-lemma multiplicity_one_nat [simp]: "multiplicity p (Suc 0) = 0"
-  unfolding One_nat_def [symmetric] by (rule multiplicity_one)
-
-lemma multiplicity_eq_nat:
-  fixes x and y::nat
-  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
-  shows "x = y"
-  using multiplicity_eq_imp_eq[of x y] assms by simp
-
-lemma multiplicity_eq_int:
-  fixes x y :: int
-  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
-  shows "x = y"
-  using multiplicity_eq_imp_eq[of x y] assms by simp
-
-lemma multiplicity_prod_prime_powers:
-  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> prime x" "prime p"
-  shows   "multiplicity p (\<Prod>p \<in> S. p ^ f p) = (if p \<in> S then f p else 0)"
-proof -
-  define g where "g = (\<lambda>x. if x \<in> S then f x else 0)"
-  define A where "A = Abs_multiset g"
-  have "{x. g x > 0} \<subseteq> S" by (auto simp: g_def)
-  from finite_subset[OF this assms(1)] have [simp]: "g :  multiset"
-    by (simp add: multiset_def)
-  from assms have count_A: "count A x = g x" for x unfolding A_def
-    by simp
-  have set_mset_A: "set_mset A = {x\<in>S. f x > 0}"
-    unfolding set_mset_def count_A by (auto simp: g_def)
-  with assms have prime: "prime x" if "x \<in># A" for x using that by auto
-  from set_mset_A assms have "(\<Prod>p \<in> S. p ^ f p) = (\<Prod>p \<in> S. p ^ g p) "
-    by (intro prod.cong) (auto simp: g_def)
-  also from set_mset_A assms have "\<dots> = (\<Prod>p \<in> set_mset A. p ^ g p)"
-    by (intro prod.mono_neutral_right) (auto simp: g_def set_mset_A)
-  also have "\<dots> = prod_mset A"
-    by (auto simp: prod_mset_multiplicity count_A set_mset_A intro!: prod.cong)
-  also from assms have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
-    by (subst prime_elem_multiplicity_prod_mset_distrib) (auto dest: prime)
-  also from assms have "image_mset (multiplicity p) A = image_mset (\<lambda>x. if x = p then 1 else 0) A"
-    by (intro image_mset_cong) (auto simp: prime_multiplicity_other dest: prime)
-  also have "sum_mset \<dots> = (if p \<in> S then f p else 0)" by (simp add: sum_mset_delta count_A g_def)
-  finally show ?thesis .
-qed
-
-lemma prime_factorization_prod_mset:
-  assumes "0 \<notin># A"
-  shows "prime_factorization (prod_mset A) = \<Union>#(image_mset prime_factorization A)"
-  using assms by (induct A) (auto simp add: prime_factorization_mult)
-
-lemma prime_factors_prod:
-  assumes "finite A" and "0 \<notin> f ` A"
-  shows "prime_factors (prod f A) = UNION A (prime_factors \<circ> f)"
-  using assms by (simp add: prod_unfold_prod_mset prime_factorization_prod_mset)
-
-lemma prime_factors_fact:
-  "prime_factors (fact n) = {p \<in> {2..n}. prime p}" (is "?M = ?N")
-proof (rule set_eqI)
-  fix p
-  { fix m :: nat
-    assume "p \<in> prime_factors m"
-    then have "prime p" and "p dvd m" by auto
-    moreover assume "m > 0" 
-    ultimately have "2 \<le> p" and "p \<le> m"
-      by (auto intro: prime_ge_2_nat dest: dvd_imp_le)
-    moreover assume "m \<le> n"
-    ultimately have "2 \<le> p" and "p \<le> n"
-      by (auto intro: order_trans)
-  } note * = this
-  show "p \<in> ?M \<longleftrightarrow> p \<in> ?N"
-    by (auto simp add: fact_prod prime_factors_prod Suc_le_eq dest!: prime_prime_factors intro: *)
-qed
-
-lemma prime_dvd_fact_iff:
-  assumes "prime p"
-  shows "p dvd fact n \<longleftrightarrow> p \<le> n"
-  using assms
-  by (auto simp add: prime_factorization_subset_iff_dvd [symmetric]
-    prime_factorization_prime prime_factors_fact prime_ge_2_nat)
-
-(* TODO Legacy names *)
-lemmas prime_imp_coprime_nat = prime_imp_coprime[where ?'a = nat]
-lemmas prime_imp_coprime_int = prime_imp_coprime[where ?'a = int]
-lemmas prime_dvd_mult_nat = prime_dvd_mult_iff[where ?'a = nat]
-lemmas prime_dvd_mult_int = prime_dvd_mult_iff[where ?'a = int]
-lemmas prime_dvd_mult_eq_nat = prime_dvd_mult_iff[where ?'a = nat]
-lemmas prime_dvd_mult_eq_int = prime_dvd_mult_iff[where ?'a = int]
-lemmas prime_dvd_power_nat = prime_dvd_power[where ?'a = nat]
-lemmas prime_dvd_power_int = prime_dvd_power[where ?'a = int]
-lemmas prime_dvd_power_nat_iff = prime_dvd_power_iff[where ?'a = nat]
-lemmas prime_dvd_power_int_iff = prime_dvd_power_iff[where ?'a = int]
-lemmas prime_imp_power_coprime_nat = prime_imp_power_coprime[where ?'a = nat]
-lemmas prime_imp_power_coprime_int = prime_imp_power_coprime[where ?'a = int]
-lemmas primes_coprime_nat = primes_coprime[where ?'a = nat]
-lemmas primes_coprime_int = primes_coprime[where ?'a = nat]
-lemmas prime_divprod_pow_nat = prime_elem_divprod_pow[where ?'a = nat]
-lemmas prime_exp = prime_elem_power_iff[where ?'a = nat]
-
-text \<open>Code generation\<close>
-  
-context
-begin
-
-qualified definition prime_nat :: "nat \<Rightarrow> bool"
-  where [simp, code_abbrev]: "prime_nat = prime"
-
-lemma prime_nat_naive [code]:
-  "prime_nat p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
-  by (auto simp add: prime_nat_iff')
-
-qualified definition prime_int :: "int \<Rightarrow> bool"
-  where [simp, code_abbrev]: "prime_int = prime"
-
-lemma prime_int_naive [code]:
-  "prime_int p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
-  by (auto simp add: prime_int_iff')
-
-lemma "prime(997::nat)" by eval
-
-lemma "prime(997::int)" by eval
-  
-end
-
-end
--- a/src/HOL/Proofs/Extraction/Euclid.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/Proofs/Extraction/Euclid.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -8,7 +8,7 @@
 
 theory Euclid
 imports
-  "~~/src/HOL/Number_Theory/Primes"
+  "~~/src/HOL/Computational_Algebra/Primes"
   Util
   "~~/src/HOL/Library/Code_Target_Numeral"
 begin
--- a/src/HOL/ROOT	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/ROOT	Thu Apr 06 21:37:13 2017 +0200
@@ -32,10 +32,8 @@
   theories
     Library
     (*conflicting type class instantiations and dependent applications*)
-    Field_as_Ring
     Finite_Lattice
     List_lexord
-    Polynomial_Factorial
     Prefix_Order
     Product_Lexorder
     Product_Order
@@ -71,6 +69,13 @@
     Approximations
     Circle_Area
 
+session "HOL-Computation_Algebra" in "Computational_Algebra" = HOL +
+  theories
+    Computational_Algebra
+    (*conflicting type class instantiations and dependent applications*)
+    Field_as_Ring
+    Polynomial_Factorial
+
 session "HOL-Hahn_Banach" in Hahn_Banach = HOL +
   description {*
     Author:     Gertrud Bauer, TU Munich
@@ -301,7 +306,7 @@
   theories [document = false]
     (* Preliminaries from set and number theory *)
     "~~/src/HOL/Library/FuncSet"
-    "~~/src/HOL/Number_Theory/Primes"
+    "~~/src/HOL/Computational_Algebra/Primes"
     "~~/src/HOL/Library/Permutation"
   theories
     (* Orders and Lattices *)
@@ -418,7 +423,7 @@
   theories [document = false]
     "~~/src/HOL/Library/Code_Target_Numeral"
     "~~/src/HOL/Library/Monad_Syntax"
-    "~~/src/HOL/Number_Theory/Primes"
+    "~~/src/HOL/Computational_Algebra/Primes"
     "~~/src/HOL/Library/State_Monad"
   theories
     Greatest_Common_Divisor
@@ -635,7 +640,7 @@
   options [quick_and_dirty]
   theories [document = false]
     "~~/src/HOL/Library/Lattice_Syntax"
-    "../Number_Theory/Primes"
+    "../Computational_Algebra/Primes"
   theories
     Knaster_Tarski
     Peirce
--- a/src/HOL/ex/Sqrt.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/ex/Sqrt.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -5,7 +5,7 @@
 section \<open>Square roots of primes are irrational\<close>
 
 theory Sqrt
-imports Complex_Main "~~/src/HOL/Number_Theory/Primes"
+imports Complex_Main "~~/src/HOL/Computational_Algebra/Primes"
 begin
 
 text \<open>The square root of any prime number (including 2) is irrational.\<close>
--- a/src/HOL/ex/Sqrt_Script.thy	Thu Apr 06 08:33:37 2017 +0200
+++ b/src/HOL/ex/Sqrt_Script.thy	Thu Apr 06 21:37:13 2017 +0200
@@ -6,7 +6,7 @@
 section \<open>Square roots of primes are irrational (script version)\<close>
 
 theory Sqrt_Script
-imports Complex_Main "~~/src/HOL/Number_Theory/Primes"
+imports Complex_Main "~~/src/HOL/Computational_Algebra/Primes"
 begin
 
 text \<open>