diff --git a/src/HOL/Computational_Algebra/Polynomial.thy b/src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy +++ b/src/HOL/Computational_Algebra/Polynomial.thy @@ -1,5125 +1,5125 @@ (* Title: HOL/Computational_Algebra/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial imports Complex_Main "HOL-Library.More_List" "HOL-Library.Infinite_Set" Factorial_Ring begin context semidom_modulo begin lemma not_dvd_imp_mod_neq_0: \a mod b \ 0\ if \\ b dvd a\ using that mod_0_imp_dvd [of a b] by blast end subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) where "x ## xs = (if xs = [] \ 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 \ 0 \ x ## xs = x # xs" by (simp add: cCons_def) lemma strip_while_not_0_Cons_eq [simp]: "strip_while (\x. x = 0) (x # xs) = x ## strip_while (\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 \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" morphisms coeff Abs_poly by (auto intro!: ALL_MOST) setup_lifting type_definition_poly lemma poly_eq_iff: "p = q \ (\n. coeff p n = coeff q n)" by (simp add: coeff_inject [symmetric] fun_eq_iff) lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" by (simp add: poly_eq_iff) lemma MOST_coeff_eq_0: "\\<^sub>\ n. coeff p n = 0" using coeff [of p] by simp subsection \Degree of a polynomial\ definition degree :: "'a::zero poly \ nat" where "degree p = (LEAST n. \i>n. coeff p i = 0)" lemma coeff_eq_0: assumes "degree p < n" shows "coeff p n = 0" proof - have "\n. \i>n. coeff p i = 0" using MOST_coeff_eq_0 by (simp add: MOST_nat) then have "\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 \ 0 \ n \ degree p" by (erule contrapos_np, rule coeff_eq_0, simp) lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" unfolding degree_def by (erule Least_le) lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" unfolding degree_def by (drule not_less_Least, simp) subsection \The zero polynomial\ instantiation poly :: (zero) zero begin lift_definition zero_poly :: "'a poly" is "\_. 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 \ 0" shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from \p \ 0\ obtain n where "coeff p n \ 0" by (auto simp add: poly_eq_iff) then have "n \ degree p" by (rule le_degree) with \coeff p n \ 0\ and \degree p = 0\ show "coeff p (degree p) \ 0" by simp next case (Suc n) from \degree p = Suc n\ have "n < degree p" by simp then have "\i>n. coeff p i \ 0" by (rule less_degree_imp) then obtain i where "n < i" and "coeff p i \ 0" by blast from \degree p = Suc n\ and \n < i\ have "degree p \ i" by simp also from \coeff p i \ 0\ have "i \ degree p" by (rule le_degree) finally have "degree p = i" . with \coeff p i \ 0\ show "coeff p (degree p) \ 0" by simp qed lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0") (simp_all add: leading_coeff_neq_0) lemma eq_zero_or_degree_less: assumes "degree p \ n" and "coeff p n = 0" shows "p = 0 \ degree p < n" proof (cases n) case 0 with \degree p \ n\ and \coeff p n = 0\ have "coeff p (degree p) = 0" by simp then have "p = 0" by simp then show ?thesis .. next case (Suc m) from \degree p \ n\ have "\i>n. coeff p i = 0" by (simp add: coeff_eq_0) with \coeff p n = 0\ have "\i\n. coeff p i = 0" by (simp add: le_less) with \n = Suc m\ have "\i>m. coeff p i = 0" by (simp add: less_eq_Suc_le) then have "degree p \ m" by (rule degree_le) with \n = Suc m\ 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 \ degree rrr \ dr \ degree rrr \ dr - 1" using eq_zero_or_degree_less by fastforce subsection \List-style constructor for polynomials\ lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" is "\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) \ Suc (degree p)" by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) lemma degree_pCons_eq: "p \ 0 \ degree (pCons a p) = Suc (degree p)" by (simp add: degree_pCons_le le_antisym le_degree) lemma degree_pCons_0: "degree (pCons a 0) = 0" proof - have "degree (pCons a 0) \ Suc 0" by (metis (no_types) degree_0 degree_pCons_le) then show ?thesis by (metis coeff_0 coeff_pCons_Suc degree_0 eq_zero_or_degree_less less_Suc0) qed lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" by (simp add: degree_pCons_0 degree_pCons_eq) 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 \ a = b \ 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 \ a = 0 \ 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 (\n. coeff p (Suc n)))" by transfer (simp_all add: MOST_inj[where f=Suc and P="\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: "\a p. a \ 0 \ p \ 0 \ P p \ 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 \p = pCons a q\ have "degree q < degree p" by simp then show "P q" by (rule less.hyps) qed have "P (pCons a q)" proof (cases "a \ 0 \ q \ 0") case True with \P q\ show ?thesis by (auto intro: pCons) next case False with zero show ?thesis by simp qed with \p = pCons a q\ 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 \Quickcheck generator for polynomials\ quickcheck_generator poly constructors: "0 :: _ poly", pCons subsection \List-style syntax for polynomials\ syntax "_poly" :: "args \ 'a poly" ("[:(_):]") translations "[:x, xs:]" \ "CONST pCons x [:xs:]" "[:x:]" \ "CONST pCons x 0" "[:x:]" \ "CONST pCons x (_constrain 0 t)" subsection \Representation of polynomials by lists of coefficients\ primrec Poly :: "'a::zero list \ '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 \ (\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 \ Poly (rev (tl as)) = Poly (rev as)" by (cases as) simp_all lemma degree_Poly: "degree (Poly xs) \ 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 \ 'a::zero list" where "coeffs p = (if p = 0 then [] else map (\i. coeff p i) [0 ..< Suc (degree p)])" lemma coeffs_eq_Nil [simp]: "coeffs p = [] \ p = 0" by (simp add: coeffs_def) lemma not_0_coeffs_not_Nil: "p \ 0 \ coeffs p \ []" 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 *: "\m\set ms. m > 0 \ map (case_nat x f) ms = map f (map (\n. n - 1) ms)" for ms :: "nat list" and f :: "nat \ '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 \ 0 \ length (coeffs p) = degree p + 1" by (simp add: coeffs_def) lemma coeffs_nth: "p \ 0 \ n \ degree p \ coeffs p ! n = coeff p n" by (auto simp: coeffs_def simp del: upt_Suc) lemma coeff_in_coeffs: "p \ 0 \ n \ degree p \ coeff p n \ set (coeffs p)" using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) lemma not_0_cCons_eq [simp]: "p \ 0 \ 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 "(\n. as \ replicate n 0) \ (\a\set as. a \ 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 \ coeffs p = coeffs q" (is "?P \ ?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: "\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 \ 0 \ 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) lemma set_coeffs_subset_singleton_0_iff [simp]: "set (coeffs p) \ {0} \ p = 0" by (auto simp add: coeffs_def intro: classical) lemma set_coeffs_not_only_0 [simp]: "set (coeffs p) \ {0}" by (auto simp add: set_eq_subset) lemma forall_coeffs_conv: "(\n. P (coeff p n)) \ (\c \ set (coeffs p). P c)" if "P 0" using that by (auto simp add: coeffs_def) (metis atLeastLessThan_iff coeff_eq_0 not_less_iff_gr_or_eq zero_le) instantiation poly :: ("{zero, equal}") equal begin definition [code]: "HOL.equal (p::'a poly) q \ 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 \ True" by (fact equal_refl) definition is_zero :: "'a::zero poly \ bool" where [code]: "is_zero p \ List.null (coeffs p)" lemma is_zero_null [code_abbrev]: "is_zero p \ p = 0" by (simp add: is_zero_def null_def) subsubsection \Reconstructing the polynomial from the list\ \ \contributed by Sebastiaan J.C. Joosten and René Thiemann\ definition poly_of_list :: "'a::comm_monoid_add list \ '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 \Fold combinator for polynomials\ definition fold_coeffs :: "('a::zero \ 'b \ 'b) \ 'a poly \ 'b \ '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 \ fold_coeffs f (pCons a p) = f a \ 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 \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_not_0_0_eq [simp]: "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: \'a::comm_semiring_0 poly \ 'a \ 'a\ where \poly p a = horner_sum id a (coeffs p)\ lemma poly_eq_fold_coeffs: \poly p = fold_coeffs (\a f x. a + x * f x) p (\x. 0)\ by (induction p) (auto simp add: fun_eq_iff poly_def) 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 \ a = 0") (auto simp add: poly_def) lemma poly_altdef: "poly p x = (\i\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 * (\i\degree p. coeff p i * x ^ i) = coeff ?p' 0 * x^0 + (\i\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 \p \ 0\, 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 \Monomials\ lift_definition monom :: "'a \ nat \ 'a::zero poly" is "\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 = [:a:]" 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 \ a = 0" by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" by (metis coeff_monom leading_coeff_0_iff) 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 \ 0 \ fold_coeffs f (monom a n) = f 0 ^^ n \ 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_eq_fold_coeffs) lemma monom_eq_iff': "monom c n = monom d m \ c = d \ (c = 0 \ n = m)" by (auto simp: poly_eq_iff) lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) subsection \Leading coefficient\ abbreviation lead_coeff:: "'a::zero poly \ 'a" where "lead_coeff p \ coeff p (degree p)" lemma lead_coeff_pCons[simp]: "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" "p = 0 \ 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) lemma last_coeffs_eq_coeff_degree: "last (coeffs p) = lead_coeff p" if "p \ 0" using that by (simp add: coeffs_def) subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add begin lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n + coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\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 \ 'a poly \ 'a poly" is "\p q n. coeff p n - coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\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 \ 'a poly" is "\p n. - coeff p n" proof - fix p :: "'a poly" show "\\<^sub>\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) \ max (degree p) (degree q)" by (rule degree_le) (auto simp add: coeff_eq_0) lemma degree_add_le: "degree p \ n \ degree q \ n \ degree (p + q) \ n" by (auto intro: order_trans degree_add_le_max) lemma degree_add_less: "degree p < n \ degree q < n \ degree (p + q) < n" by (auto intro: le_less_trans degree_add_le_max) lemma degree_add_eq_right: assumes "degree p < degree q" shows "degree (p + q) = degree q" proof (cases "q = 0") case False show ?thesis proof (rule order_antisym) show "degree (p + q) \ degree q" by (simp add: assms degree_add_le order.strict_implies_order) show "degree q \ degree (p + q)" by (simp add: False assms coeff_eq_0 le_degree) qed qed (use assms in auto) lemma degree_add_eq_left: "degree q < degree p \ 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 \ 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) \ 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 \ n \ degree q \ n \ degree (p - q) \ 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 \ degree q < n \ 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 (\x\A. p x) i = (\x\A. coeff (p x) i)" by (induct A rule: infinite_finite_induct) simp_all lemma monom_sum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" by (rule poly_eqI) (simp add: coeff_sum) fun plus_coeffs :: "'a::comm_monoid_add list \ 'a list \ '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 \ 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" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q) (simp add: algebra_simps) qed auto 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 (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_sum_le: "finite S \ (\p. p \ S \ degree (f p) \ n) \ degree (sum f S) \ 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) \ n" "degree (f p) \ 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 \ n" shows "(\i\n. monom (coeff p i) i) = p" proof - have eq: "\i. {..n} \ {i} = (if i \ 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="\x. x * a" for a]) qed lemma poly_as_sum_of_monoms: "(\i\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 \Multiplication by a constant, polynomial multiplication and the unit polynomial\ lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" is "\a p n. a * coeff p n" proof - fix a :: 'a and p :: "'a poly" show "\\<^sub>\ 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) \ 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 ((*) 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 \ a = 0 \ 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 \ times a = HOL.eq (0::'a)" if "a \ 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 \ 0" shows "smult a p = smult b q \ smult (a / b) p = q" (is "?lhs \ ?rhs") proof assume ?lhs also from assms have "smult (inverse b) \ = 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 (\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 \ 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 \ 0" and "q \ 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 \p \ 0\ \q \ 0\ have "coeff p (degree p) * coeff q (degree q) \ 0" by simp finally have "\n. coeff (p * q) n \ 0" .. then show "p * q \ 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 = (\i\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) \ degree p + degree q" proof (rule degree_le) show "\i>degree p + degree q. coeff (p * q) i = 0" by (induct p) (simp_all add: coeff_eq_0 coeff_pCons split: nat.split) qed 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 lift_definition one_poly :: "'a poly" is "\n. of_bool (n = 0)" by (rule MOST_SucD) simp lemma coeff_1 [simp]: "coeff 1 n = of_bool (n = 0)" by (simp add: one_poly.rep_eq) lemma one_pCons: "1 = [:1:]" by (simp add: poly_eq_iff coeff_pCons split: nat.splits) lemma pCons_one: "[:1:] = 1" by (simp add: one_pCons) instance by standard (simp_all add: one_pCons) end lemma poly_1 [simp]: "poly 1 x = 1" by (simp add: one_pCons) lemma one_poly_eq_simps [simp]: "1 = [:1:] \ True" "[:1:] = 1 \ True" by (simp_all add: one_pCons) lemma degree_1 [simp]: "degree 1 = 0" by (simp add: one_pCons) lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]" by (simp add: one_pCons) lemma smult_one [simp]: "smult c 1 = [:c:]" by (simp add: one_pCons) lemma monom_eq_1 [simp]: "monom 1 0 = 1" by (simp add: monom_0 one_pCons) lemma monom_eq_1_iff: "monom c n = 1 \ c = 1 \ n = 0" using monom_eq_const_iff [of c n 1] by auto lemma monom_altdef: "monom c n = smult c ([:0, 1:] ^ n)" by (induct n) (simp_all add: monom_0 monom_Suc) 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 degree_power_le: "degree (p ^ n) \ 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_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 (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_prod_sum_le: "finite S \ degree (prod f S) \ sum (degree \ 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 (\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 = (\i\k. (if n = i then c else 0) * coeff p (k - i))" by (simp add: coeff_mult) also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" by (intro sum.cong) simp_all also have "\ = (if k < n then 0 else c * coeff p (k - n))" by simp finally show ?thesis . qed lemma monom_1_dvd_iff': "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" have "\\<^sub>\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 \Mapping polynomials\ definition map_poly :: "('a :: zero \ 'b :: zero) \ 'a poly \ '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 \ map_poly f 1 = 1" by (simp add: map_poly_def one_pCons) 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 ((=) 0) (map f (coeffs p))" by (simp add: map_poly_def) lemma coeffs_map_poly': assumes "\x. x \ 0 \ f x \ 0" shows "coeffs (map_poly f p) = map f (coeffs p)" using assms by (auto simp add: coeffs_map_poly strip_while_idem_iff last_coeffs_eq_coeff_degree no_trailing_unfold last_map) lemma set_coeffs_map_poly: "(\x. f x = 0 \ x = 0) \ set (coeffs (map_poly f p)) = f ` set (coeffs p)" by (simp add: coeffs_map_poly') lemma degree_map_poly: assumes "\x. x \ 0 \ f x \ 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" "\x. x \ set (coeffs p) \ x \ 0 \ f x \ 0" shows "map_poly f p = 0 \ 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 "\ = 0 \ coeff p n = 0" proof (cases "n < length (coeffs p)") case True then have "coeff p n \ set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc) with assms show "f (coeff p n) = 0 \ 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""\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 \ 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 (\x. x) p = p" by (simp add: map_poly_def) lemma map_poly_cong: assumes "(\x. x \ set (coeffs p) \ 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 \ 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 "\x. x \ set (coeffs p) \ f x = x" shows "map_poly f p = p" using map_poly_cong[OF assms, of _ id] by simp lemma map_poly_idI': assumes "\x. x \ set (coeffs p) \ 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 (\x. c * x) p" by (intro poly_eqI) (simp_all add: coeff_map_poly) subsection \Conversions\ lemma of_nat_poly: "of_nat n = [:of_nat n:]" by (induct n) (simp_all add: one_pCons) lemma of_nat_monom: "of_nat n = monom (of_nat n) 0" by (simp add: of_nat_poly monom_0) 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" by (simp add: of_nat_poly) lemma of_int_poly: "of_int k = [:of_int k:]" by (simp only: of_int_of_nat of_nat_poly) simp lemma of_int_monom: "of_int k = monom (of_int k) 0" by (simp add: of_int_poly monom_0) 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" by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" proof - have "numeral n = of_nat (numeral n)" by simp also have "\ = [:of_nat (numeral n):]" by (simp add: of_nat_poly) finally show ?thesis by simp qed lemma numeral_monom: "numeral n = monom (numeral n) 0" by (simp add: numeral_poly monom_0) lemma degree_numeral [simp]: "degree (numeral n) = 0" by (simp add: numeral_poly) lemma lead_coeff_numeral [simp]: "lead_coeff (numeral n) = numeral n" by (simp add: numeral_poly) subsection \Lemmas about divisibility\ 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 \ a \ 0 \ p dvd q" for a :: "'a::field" by (drule dvd_smult [where a="inverse a"]) simp lemma dvd_smult_iff: "a \ 0 \ p dvd smult a q \ 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 \ a \ 0 \ 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 \ (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 \ c dvd 1 \ p dvd 1" proof - have "smult c p = [:c:] * p" by simp also have "\ dvd 1 \ c dvd 1 \ 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 "\ = 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_pCons ac_simps) then have "[:c:] dvd 1" by (rule dvdI) from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp qed finally show ?thesis . qed subsection \Polynomials form an integral domain\ instance poly :: (idom) idom .. instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0 by standard (auto simp add: of_nat_poly intro: injI) lemma degree_mult_eq: "p \ 0 \ q \ 0 \ 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 dvd_imp_degree: \degree x \ degree y\ if \x dvd y\ \x \ 0\ \y \ 0\ for x y :: \'a::{comm_semiring_1,semiring_no_zero_divisors} poly\ proof - from \x dvd y\ obtain z where \y = x * z\ .. with \x \ 0\ \y \ 0\ show ?thesis by (simp add: degree_mult_eq) qed lemma degree_prod_eq_sum_degree: fixes A :: "'a set" and f :: "'a \ 'b::idom poly" assumes f0: "\i\A. f i \ 0" shows "degree (\i\A. (f i)) = (\i\A. degree (f i))" using assms by (induction A rule: infinite_finite_induct) (auto simp: degree_mult_eq) lemma degree_mult_eq_0: "degree (p * q) = 0 \ p = 0 \ q = 0 \ (p \ 0 \ q \ 0 \ degree p = 0 \ degree q = 0)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (auto simp: degree_mult_eq) lemma degree_power_eq: "p \ 0 \ degree ((p :: 'a :: idom poly) ^ n) = n * degree p" by (induction n) (simp_all add: degree_mult_eq) lemma degree_mult_right_le: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "q \ 0" shows "degree p \ 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 \ q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) lemma dvd_imp_degree_le: "p dvd q \ q \ 0 \ degree p \ 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 \ degree q \ 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 \ (\n. c dvd coeff p n)" proof (cases "c = 0 \ p = 0") case True then show ?thesis by (auto intro!: poly_eqI) next case False show ?thesis proof assume "[:c:] dvd p" then show "\n. c dvd coeff p n" by (auto simp: coeffs_def) next assume *: "\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 (\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:] \ 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 \ q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) lemma lead_coeff_prod: "lead_coeff (prod f A) = (\x\A. lead_coeff (f x))" for f :: "'a \ 'b::{comm_semiring_1, semiring_no_zero_divisors} poly" by (induction A rule: infinite_finite_induct) (auto simp: lead_coeff_mult) 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 \ = 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 \Polynomials form an ordered integral domain\ definition pos_poly :: "'a::linordered_semidom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" lemma pos_poly_pCons: "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" by (simp add: pos_poly_def) lemma not_pos_poly_0 [simp]: "\ pos_poly 0" by (simp add: pos_poly_def) lemma pos_poly_add: "pos_poly p \ pos_poly q \ pos_poly (p + q)" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q; force simp add: pos_poly_pCons add_pos_pos) qed auto lemma pos_poly_mult: "pos_poly p \ pos_poly q \ pos_poly (p * q)" by (simp add: pos_poly_def coeff_degree_mult) lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" for p :: "'a::linordered_idom poly" by (induct p) (auto simp: pos_poly_pCons) lemma pos_poly_coeffs [code]: "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?lhs \ ?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 \ 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 \ pos_poly (y - x)" definition "x \ y \ x = y \ pos_poly (y - x)" definition "\x::'a poly\ = (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 \ x \ y \ \ y \ x" unfolding less_eq_poly_def less_poly_def using pos_poly_add by force then show "x \ y \ y \ x \ x = y" using less_eq_poly_def less_poly_def by force show "x \ x" by (simp add: less_eq_poly_def) show "x \ y \ y \ z \ x \ z" using less_eq_poly_def pos_poly_add by fastforce show "x \ y \ z + x \ z + y" by (simp add: less_eq_poly_def) show "x \ y \ y \ x" unfolding less_eq_poly_def using pos_poly_total [of "x - y"] by auto show "x < y \ 0 < z \ z * x < z * y" by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) show "\x\ = (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 \TODO: Simplification rules for comparisons\ subsection \Synthetic division and polynomial roots\ subsubsection \Synthetic division\ text \Synthetic division is simply division by the linear polynomial \<^term>\x - c\.\ definition synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" where "synthetic_divmod p c = fold_coeffs (\a (q, r). (pCons r q, a + c * r)) p (0, 0)" definition synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ '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 = (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" by (cases "p = 0 \ 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 \ 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 \ 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 \ r = poly p c \ q = synthetic_div p c" proof (induction p arbitrary: q r) case 0 then show ?case using synthetic_div_unique_lemma by fastforce next case (pCons a p) then show ?case by (cases q; force) qed 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 \Polynomial roots\ lemma poly_eq_0_iff_dvd: "poly p c = 0 \ [:- c, 1:] dvd p" (is "?lhs \ ?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 \ poly p (- c) = 0" for c :: "'a::comm_ring_1" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: "p \ 0 \ finite {x. poly p x = 0}" for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly" proof (induct n \ "degree p" arbitrary: p) case 0 then obtain a where "a \ 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 "\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 \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from this \k \ 0\ 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 \ p = q" (is "?lhs \ ?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 \ p = 0" for p :: "'a poly" proof (cases "p = 0") case False then show ?thesis by (auto simp add: infinite_UNIV_char_0 dest: poly_roots_finite) qed auto from \?lhs\ and this [of "p - q"] show ?rhs by auto qed lemma poly_all_0_iff_0: "(\x. poly p x = 0) \ 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 \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" for a :: "'a::comm_semiring_1" proof (induct n) case (Suc n) have "degree ([:a, 1:] ^ n) \ 1 * n" by (metis One_nat_def degree_pCons_eq_if degree_power_le one_neq_zero one_pCons) then have "coeff ([:a, 1:] ^ n) (Suc n) = 0" by (simp add: coeff_eq_0) then show ?case using Suc.hyps by fastforce qed auto lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" for a :: "'a::comm_semiring_1" proof (rule order_antisym) show "degree ([:a, 1:] ^ n) \ n" by (metis One_nat_def degree_pCons_eq_if degree_power_le mult.left_neutral one_neq_zero one_pCons) qed (simp add: coeff_linear_power le_degree) lemma order_1: "[:-a, 1:] ^ order a p dvd p" proof (cases "p = 0") case False show ?thesis proof (cases "order a p") case (Suc n) then show ?thesis by (metis lessI not_less_Least order_def) qed auto qed auto lemma order_2: assumes "p \ 0" shows "\ [:-a, 1:] ^ Suc (order a p) dvd p" proof - have False if "[:- a, 1:] ^ Suc (degree p) dvd p" using dvd_imp_degree_le [OF that] by (metis Suc_n_not_le_n assms degree_linear_power) then show ?thesis unfolding order_def by (metis (no_types, lifting) LeastI) qed lemma order: "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" by (rule conjI [OF order_1 order_2]) lemma order_degree: assumes p: "p \ 0" shows "order a p \ 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 "\ \ degree p" by (rule dvd_imp_degree_le) finally show ?thesis . qed lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (metis One_nat_def order_2 poly_eq_0_iff_dvd power_one_right) show "?rhs \ ?lhs" by (meson dvd_power dvd_trans neq0_conv order_1 poly_0 poly_eq_0_iff_dvd) qed lemma order_0I: "poly p a \ 0 \ 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" "\ [:-a, 1:] ^ Suc n dvd p" shows "order a p = n" unfolding Polynomial.order_def by (metis (mono_tags, lifting) Least_equality assms not_less_eq_eq power_le_dvd) lemma order_mult: assumes "p * q \ 0" shows "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: "\u. t dvd u \ poly u a = 0" by (simp add: t_def dvd_iff_poly_eq_0) have dvd: "t ^ i dvd p" "t ^ j dvd q" and "\ t ^ Suc i dvd p" "\ t ^ Suc j dvd q" using assms i_def j_def order_1 order_2 t_def by auto then have "\ t ^ Suc(i + j) dvd p * q" by (elim dvdE) (simp add: power_add t_dvd_iff) moreover have "t ^ (i + j) dvd p * q" using dvd by (simp add: mult_dvd_mono power_add) ultimately show "order a (p * q) = i + j" using order_unique_lemma t_def by blast qed lemma order_smult: assumes "c \ 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 \ = 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 text \Next three 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_uminus[simp]: "order x (-p) = order x p" by (metis neg_equal_0_iff_equal order_smult smult_1_left smult_minus_left) 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, opaque_lifting) 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) \ degree ([:- a, 1:])" by (rule dvd_imp_degree_le) auto then show False by auto next fix y assume *: "\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" show "1 \ y" proof (rule ccontr) assume "\ 1 \ 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 \ 0 \ 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 \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" by (auto simp: order_mult) text \Now justify the standard squarefree decomposition, i.e. \f / gcd f f'\.\ lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" by (meson dvd_0_right not_less_eq_eq order_1 order_2 power_le_dvd) lemma order_decomp: assumes "p \ 0" shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" proof - from assms have *: "[:- a, 1:] ^ order a p dvd p" and **: "\ [:- 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 "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" by simp then have "\ [:- 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 "\ [:- a, 1:] dvd q" by auto with q show ?thesis by blast qed lemma monom_1_dvd_iff: "p \ 0 \ monom 1 n dvd p \ n \ order 0 p" using order_divides[of 0 n p] by (simp add: monom_altdef) subsection \Additional induction rules on polynomials\ text \ An induction rule for induction over the roots of a polynomial with a certain property. (e.g. all positive roots) \ lemma poly_root_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "Q 0" and "\p. (\a. P a \ poly p a \ 0) \ Q p" and "\a p. P a \ Q p \ 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 "\a. P a \ 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 \ 0" by auto have "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) (simp_all add: \q \ 0\) 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 ((=) a) (replicate n a @ ys) = dropWhile ((=) 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 \ An induction rule for simultaneous induction over two polynomials, prepending one coefficient in each step. \ lemma poly_induct2 [case_names 0 pCons]: assumes "P 0 0" "\a p b q. P p q \ 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 \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" notation pcompose (infixl "\\<^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 \ 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_pCons 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) \ degree p * degree q" proof (induction p) case (pCons a p) then show ?case proof (clarsimp simp add: pcompose_pCons) assume "degree (p \\<^sub>p q) \ degree p * degree q" "p \ 0" then have "degree (q * p \\<^sub>p q) \ degree q + degree p * degree q" by (meson add_le_cancel_left degree_mult_le dual_order.trans pCons.IH) then show "degree ([:a:] + q * p \\<^sub>p q) \ degree q + degree p * degree q" by (simp add: degree_add_le) qed qed auto 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 "\ + 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 (\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 (\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 \ pcompose p q = 0" by (auto simp add: degree_mult_eq_0) moreover have False if "pcompose p q = 0" "degree q \ 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 \pcompose p q = 0\ \p \ 0\ 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 "\ \ 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 \ 0" "q \ 0" "pcompose p q \ 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 \q\0\ \pcompose p q\0\] 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 \p = [:a:]\ 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 _ \degree q > 0\] 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 \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" by (rule lead_coeff_add_le) then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" by (simp add: pcompose_pCons) also have "\ = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "\ = lead_coeff p * lead_coeff q ^ (degree p + 1)" by (auto simp: mult_ac) finally show ?thesis by auto qed qed subsection \Closure properties of coefficients\ context fixes R :: "'a :: comm_semiring_1 set" assumes R_0: "0 \ R" assumes R_plus: "\x y. x \ R \ y \ R \ x + y \ R" assumes R_mult: "\x y. x \ R \ y \ R \ x * y \ R" begin lemma coeff_mult_semiring_closed: assumes "\i. coeff p i \ R" "\i. coeff q i \ R" shows "coeff (p * q) i \ R" proof - have R_sum: "sum f A \ R" if "\x. x \ A \ f x \ R" for A and f :: "nat \ 'a" using that by (induction A rule: infinite_finite_induct) (auto intro: R_0 R_plus) show ?thesis unfolding coeff_mult by (auto intro!: R_sum R_mult assms) qed lemma coeff_pcompose_semiring_closed: assumes "\i. coeff p i \ R" "\i. coeff q i \ R" shows "coeff (pcompose p q) i \ R" using assms(1) proof (induction p arbitrary: i) case (pCons a p i) have [simp]: "a \ R" using pCons.prems[of 0] by auto have "coeff p i \ R" for i using pCons.prems[of "Suc i"] by auto hence "coeff (p \\<^sub>p q) i \ R" for i using pCons.prems by (intro pCons.IH) thus ?case by (auto simp: pcompose_pCons coeff_pCons split: nat.splits intro!: assms R_plus coeff_mult_semiring_closed) qed auto end subsection \Shifting polynomials\ definition poly_shift :: "nat \ 'a::zero poly \ 'a poly" where "poly_shift n p = Abs_poly (\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 "\k>m. coeff p k = 0" by (auto simp: MOST_nat) then have "\k>m. coeff p (k + n) = 0" by auto then have "\\<^sub>\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 = (\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 \ 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 \Truncating polynomials\ definition poly_cutoff where "poly_cutoff n p = Abs_poly (\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 ((=) 0) (take n (coeffs p))" proof (cases "strip_while ((=) 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 ((=) 0) (strip_while ((=) 0) (take n (coeffs p)))" by simp with False have "last (strip_while ((=) 0) (take n (coeffs p))) \ 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 \Reflecting polynomials\ definition reflect_poly :: "'a::zero poly \ 'a poly" where "reflect_poly p = Poly (rev (coeffs p))" lemma coeffs_reflect_poly [code abstract]: "coeffs (reflect_poly p) = rev (dropWhile ((=) 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_pCons) 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 \ p = 0" by (simp add: coeff_reflect_poly) lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly poly_0_coeff_0) lemma reflect_poly_pCons': "p \ 0 \ 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 \ 0 \ 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 \ 0 \ 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) \ degree p" by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) lemma reflect_poly_pCons: "a \ 0 \ 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 \ 0 \ 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 \ q = 0") case False then have [simp]: "p \ 0" "q \ 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 \ degree (p * q)") case True define A where "A = {..i} \ {i - degree q..degree p}" define B where "B = {..degree p} \ {degree p - i..degree (p*q) - i}" let ?f = "\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 "\ = (\j\degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" by (simp add: coeff_mult) also have "\ = (\j\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 "\ = (\j\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 "\ = (\j\i. if j \ {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 "\ = 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 (\x. reflect_poly (f x)) A" for f :: "_ \ _::{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 \ 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 \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ '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_pCons 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} \ 'a list \ '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 \ '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 = [] \ f * x = 0") (auto simp: cCons_def) next case n: (Suc m) show ?thesis proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ 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 "\ = (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 = [] \ g + of_nat m = 0 \ 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 \ 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 \auto simp: field_simps\) ultimately show ?thesis by simp qed qed qed 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 (\i. coeff p (Suc i)) [0.. degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof (cases "degree p") case 0 then show ?thesis by (metis degree_eq_zeroE pderiv.simps) next case (Suc n) then show ?thesis using coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff by (metis coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed lemma degree_pderiv: "degree (pderiv p) = degree p - 1" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof - have "degree p - 1 \ degree (pderiv p)" proof (cases "degree p") case (Suc n) then show ?thesis by (metis coeff_pderiv degree_0 diff_Suc_1 le_degree leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed auto moreover have "\i>degree p - 1. coeff (pderiv p) i = 0" by (simp add: coeff_eq_0 coeff_pderiv) ultimately show ?thesis using order_antisym [OF degree_le] by blast qed lemma not_dvd_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" assumes "degree p \ 0" shows "\ 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 \ 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 \ 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" proof (induction n) case (Suc n) then show ?case by (simp add: pderiv_mult smult_add_left algebra_simps) qed auto lemma pderiv_pcompose: "pderiv (pcompose p q) = pcompose (pderiv p) q * pderiv q" by (induction p rule: pCons_induct) (auto simp: pcompose_pCons pderiv_add pderiv_mult pderiv_pCons pcompose_add algebra_simps) lemma pderiv_prod: "pderiv (prod f (as)) = (\a\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" "\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 \ as" for b proof - from \a \ as\ 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 (\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 \ DERIV (\x. a + f x :: 'a::real_normed_field) x :> D" by (rule DERIV_cong, rule DERIV_add) auto lemma poly_DERIV [simp]: "DERIV (\x. poly p x) x :> poly (pderiv p) x" by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) lemma poly_isCont[simp]: fixes x::"'a::real_normed_field" shows "isCont (\x. poly p x) x" by (rule poly_DERIV [THEN DERIV_isCont]) lemma tendsto_poly [tendsto_intros]: "(f \ a) F \ ((\x. poly p (f x)) \ poly p a) F" for f :: "_ \ 'a::real_normed_field" by (rule isCont_tendsto_compose [OF poly_isCont]) lemma continuous_within_poly: "continuous (at z within s) (poly p)" for z :: "'a::{real_normed_field}" by (simp add: continuous_within tendsto_poly) lemma continuous_poly [continuous_intros]: "continuous F f \ continuous F (\x. poly p (f x))" for f :: "_ \ 'a::real_normed_field" unfolding continuous_def by (rule tendsto_poly) lemma continuous_on_poly [continuous_intros]: fixes p :: "'a :: {real_normed_field} poly" assumes "continuous_on A f" shows "continuous_on A (\x. poly p (f x))" by (metis DERIV_continuous_on assms continuous_on_compose2 poly_DERIV subset_UNIV) text \Consequences of the derivative theorem above.\ lemma poly_differentiable[simp]: "(\x. poly p x) differentiable (at x)" for x :: real by (simp add: real_differentiable_def) (blast intro: poly_DERIV) lemma poly_IVT_pos: "a < b \ poly p a < 0 \ 0 < poly p b \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using IVT [of "poly p" a 0 b] by (auto simp add: order_le_less) lemma poly_IVT_neg: "a < b \ 0 < poly p a \ poly p b < 0 \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using poly_IVT_pos [where p = "- p"] by simp lemma poly_IVT: "a < b \ poly p a * poly p b < 0 \ \x>a. x < b \ 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 \ \x. a < x \ x < b \ poly p b - poly p a = (b - a) * poly (pderiv p) x" for a b :: real by (simp add: MVT2) lemma poly_MVT': fixes a b :: real assumes "{min a b..max a b} \ A" shows "\x\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] obtain x where "a < x" "x < b" "poly p b - poly p a = (b - a) * poly (pderiv p) x" by auto then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) next case greater from poly_MVT[OF greater, of p] obtain x where "b < x" "x < a" "poly p a - poly p b = (a - b) * poly (pderiv p) x" by auto 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 "\n. \ x \ n. poly p x \ 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 \ 0" "p = 0" | "p \ 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: "\x\n1. lead_coeff p \ poly p x" by auto from pCons(3) \p \ 0\ have gt_0: "lead_coeff p > 0" by auto define n where "n = max n1 (1 + \a\ / lead_coeff p)" have "lead_coeff (pCons a p) \ poly (pCons a p) x" if "n \ x" for x proof - from gte_lcoeff that have "lead_coeff p \ poly p x" by (auto simp: n_def) with gt_0 have "\a\ / lead_coeff p \ \a\ / poly p x" and "poly p x > 0" by (auto intro: frac_le) with \n \ x\[unfolded n_def] have "x \ 1 + \a\ / poly p x" by auto with \lead_coeff p \ poly p x\ \poly p x > 0\ \p \ 0\ show "lead_coeff (pCons a p) \ 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 \ 0" and pe: "p = [:- a, 1:] ^ n * q" and nd: "\ [:- a, 1:] dvd q" shows "n = Suc (order a (pderiv p))" proof - from assms have "pderiv ([:- a, 1:] ^ n * q) \ 0" by auto from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" by (cases n) auto have "order a (pderiv ([:- a, 1:] ^ Suc n' * q)) = n'" proof (rule order_unique_lemma) show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 proof (rule dvd_add) show "[:- a, 1:] ^ n' dvd [:- a, 1:] ^ Suc n' * pderiv q" by (metis dvdI dvd_mult2 power_Suc2) show "[:- a, 1:] ^ n' dvd smult (of_nat (Suc n')) (q * [:- a, 1:] ^ n')" by (metis dvd_smult dvd_triv_right) qed have "k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" for k l by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) then show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 by (metis nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) qed then show ?thesis by (metis \n = Suc n'\ pe) qed lemma order_pderiv: "order a p = Suc (order a (pderiv p))" if "pderiv p \ 0" "order a p \ 0" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False obtain q where "p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" using False order_decomp by blast then show ?thesis using lemma_order_pderiv that by blast qed (use that in auto) lemma poly_squarefree_decomp_order: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 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: "\ ?thesis" from \pderiv p \ 0\ have "p \ 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 \ 0" by (auto split: if_splits) from \pderiv p \ 0\ \pderiv p = e * d\ have oapp: "order a (pderiv p) = order a e + order a d" by (simp add: order_mult) from \pderiv p \ 0\ \order a p \ 0\ have oap: "order a p = Suc (order a (pderiv p))" by (rule order_pderiv) from \p \ 0\ \p = q * d\ have "d \ 0" by simp have "[:- a, 1:] ^ order a (pderiv p) dvd r * p" by (metis dvd_trans dvd_triv_right oap order_1 power_Suc) then have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" by (simp add: d order_1) with \d \ 0\ have "order a (pderiv p) \ order a d" by (simp add: order_divides) show ?thesis using \order a p = order a q + order a d\ and oapp oap and \order a (pderiv p) \ order a d\ by auto qed lemma poly_squarefree_decomp_order2: "pderiv p \ 0 \ p = q * d \ pderiv p = e * d \ d = r * p + s * pderiv p \ \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 \ 0 \ order a p \ 0 \ order a (pderiv p) = n \ order a p = Suc n" for p :: "'a::field_char_0 poly" by (auto dest: order_pderiv) definition rsquarefree :: "'a::idom poly \ bool" where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" lemma pderiv_iszero: "pderiv p = 0 \ \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 \ (\a. \ (poly p a = 0 \ poly (pderiv p) a = 0))" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False show ?thesis proof (cases "pderiv p = 0") case True with \p \ 0\ pderiv_iszero show ?thesis by (force simp add: order_0I rsquarefree_def) next case False with \p \ 0\ order_pderiv2 show ?thesis by (force simp add: rsquarefree_def order_root) qed qed (simp add: rsquarefree_def) lemma poly_squarefree_decomp: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and "p = q * d" and "pderiv p = e * d" and "d = r * p + s * pderiv p" shows "rsquarefree q \ (\a. poly q a = 0 \ poly p a = 0)" proof - from \pderiv p \ 0\ have "p \ 0" by auto with \p = q * d\ have "q \ 0" by simp from assms have "\a. order a q = (if order a p = 0 then 0 else 1)" by (rule poly_squarefree_decomp_order2) with \p \ 0\ \q \ 0\ show ?thesis by (simp add: rsquarefree_def order_root) qed subsection \Algebraic numbers\ text \ 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. \ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" lemma algebraicI: "(\i. coeff p i \ \) \ p \ 0 \ poly p x = 0 \ algebraic x" unfolding algebraic_def by blast lemma algebraicE: assumes "algebraic x" obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast lemma algebraic_altdef: "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" for p :: "'a::field_char_0 poly" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \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 \ 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 \ \" for n proof (cases "n \ 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 * \ = of_rat (of_int (a*d) / of_int b)" by (simp add: of_rat_mult of_rat_divide) also from nz True have "b \ 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 \ (\ :: rat set)" by (rule of_int_divide_in_Ints) then have "of_rat (of_int (a * d) / of_int b) \ \" 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') \ {0<..}" unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) then have "d \ 0" unfolding d_def by (induct cs') simp_all with nz have "p' \ 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 \ \" "poly p x = 0" "p \ 0" for i by (force simp: algebraic_def) moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp ultimately show "\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0" by auto qed subsection \Algebraic integers\ inductive algebraic_int :: "'a :: field \ bool" where "\lead_coeff p = 1; \i. coeff p i \ \; poly p x = 0\ \ algebraic_int x" lemma algebraic_int_altdef_ipoly: fixes x :: "'a :: field_char_0" shows "algebraic_int x \ (\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" proof assume "algebraic_int x" then obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto elim: algebraic_int.cases) define the_int where "the_int = (\x::'a. THE r. x = of_int r)" define p' where "p' = map_poly the_int p" have of_int_the_int: "of_int (the_int x) = x" if "x \ \" for x unfolding the_int_def by (rule sym, rule theI') (insert that, auto simp: Ints_def) have the_int_0_iff: "the_int x = 0 \ x = 0" if "x \ \" for x using of_int_the_int[OF that] by auto have [simp]: "the_int 0 = 0" by (subst the_int_0_iff) auto have "map_poly of_int p' = map_poly (of_int \ the_int) p" by (simp add: p'_def map_poly_map_poly) also from p of_int_the_int have "\ = p" by (subst poly_eq_iff) (auto simp: coeff_map_poly) finally have p_p': "map_poly of_int p' = p" . show "(\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" proof (intro exI conjI notI) from p show "poly (map_poly of_int p') x = 0" by (simp add: p_p') next show "lead_coeff p' = 1" using p by (simp flip: p_p' add: degree_map_poly coeff_map_poly) qed next assume "\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1" then obtain p where p: "poly (map_poly of_int p) x = 0" "lead_coeff p = 1" by auto define p' where "p' = (map_poly of_int p :: 'a poly)" from p have "lead_coeff p' = 1" "poly p' x = 0" "\i. coeff p' i \ \" by (auto simp: p'_def coeff_map_poly degree_map_poly) thus "algebraic_int x" by (intro algebraic_int.intros) qed theorem rational_algebraic_int_is_int: assumes "algebraic_int x" and "x \ \" shows "x \ \" proof - from assms(2) obtain a b where ab: "b > 0" "Rings.coprime a b" and x_eq: "x = of_int a / of_int b" by (auto elim: Rats_cases') from \b > 0\ have [simp]: "b \ 0" by auto from assms(1) obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) define q :: "'a poly" where "q = [:-of_int a, of_int b:]" have "poly q x = 0" "q \ 0" "\i. coeff q i \ \" by (auto simp: x_eq q_def coeff_pCons split: nat.splits) define n where "n = degree p" have "n > 0" using p by (intro Nat.gr0I) (auto simp: n_def elim!: degree_eq_zeroE) have "(\i \" using p by auto then obtain R where R: "of_int R = (\i = (\i\n. coeff p i * x ^ i * of_int b ^ n)" by (simp add: poly_altdef n_def sum_distrib_right) also have "\ = (\i\n. coeff p i * of_int (a ^ i * b ^ (n - i)))" by (intro sum.cong) (auto simp: x_eq field_simps simp flip: power_add) also have "{..n} = insert n {..n > 0\ by auto also have "(\i\\. coeff p i * of_int (a ^ i * b ^ (n - i))) = coeff p n * of_int (a ^ n) + (\iii = of_int (b * R)" by (simp add: R sum_distrib_left sum_distrib_right mult_ac) finally have "of_int (a ^ n) = (-of_int (b * R) :: 'a)" by (auto simp: add_eq_0_iff) hence "a ^ n = -b * R" by (simp flip: of_int_mult of_int_power of_int_minus) hence "b dvd a ^ n" by simp with \Rings.coprime a b\ have "b dvd 1" by (meson coprime_power_left_iff dvd_refl not_coprimeI) with x_eq and \b > 0\ show ?thesis by auto qed lemma algebraic_int_imp_algebraic [dest]: "algebraic_int x \ algebraic x" by (auto simp: algebraic_int.simps algebraic_def) lemma int_imp_algebraic_int: assumes "x \ \" shows "algebraic_int x" proof show "\i. coeff [:-x, 1:] i \ \" using assms by (auto simp: coeff_pCons split: nat.splits) qed auto lemma algebraic_int_0 [simp, intro]: "algebraic_int 0" and algebraic_int_1 [simp, intro]: "algebraic_int 1" and algebraic_int_numeral [simp, intro]: "algebraic_int (numeral n)" and algebraic_int_of_nat [simp, intro]: "algebraic_int (of_nat k)" and algebraic_int_of_int [simp, intro]: "algebraic_int (of_int m)" by (simp_all add: int_imp_algebraic_int) lemma algebraic_int_ii [simp, intro]: "algebraic_int \" proof show "poly [:1, 0, 1:] \ = 0" by simp qed (auto simp: coeff_pCons split: nat.splits) lemma algebraic_int_minus [intro]: assumes "algebraic_int x" shows "algebraic_int (-x)" proof - from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) define s where "s = (if even (degree p) then 1 else -1 :: 'a)" define q where "q = Polynomial.smult s (pcompose p [:0, -1:])" have "lead_coeff q = s * lead_coeff (pcompose p [:0, -1:])" by (simp add: q_def) also have "lead_coeff (pcompose p [:0, -1:]) = lead_coeff p * (- 1) ^ degree p" by (subst lead_coeff_comp) auto finally have "poly q (-x) = 0" and "lead_coeff q = 1" using p by (auto simp: q_def poly_pcompose s_def) moreover have "coeff q i \ \" for i proof - have "coeff (pcompose p [:0, -1:]) i \ \" using p by (intro coeff_pcompose_semiring_closed) (auto simp: coeff_pCons split: nat.splits) thus ?thesis by (simp add: q_def s_def) qed ultimately show ?thesis by (auto simp: algebraic_int.simps) qed lemma algebraic_int_minus_iff [simp]: "algebraic_int (-x) \ algebraic_int (x :: 'a :: field_char_0)" using algebraic_int_minus[of x] algebraic_int_minus[of "-x"] by auto lemma algebraic_int_inverse [intro]: assumes "poly p x = 0" and "\i. coeff p i \ \" and "coeff p 0 = 1" shows "algebraic_int (inverse x)" proof from assms have [simp]: "x \ 0" by (auto simp: poly_0_coeff_0) show "poly (reflect_poly p) (inverse x) = 0" using assms by (simp add: poly_reflect_poly_nz) qed (use assms in \auto simp: coeff_reflect_poly\) lemma algebraic_int_root: assumes "algebraic_int y" and "poly p x = y" and "\i. coeff p i \ \" and "lead_coeff p = 1" and "degree p > 0" shows "algebraic_int x" proof - from assms obtain q where q: "poly q y = 0" "\i. coeff q i \ \" "lead_coeff q = 1" by (auto simp: algebraic_int.simps) show ?thesis proof from assms q show "lead_coeff (pcompose q p) = 1" by (subst lead_coeff_comp) auto from assms q show "\i. coeff (pcompose q p) i \ \" by (intro allI coeff_pcompose_semiring_closed) auto show "poly (pcompose q p) x = 0" using assms q by (simp add: poly_pcompose) qed qed lemma algebraic_int_abs_real [simp]: "algebraic_int \x :: real\ \ algebraic_int x" by (auto simp: abs_if) lemma algebraic_int_nth_root_real [intro]: assumes "algebraic_int x" shows "algebraic_int (root n x)" proof (cases "n = 0") case False show ?thesis proof (rule algebraic_int_root) show "poly (monom 1 n) (root n x) = (if even n then \x\ else x)" using sgn_power_root[of n x] False by (auto simp add: poly_monom sgn_if split: if_splits) qed (use False assms in \auto simp: degree_monom_eq\) qed auto lemma algebraic_int_sqrt [intro]: "algebraic_int x \ algebraic_int (sqrt x)" by (auto simp: sqrt_def) lemma algebraic_int_csqrt [intro]: "algebraic_int x \ algebraic_int (csqrt x)" by (rule algebraic_int_root[where p = "monom 1 2"]) (auto simp: poly_monom degree_monom_eq) lemma poly_map_poly_cnj [simp]: "poly (map_poly cnj p) x = cnj (poly p (cnj x))" by (induction p) (auto simp: map_poly_pCons) lemma algebraic_int_cnj [intro]: assumes "algebraic_int x" shows "algebraic_int (cnj x)" proof - from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) show ?thesis proof show "poly (map_poly cnj p) (cnj x) = 0" using p by simp show "lead_coeff (map_poly cnj p) = 1" using p by (simp add: coeff_map_poly degree_map_poly) show "\i. coeff (map_poly cnj p) i \ \" using p by (auto simp: coeff_map_poly) qed qed lemma algebraic_int_cnj_iff [simp]: "algebraic_int (cnj x) \ algebraic_int x" using algebraic_int_cnj[of x] algebraic_int_cnj[of "cnj x"] by auto lemma algebraic_int_of_real [intro]: assumes "algebraic_int x" shows "algebraic_int (of_real x)" proof - from assms obtain p where p: "poly p x = 0" "\i. coeff p i \ \" "lead_coeff p = 1" by (auto simp: algebraic_int.simps) show "algebraic_int (of_real x :: 'a)" proof have "poly (map_poly of_real p) (of_real x) = (of_real (poly p x) :: 'a)" by (induction p) (auto simp: map_poly_pCons) thus "poly (map_poly of_real p) (of_real x) = (0 :: 'a)" using p by simp qed (use p in \auto simp: coeff_map_poly degree_map_poly\) qed lemma algebraic_int_of_real_iff [simp]: "algebraic_int (of_real x :: 'a :: {field_char_0, real_algebra_1}) \ algebraic_int x" proof assume "algebraic_int (of_real x :: 'a)" then obtain p where p: "poly (map_poly of_int p) (of_real x :: 'a) = 0" "lead_coeff p = 1" by (auto simp: algebraic_int_altdef_ipoly) show "algebraic_int x" unfolding algebraic_int_altdef_ipoly proof (intro exI[of _ p] conjI) have "of_real (poly (map_poly real_of_int p) x) = poly (map_poly of_int p) (of_real x :: 'a)" by (induction p) (auto simp: map_poly_pCons) also note p(1) finally show "poly (map_poly real_of_int p) x = 0" by simp qed (use p in auto) qed auto subsection \Division of polynomials\ subsubsection \Division in general\ instantiation poly :: (idom_divide) idom_divide begin fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ '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 \ a * lc = cr then \ \\False \\ 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 \ 'a poly \ '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 \ 0" "lc = coeff d (degree d)" and "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ 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 \ 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 "\ = 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 \ 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 "\ = 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 \ dr" by auto have deg_bd: "degree (b * d) \ dr" unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?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 \ 0" shows "(f * g) div g = (f :: 'a poly)" proof - have len: "length (coeffs f) = Suc (degree f)" if "f \ 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 \ 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 (\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 "\ 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 \ 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 \ 0 \ 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 \ 0" shows "is_unit p \ degree p = 0" (is "?lhs \ ?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 \ 0" "p * q = 1" .. then have "degree (p * q) = degree 1" by simp with \p \ 0\ \q \ 0\ 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) \ p = 0 \ a \ 0" for p :: "'a::field poly" by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) lemma is_unit_monom_trivial: "is_unit p \ 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 \ c dvd 1" for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (auto simp: one_pCons) 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 \ 0" and "q \ 0" by auto from \1 = p * q\ have "degree 1 = degree (p * q)" by simp also from \p \ 0\ and \q \ 0\ have "\ = 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 \p dvd 1\ 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 \ 0" proof - obtain a q where "p = pCons a q" by (cases p) with assms have "p = [:a:]" and "a \ 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 \ (\c. p = [:c:] \ 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 \Pseudo-Division\ text \This part is by René Thiemann and Akihisa Yamada.\ fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly \ '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 \ 'a poly \ 'a poly \ 'a poly" where "pseudo_divmod p q \ 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 \ 0" "lc = coeff d (degree d)" and "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" shows "(r' = 0 \ degree r' < degree d) \ 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 "\ = 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 \ dr" using Suc(2) degree_smult_le dual_order.trans by blast have deg_bd: "degree (b * d) \ dr" unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" using degree_diff_le[OF deg_rr deg_bd] by auto with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?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 "\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 \ 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 \ 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 \ 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 \ degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) note main' = main[OF this] then show "r = 0 \ 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 \ 'a poly \ '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 \ pseudo_mod f g" assumes g: "g \ 0" shows "\a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ 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 \ degree r < degree g" by (auto simp: a_def) show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id show "\a q. a \ 0 \ smult a f = g * q + r" by auto qed lemma fst_pseudo_divmod_main_as_divide_poly_main: assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" proof (induct n arbitrary: q r dr) case 0 then show ?case by simp next case (Suc n) note lc0 = leading_coeff_neq_0[OF d, folded lc] then have "pseudo_divmod_main lc q r d dr (Suc n) = pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" by (simp add: Let_def ac_simps) also have "fst \ = divide_poly_main lc (smult (lc^n) (smult lc q + monom (coeff r dr) n)) (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) d (dr - 1) n" by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) also have "\ = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" unfolding smult_monom smult_distribs mult_smult_left[symmetric] using lc0 by (simp add: Let_def ac_simps) finally show ?case . qed subsubsection \Division in polynomials over fields\ lemma pseudo_divmod_field: fixes g :: "'a::field poly" assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" defines "c \ 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 \ 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) \ = 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 \ 0" defines lc: "lc \ 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' \ 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 \ 'a poly" where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" definition normalize_poly :: "'a poly \ '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 \ 0" by simp then have *: "unit_factor (lead_coeff p) \ 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 (\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 \ 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) next fix a b :: "'a poly" assume "is_unit a" thus "unit_factor (a * b) = a * unit_factor b" by (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult elim!: is_unit_polyE) qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) end instance poly :: ("{semidom_divide_unit_factor,idom_divide,normalization_semidom_multiplicative}") normalization_semidom_multiplicative by intro_classes (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult) lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\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 \ 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: fixes c :: "'a :: {normalization_semidom_multiplicative, idom_divide}" shows "normalize (smult c p) = smult (normalize c) (normalize p)" proof - have "smult c p = [:c:] * p" by simp also have "normalize \ = smult (normalize c) (normalize p)" by (subst normalize_mult) (simp add: normalize_const_poly) finally show ?thesis . qed instantiation poly :: (field) idom_modulo begin definition modulo_poly :: "'a poly \ 'a poly \ '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 pseudo_divmod_eq_div_mod: \pseudo_divmod f g = (f div g, f mod g)\ if \lead_coeff g = 1\ using that by (auto simp add: divide_poly_field mod_poly_def pseudo_mod_def) lemma degree_mod_less_degree: \degree (x mod y) < degree y\ if \y \ 0\ \\ y dvd x\ proof - from pseudo_mod(2) [of y] \y \ 0\ have *: \pseudo_mod f y \ 0 \ degree (pseudo_mod f y) < degree y\ for f by blast from \\ y dvd x\ have \x mod y \ 0\ by blast with \y \ 0\ show ?thesis by (auto simp add: mod_poly_def intro: *) qed instantiation poly :: (field) unique_euclidean_ring begin definition euclidean_size_poly :: "'a poly \ nat" where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" definition division_segment_poly :: "'a poly \ 'a poly" where [simp]: "division_segment_poly p = 1" instance proof show \(q * p + r) div p = q\ if \p \ 0\ and \euclidean_size r < euclidean_size p\ for q p r :: \'a poly\ proof (cases \r = 0\) case True with that show ?thesis by simp next case False with \p \ 0\ \euclidean_size r < euclidean_size p\ have \degree r < degree p\ by (simp add: euclidean_size_poly_def) with \r \ 0\ have \\ p dvd r\ by (auto dest: dvd_imp_degree) have \(q * p + r) div p = q \ (q * p + r) mod p = r\ proof (rule ccontr) assume \\ ?thesis\ moreover have *: \((q * p + r) div p - q) * p = r - (q * p + r) mod p\ by (simp add: algebra_simps) ultimately have \(q * p + r) div p \ q\ and \(q * p + r) mod p \ r\ using \p \ 0\ by auto from \\ p dvd r\ have \\ p dvd (q * p + r)\ by simp with \p \ 0\ have \degree ((q * p + r) mod p) < degree p\ by (rule degree_mod_less_degree) with \degree r < degree p\ \(q * p + r) mod p \ r\ have \degree (r - (q * p + r) mod p) < degree p\ by (auto intro: degree_diff_less) also have \degree p \ degree ((q * p + r) div p - q) + degree p\ by simp also from \(q * p + r) div p \ q\ \p \ 0\ have \\ = degree (((q * p + r) div p - q) * p)\ by (simp add: degree_mult_eq) also from * have \\ = degree (r - (q * p + r) mod p)\ by simp finally have \degree (r - (q * p + r) mod p) < degree (r - (q * p + r) mod p)\ . then show False by simp qed then show \(q * p + r) div p = q\ .. qed qed (auto simp: euclidean_size_poly_def degree_mult_eq power_add intro: degree_mod_less_degree) end lemma euclidean_relation_polyI [case_names by0 divides euclidean_relation]: \(x div y, x mod y) = (q, r)\ if by0: \y = 0 \ q = 0 \ r = x\ and divides: \y \ 0 \ y dvd x \ r = 0 \ x = q * y\ and euclidean_relation: \y \ 0 \ \ y dvd x \ degree r < degree y \ x = q * y + r\ by (rule euclidean_relationI) (use that in \simp_all add: euclidean_size_poly_def\) lemma div_poly_eq_0_iff: \x div y = 0 \ x = 0 \ y = 0 \ degree x < degree y\ for x y :: \'a::field poly\ by (simp add: unique_euclidean_semiring_class.div_eq_0_iff euclidean_size_poly_def) lemma div_poly_less: \x div y = 0\ if \degree x < degree y\ for x y :: \'a::field poly\ using that by (simp add: div_poly_eq_0_iff) lemma mod_poly_less: \x mod y = x\ if \degree x < degree y\ using that by (simp add: mod_eq_self_iff_div_eq_0 div_poly_eq_0_iff) lemma degree_div_less: \degree (x div y) < degree x\ if \degree x > 0\ \degree y > 0\ for x y :: \'a::field poly\ proof (cases \x div y = 0\) case True with \degree x > 0\ show ?thesis by simp next case False from that have \x \ 0\ \y \ 0\ and *: \degree (x div y * y + x mod y) > 0\ by auto show ?thesis proof (cases \y dvd x\) case True then obtain z where \x = y * z\ .. then have \degree (x div y) < degree (x div y * y)\ using \y \ 0\ \x \ 0\ \degree y > 0\ by (simp add: degree_mult_eq) with \y dvd x\ show ?thesis by simp next case False with \y \ 0\ have \degree (x mod y) < degree y\ by (rule degree_mod_less_degree) with \y \ 0\ \x div y \ 0\ have \degree (x mod y) < degree (x div y * y)\ by (simp add: degree_mult_eq) then have \degree (x div y * y + x mod y) = degree (x div y * y)\ by (rule degree_add_eq_left) with \y \ 0\ \x div y \ 0\ \degree y > 0\ show ?thesis by (simp add: degree_mult_eq) qed qed lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" by (rule degree_mod_less_degree) auto lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using degree_mod_less' by blast lemma div_smult_left: \smult a x div y = smult a (x div y)\ (is ?Q) and mod_smult_left: \smult a x mod y = smult a (x mod y)\ (is ?R) for x y :: \'a::field poly\ proof - have \(smult a x div y, smult a x mod y) = (smult a (x div y), smult a (x mod y))\ proof (cases \a = 0\) case True then show ?thesis by simp next case False - then show ?thesis - by (cases y \smult a (x div y)\ \smult a (x mod y)\ \smult a x\ rule: euclidean_relation_polyI) - (simp_all add: dvd_smult_iff degree_mod_less_degree flip: smult_add_right) + show ?thesis + by (rule euclidean_relation_polyI) + (use False in \simp_all add: dvd_smult_iff degree_mod_less_degree flip: smult_add_right\) qed then show ?Q and ?R by simp_all qed 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 poly_div_add_left: \(x + y) div z = x div z + y div z\ (is ?Q) and poly_mod_add_left: \(x + y) mod z = x mod z + y mod z\ (is ?R) for x y z :: \'a::field poly\ proof - have \((x + y) div z, (x + y) mod z) = (x div z + y div z, x mod z + y mod z)\ - proof (cases z \x div z + y div z\ \x mod z + y mod z\ \x + y\ rule: euclidean_relation_polyI) + proof (induction rule: euclidean_relation_polyI) case by0 then show ?case by simp next case divides then obtain w where \x + y = z * w\ by blast then have y: \y = z * w - x\ by (simp add: algebra_simps) from \z \ 0\ show ?case using mod_mult_self4 [of z w \- x\] div_mult_self4 [of z w \- x\] by (simp add: algebra_simps y) next case euclidean_relation then have \degree (x mod z + y mod z) < degree z\ using degree_mod_less_degree [of z x] degree_mod_less_degree [of z y] dvd_add_right_iff [of z x y] dvd_add_left_iff [of z y x] by (cases \z dvd x \ z dvd y\) (auto intro: degree_add_less) moreover have \x + y = (x div z + y div z) * z + (x mod z + y mod z)\ by (simp add: algebra_simps) ultimately show ?case by simp qed then show ?Q and ?R by simp_all qed 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 div_smult_right: \x div smult a y = smult (inverse a) (x div y)\ (is ?Q) and mod_smult_right: \x mod smult a y = (if a = 0 then x else x mod y)\ (is ?R) proof - have \(x div smult a y, x mod smult a y) = (smult (inverse a) (x div y), (if a = 0 then x else x mod y))\ - proof (cases \smult a y\ \smult (inverse a) (x div y)\ \(if a = 0 then x else x mod y)\ x rule: euclidean_relation_polyI) + proof (induction rule: euclidean_relation_polyI) case by0 then show ?case by auto next case divides moreover define w where \w = x div y\ ultimately have \x = y * w\ by (simp add: smult_dvd_iff) with divides show ?case by simp next case euclidean_relation then show ?case by (simp add: smult_dvd_iff degree_mod_less_degree) qed then show ?Q and ?R by simp_all qed 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 poly_div_mult_right: \x div (y * z) = (x div y) div z\ (is ?Q) and poly_mod_mult_right: \x mod (y * z) = y * (x div y mod z) + x mod y\ (is ?R) for x y z :: \'a::field poly\ proof - have \(x div (y * z), x mod (y * z)) = ((x div y) div z, y * (x div y mod z) + x mod y)\ - proof (cases \y * z\ \(x div y) div z\ \y * (x div y mod z) + x mod y\ x rule: euclidean_relation_polyI) + proof (induction rule: euclidean_relation_polyI) case by0 then show ?case by auto next case divides then show ?case by auto next case euclidean_relation then have \y \ 0\ \z \ 0\ by simp_all with \\ y * z dvd x\ have \degree (y * (x div y mod z) + x mod y) < degree (y * z)\ using degree_mod_less_degree [of y x] degree_mod_less_degree [of z \x div y\] degree_add_eq_left [of \x mod y\ \y * (x div y mod z)\] by (cases \z dvd x div y\; cases \y dvd x\) (auto simp add: degree_mult_eq not_dvd_imp_mod_neq_0 dvd_div_iff_mult) moreover have \x = x div y div z * (y * z) + (y * (x div y mod z) + x mod y)\ by (simp add: field_simps flip: distrib_left) ultimately show ?case by simp qed then show ?Q and ?R by simp_all qed lemma dvd_pCons_imp_dvd_pCons_mod: \y dvd pCons a (x mod y)\ if \y dvd pCons a x\ proof - have \pCons a x = pCons a (x div y * y + x mod y)\ by simp also have \\ = pCons 0 (x div y * y) + pCons a (x mod y)\ by simp also have \pCons 0 (x div y * y) = (x div y * monom 1 (Suc 0)) * y\ by (simp add: monom_Suc) finally show \y dvd pCons a (x mod y)\ using \y dvd pCons a x\ by simp qed lemma degree_less_if_less_eqI: \degree x < degree y\ if \degree x \ degree y\ \coeff x (degree y) = 0\ \x \ 0\ proof (cases \degree x = degree y\) case True with \coeff x (degree y) = 0\ have \lead_coeff x = 0\ by simp then have \x = 0\ by simp with \x \ 0\ show ?thesis by simp next case False with \degree x \ degree y\ show ?thesis by simp 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))\ (is ?Q) and 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)\ (is ?R) for x y :: \'a::field poly\ proof - have \?Q\ and \?R\ if \q = 0\ using that by simp_all moreover have \?Q\ and \?R\ if \q \ 0\ proof - define b where \b = coeff (pCons a (p mod q)) (degree q) / lead_coeff q\ have \(pCons a p div q, pCons a p mod q) = (pCons b (p div q), (pCons a (p mod q) - smult b q))\ (is \_ = (?q, ?r)\) - proof (cases q ?q ?r \pCons a p\ rule: euclidean_relation_polyI) + proof (induction rule: euclidean_relation_polyI) case by0 with \q \ 0\ show ?case by simp next case divides show ?case proof (cases \pCons a (p mod q) = 0\) case True then show ?thesis by (auto simp add: b_def) next case False have \q dvd pCons a (p mod q)\ using \q dvd pCons a p\ by (rule dvd_pCons_imp_dvd_pCons_mod) then obtain s where *: \pCons a (p mod q) = q * s\ .. with False have \s \ 0\ by auto from \q \ 0\ have \degree (pCons a (p mod q)) \ degree q\ by (auto simp add: Suc_le_eq intro: degree_mod_less_degree) moreover from \s \ 0\ have \degree q \ degree (pCons a (p mod q))\ by (simp add: degree_mult_right_le *) ultimately have \degree (pCons a (p mod q)) = degree q\ by (rule order.antisym) with \s \ 0\ \q \ 0\ have \degree s = 0\ by (simp add: * degree_mult_eq) then obtain c where \s = [:c:]\ by (rule degree_eq_zeroE) also have \c = b\ using \q \ 0\ by (simp add: b_def * \s = [:c:]\) finally have \smult b q = pCons a (p mod q)\ by (simp add: *) then show ?thesis by simp qed next case euclidean_relation then have \degree q > 0\ using is_unit_iff_degree by blast from \q \ 0\ have \degree (pCons a (p mod q)) \ degree q\ by (auto simp add: Suc_le_eq intro: degree_mod_less_degree) moreover have \degree (smult b q) \ degree q\ by (rule degree_smult_le) ultimately have \degree (pCons a (p mod q) - smult b q) \ degree q\ by (rule degree_diff_le) moreover have \coeff (pCons a (p mod q) - smult b q) (degree q) = 0\ using \degree q > 0\ by (auto simp add: b_def) ultimately have \degree (pCons a (p mod q) - smult b q) < degree q\ using \degree q > 0\ by (cases \pCons a (p mod q) = smult b q\) (auto intro: degree_less_if_less_eqI) then show ?case by simp qed with \q \ 0\ show ?Q and ?R by (simp_all add: b_def) qed ultimately show ?Q and ?R by simp_all qed lemma div_mod_fold_coeffs: "(p div q, p mod q) = (if q = 0 then (0, p) else fold_coeffs (\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 mod_pCons: fixes a :: "'a::field" and x y :: "'a::field poly" assumes y: "y \ 0" defines "b \ 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 (simp add: mod_pCons_eq) subsubsection \List-based versions for fast implementation\ (* Subsection by: Sebastiaan Joosten René Thiemann Akihisa Yamada *) fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ '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 \ 'a list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "pseudo_divmod_main_list lc q r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; qqq = cCons a (map ((*) lc) q); rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) 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 \ 'a list \ 'a list \ nat \ 'a list" where "pseudo_mod_main_list lc r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) 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 \ 'a list \ 'a list \ nat \ 'a list \ '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 ((*) 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 \ 'a list \ nat \ '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 ((*) 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 \ 'a list \ 'a list \ '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 \ 'a list \ '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 ((*) 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 ((*) a) y)) = minus_poly_rev_list x (map ((*) 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 \ length q \ 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) \ length (rev p)" by simp from this[folded 1(2,3)] have ys_xs: "length ys \ 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 "\ = 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 "\ = 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 "\ = 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 \ length r \ d \ [] \ hd (minus_poly_rev_list (map ((*) (last d)) r) (map ((*) (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 ((*) 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 \ [] \ 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 \ 0" and lc: "last d = lc" and "d \ []" 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 *: "\ Suc (length r) \ length d" by simp with \d \ []\ have "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce let ?a = "(hd (rev r))" let ?rr = "map ((*) lc) (rev r)" let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map ((*) ?a) (rev d))))" let ?qq = "cCons ?a (map ((*) 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 \r \ []\ * 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) \r \ []\ have n_ok : "n = 1 + (length ?rrr) - length d" by simp have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ 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 \r \ []\] 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 ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))) = 0" by (fold lc, subst head_minus_poly_rev_list, insert * \d \ []\, auto) from * have "length d \ 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 ((*) lc) (rev r)) (map ((*) (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) \ 0" and "last (coeffs g) = lead_coeff g" and "coeffs g \ []" 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)" by (metis hd_rev qr rev.simps(1) rev_swap) ultimately show ?thesis by (simp add: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_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: "\f g p. snd ((\(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 \Improved Code-Equations for Polynomial (Pseudo) Division\ lemma pdivmod_via_pseudo_divmod: \(f div g, f mod g) = (if g = 0 then (0, f) else let ilc = inverse (lead_coeff 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 ilc where \ilc = inverse (lead_coeff g)\ define h where \h = smult ilc g\ from False have \lead_coeff h = 1\ and \ilc \ 0\ by (auto simp: h_def ilc_def) define q r where \q = f div h\ and \r = f mod h\ with \lead_coeff h = 1\ have p: \pseudo_divmod f h = (q, r)\ by (simp add: pseudo_divmod_eq_div_mod) from \ilc \ 0\ have \(f div g, f mod g) = (smult ilc q, r)\ by (auto simp: h_def div_smult_right mod_smult_right q_def r_def) also have \(smult ilc q, r) = ?r\ using \g \ 0\ by (auto simp: Let_def p simp flip: h_def ilc_def) finally show ?thesis . 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 ((*) ilc) cg; (q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) 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 \ 0" by (auto simp: ilc_def) with False have id: "g = 0 \ False" "coeffs g = [] \ False" "last (coeffs g) = coeff g (degree g)" "coeffs (smult ilc g) = [] \ False" by (auto simp: last_coeffs_eq_coeff_degree) have id2: "hd (rev (coeffs (smult ilc g))) = 1" by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" "rev (coeffs (smult ilc g)) = rev (map ((*) 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 ((*) 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 ((*) 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 \ 'a list \ 'a list \ 'a list \ nat \ '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 ((*) 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 ((*) 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 \ 'a poly \ '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 ((*) 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 \ 'a poly \ '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 ((*) ilc) cg; q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) in poly_of_list ((map ((*) ilc) q)))" text \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 \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ lemma div_field_poly_impl[code_unfold]: "(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 \ 0" and lc: "last d = lc" and d: "d \ []" 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: "\ Suc (length r) \ length d" by simp with d have r: "r \ []" 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 \ 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 \ Poly (divide_poly_main_list lc (cCons (lcr div lc) q) (rev (rev (minus_poly_rev_list (rev rr) (rev (map ((*) (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: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ 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 \ 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 \lcg \ 0\) (auto simp: lcg ltp, simp add: degree_eq_length_coeffs) qed qed subsection \Primality and irreducibility in polynomial rings\ lemma prod_mset_const_poly: "(\x\#A. [:f x:]) = [:prod_mset (image_mset f A):]" by (induct A) (simp_all add: ac_simps) lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ 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 \ 0" "b \ 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 \ b' dvd 1" by (rule irreducibleD) with ab' show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed (insert A, auto simp: irreducible_def is_unit_poly_iff) next assume A: "irreducible [:c:]" then have "c \ 0" and "\ c dvd 1" by (auto simp add: irreducible_def is_unit_const_poly_iff) then 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 \ [:b:] dvd 1" by (rule irreducibleD) then show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed qed 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. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (metis (mono_tags, lifting) B Greatest_le_nat that) 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) = (\k\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 "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?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 \ {..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 } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ 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:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) subsection \Content and primitive part of a polynomial\ definition content :: "'a::semiring_gcd poly \ '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 \ 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 \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) also have "\ \ (\a\set (coeffs p). c dvd a)" proof safe fix n :: nat assume "\a\set (coeffs p). c dvd a" then show "c dvd coeff p n" by (cases "n \ 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 "\ \ 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 \ 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 \ set (coeffs p) \ 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) \ 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]: fixes c :: "'a :: {normalization_semidom_multiplicative, semiring_gcd}" shows "content (smult c p) = normalize c * content p" by (simp add: content_def coeffs_smult Gcd_fin_mult normalize_mult) lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" by (auto simp: content_def simp: poly_eq_iff coeffs_def) definition primitive_part :: "'a :: semiring_gcd poly \ 'a poly" where "primitive_part p = map_poly (\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 \ p = 0" proof (cases "p = 0") case True then show ?thesis by simp next case False then have "primitive_part p = map_poly (\x. x div content p) p" by (simp add: primitive_part_def) also from False have "\ = 0 \ 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]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" assumes "p \ 0" shows "content (primitive_part p) = 1" proof - have "p = smult (content p) (primitive_part p)" by simp also have "content \ = 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 :: {normalization_semidom_multiplicative, semiring_gcd} poly" where "p = smult (content p) p'" "content p' = 1" proof (cases "p = 0") case True then have "p = smult (content p) 1" "content 1 = 1" by simp_all then show ?thesis .. next case False then have "p = smult (content p) (primitive_part p)" "content (primitive_part p) = 1" by simp_all then show ?thesis .. qed lemma content_dvd_contentI [intro]: "p dvd q \ 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 \ 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 \ = degree (primitive_part p)" by (subst degree_smult_eq) simp_all finally show ?thesis .. qed lemma smult_content_normalize_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd, idom_divide} poly" shows "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 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 \ 0" "g \ 0" by auto hence "f * g \ 0" by auto { assume "\is_unit (content (f * g))" with False have "\p. p dvd content (f * g) \ prime p" by (intro prime_divisor_exists) simp_all then obtain p where "p dvd content (f * g)" "prime p" by blast from \p dvd content (f * g)\ have "[:p:] dvd f * g" by (simp add: const_poly_dvd_iff_dvd_content) moreover from \prime p\ have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly) ultimately have "[:p:] dvd f \ [: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 \prime p\ 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, normalization_semidom_multiplicative} poly" shows "content (p * q) = content p * content q" proof (cases "p * q = 0") case False then have "p \ 0" and "q \ 0" by simp_all then have *: "content (primitive_part p * primitive_part q) = 1" by (auto intro: content_1_mult) have "p * q = smult (content p) (primitive_part p) * smult (content q) (primitive_part q)" by simp also have "\ = smult (content p * content q) (primitive_part p * primitive_part q)" by (metis mult.commute mult_smult_right smult_smult) with * show ?thesis by (simp add: normalize_mult) next case True then show ?thesis by auto qed end lemma primitive_part_mult: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} 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 "\ = (p div [:content p:]) * (q div [:content q:])" by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac) also have "\ = 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, normalization_semidom_multiplicative} 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 \ = 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, normalization_semidom_multiplicative} poly" shows "p dvd q \ 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, normalization_semidom_multiplicative} poly multiset" shows "content (prod_mset A) = prod_mset (image_mset content A)" by (induction A) (simp_all add: content_mult mult_ac) lemma content_prod_eq_1_iff: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = 1 \ content p = 1 \ 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) no_notation cCons (infixr "##" 65) end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,3376 +1,3376 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (use assms in simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule iffI; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") auto lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ lemma power_mod [mod_simps]: "((a mod b) ^ n) mod b = (a ^ n) mod b" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" by (simp add: mod_mult_right_eq) with Suc show ?case by (simp add: mod_mult_left_eq mod_mult_right_eq) qed lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp qed end class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel begin subclass idom_divide .. lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" using div_mult_mult1 [of "- 1" a b] by simp lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" using mod_mult_mult1 [of "- 1" a b] by simp lemma div_minus_right: "a div (- b) = (- a) div b" using div_minus_minus [of "- a" b] by simp lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" using mod_minus_minus [of "- a" b] by simp lemma div_minus1_right [simp]: "a div (- 1) = - a" using div_minus_right [of a 1] by simp lemma mod_minus1_right [simp]: "a mod (- 1) = 0" using mod_minus_right [of a 1] by simp text \Negation respects modular equivalence.\ lemma mod_minus_eq [mod_simps]: "(- (a mod b)) mod b = (- a) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: div_mult_mod_eq) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. then show ?thesis by (simp add: mod_minus_eq) qed text \Subtraction respects modular equivalence.\ lemma mod_diff_left_eq [mod_simps]: "(a mod c - b) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp lemma mod_diff_right_eq [mod_simps]: "(a - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_eq: "(a mod c - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp lemma minus_mod_self2 [simp]: "(a - b) mod b = a mod b" using mod_diff_right_eq [of a b b] by (simp add: mod_diff_right_eq) lemma minus_mod_self1 [simp]: "(b - a) mod b = - a mod b" using mod_add_self2 [of "- a" b] by simp lemma mod_eq_dvd_iff: "a mod c = b mod c \ c dvd a - b" (is "?P \ ?Q") proof assume ?P then have "(a mod c - b mod c) mod c = 0" by simp then show ?Q by (simp add: dvd_eq_mod_eq_0 mod_simps) next assume ?Q then obtain d where d: "a - b = c * d" .. then have "a = c * d + b" by (simp add: algebra_simps) then show ?P by simp qed lemma mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: \euclidean_size (a * b) = euclidean_size a * euclidean_size b\ fixes division_segment :: \'a \ 'a\ assumes is_unit_division_segment [simp]: \is_unit (division_segment a)\ and division_segment_mult: \a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b\ and division_segment_mod: \b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b\ assumes div_bounded: \b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q\ begin lemma division_segment_not_0 [simp]: \division_segment a \ 0\ using is_unit_division_segment [of a] is_unitE [of \division_segment a\] by blast lemma euclidean_relationI [case_names by0 divides euclidean_relation]: \(a div b, a mod b) = (q, r)\ if by0: \b = 0 \ q = 0 \ r = a\ and divides: \b \ 0 \ b dvd a \ r = 0 \ a = q * b\ and euclidean_relation: \b \ 0 \ \ b dvd a \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ a = q * b + r\ proof (cases \b = 0\) case True with by0 show ?thesis by simp next case False show ?thesis proof (cases \b dvd a\) case True with \b \ 0\ divides show ?thesis by simp next case False with \b \ 0\ euclidean_relation have \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ \a = q * b + r\ by simp_all from \b \ 0\ \division_segment r = division_segment b\ \euclidean_size r < euclidean_size b\ have \(q * b + r) div b = q\ by (rule div_bounded) with \a = q * b + r\ have \q = a div b\ by simp from \a = q * b + r\ have \a div b * b + a mod b = q * b + r\ by (simp add: div_mult_mod_eq) with \q = a div b\ have \q * b + a mod b = q * b + r\ by simp then have \r = a mod b\ by simp with \q = a div b\ show ?thesis by simp qed qed subclass euclidean_semiring_cancel proof fix a b c assume \b \ 0\ have \((a + c * b) div b, (a + c * b) mod b) = (c + a div b, a mod b)\ - proof (cases b \c + a div b\ \a mod b\ \a + c * b\ rule: euclidean_relationI) + proof (induction rule: euclidean_relationI) case by0 with \b \ 0\ show ?case by simp next case divides then show ?case by (simp add: algebra_simps dvd_add_left_iff) next case euclidean_relation then have \\ b dvd a\ by (simp add: dvd_add_left_iff) have \a mod b + (b * c + b * (a div b)) = b * c + ((a div b) * b + a mod b)\ by (simp add: ac_simps) with \b \ 0\ have *: \a mod b + (b * c + b * (a div b)) = b * c + a\ by (simp add: div_mult_mod_eq) from \\ b dvd a\ euclidean_relation show ?case by (simp_all add: algebra_simps division_segment_mod mod_size_less *) qed then show \(a + c * b) div b = c + a div b\ by simp next fix a b c assume \c \ 0\ have \((c * a) div (c * b), (c * a) mod (c * b)) = (a div b, c * (a mod b))\ - proof (cases \c * b\ \a div b\ \c * (a mod b)\ \c * a\ rule: euclidean_relationI) + proof (induction rule: euclidean_relationI) case by0 with \c \ 0\ show ?case by simp next case divides then show ?case by (auto simp add: algebra_simps) next case euclidean_relation then have \b \ 0\ \a mod b \ 0\ by (simp_all add: mod_eq_0_iff_dvd) have \c * (a mod b) + b * (c * (a div b)) = c * ((a div b) * b + a mod b)\ by (simp add: algebra_simps) with \b \ 0\ have *: \c * (a mod b) + b * (c * (a div b)) = c * a\ by (simp add: div_mult_mod_eq) from \b \ 0\ \c \ 0\ have \euclidean_size c * euclidean_size (a mod b) < euclidean_size c * euclidean_size b\ using mod_size_less [of b a] by simp with euclidean_relation \b \ 0\ \a mod b \ 0\ show ?case by (simp add: algebra_simps division_segment_mult division_segment_mod euclidean_size_mult *) qed then show \(c * a) div (c * b) = a div b\ by simp qed lemma div_eq_0_iff: \a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0\ (is "_ \ ?P") if \division_segment a = division_segment b\ proof (cases \a = 0 \ b = 0\) case True then show ?thesis by auto next case False then have \a \ 0\ \b \ 0\ by simp_all have \a div b = 0 \ euclidean_size a < euclidean_size b\ proof assume \a div b = 0\ then have \a mod b = a\ using div_mult_mod_eq [of a b] by simp with \b \ 0\ mod_size_less [of b a] show \euclidean_size a < euclidean_size b\ by simp next assume \euclidean_size a < euclidean_size b\ have \(a div b, a mod b) = (0, a)\ - proof (cases b 0 a a rule: euclidean_relationI) + proof (induction rule: euclidean_relationI) case by0 show ?case by simp next case divides with \euclidean_size a < euclidean_size b\ show ?case using dvd_imp_size_le [of b a] \a \ 0\ by simp next case euclidean_relation with \euclidean_size a < euclidean_size b\ that show ?case by simp qed then show \a div b = 0\ by simp qed with \b \ 0\ show ?thesis by simp qed lemma div_mult1_eq: \(a * b) div c = a * (b div c) + a * (b mod c) div c\ proof - have *: \(a * b) mod c + (a * (c * (b div c)) + c * (a * (b mod c) div c)) = a * b\ (is \?A + (?B + ?C) = _\) proof - have \?A = a * (b mod c) mod c\ by (simp add: mod_mult_right_eq) then have \?C + ?A = a * (b mod c)\ by (simp add: mult_div_mod_eq) then have \?B + (?C + ?A) = a * (c * (b div c) + (b mod c))\ by (simp add: algebra_simps) also have \\ = a * b\ by (simp add: mult_div_mod_eq) finally show ?thesis by (simp add: algebra_simps) qed have \((a * b) div c, (a * b) mod c) = (a * (b div c) + a * (b mod c) div c, (a * b) mod c)\ - proof (cases c \a * (b div c) + a * (b mod c) div c\ \(a * b) mod c\ \a * b\ rule: euclidean_relationI) + proof (induction rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed lemma div_add1_eq: \(a + b) div c = a div c + b div c + (a mod c + b mod c) div c\ proof - have *: \(a + b) mod c + (c * (a div c) + (c * (b div c) + c * ((a mod c + b mod c) div c))) = a + b\ (is \?A + (?B + (?C + ?D)) = _\) proof - have \?A + (?B + (?C + ?D)) = ?A + ?D + (?B + ?C)\ by (simp add: ac_simps) also have \?A + ?D = (a mod c + b mod c) mod c + ?D\ by (simp add: mod_add_eq) also have \\ = a mod c + b mod c\ by (simp add: mod_mult_div_eq) finally have \?A + (?B + (?C + ?D)) = (a mod c + ?B) + (b mod c + ?C)\ by (simp add: ac_simps) then show ?thesis by (simp add: mod_mult_div_eq) qed have \((a + b) div c, (a + b) mod c) = (a div c + b div c + (a mod c + b mod c) div c, (a + b) mod c)\ - proof (cases c \a div c + b div c + (a mod c + b mod c) div c\ \(a + b) mod c\ \a + b\ rule: euclidean_relationI) + proof (induction rule: euclidean_relationI) case by0 then show ?case by simp next case divides with * show ?case by (simp add: algebra_simps) next case euclidean_relation with * show ?case by (simp add: division_segment_mod mod_size_less algebra_simps) qed then show ?thesis by simp qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: \nat \ nat\ where [simp]: \normalize = (id :: nat \ nat)\ definition unit_factor_nat :: \nat \ nat\ where \unit_factor n = of_bool (n > 0)\ for n :: nat lemma unit_factor_simps [simp]: \unit_factor 0 = (0::nat)\ \unit_factor (Suc n) = 1\ by (simp_all add: unit_factor_nat_def) definition divide_nat :: \nat \ nat \ nat\ where \m div n = (if n = 0 then 0 else Max {k. k * n \ m})\ for m n :: nat instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: \nat \ nat\ where [simp]: \euclidean_size_nat = id\ definition division_segment_nat :: \nat \ nat\ where [simp]: \division_segment n = 1\ for n :: nat definition modulo_nat :: \nat \ nat \ nat\ where \m mod n = m - (m div n * n)\ for m n :: nat instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end lemma euclidean_relation_natI [case_names by0 divides euclidean_relation]: \(m div n, m mod n) = (q, r)\ if by0: \n = 0 \ q = 0 \ r = m\ and divides: \n > 0 \ n dvd m \ r = 0 \ m = q * n\ and euclidean_relation: \n > 0 \ \ n dvd m \ r < n \ m = q * n + r\ for m n q r :: nat by (rule euclidean_relationI) (use that in simp_all) lemma div_nat_eqI: \m div n = q\ if \n * q \ m\ and \m < n * Suc q\ for m n q :: nat proof - have \(m div n, m mod n) = (q, m - n * q)\ - proof (cases n q \m - n * q\ m rule: euclidean_relation_natI) + proof (induction rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from \n dvd m\ obtain s where \m = n * s\ .. with \n > 0\ that have \s < Suc q\ by (simp only: mult_less_cancel1) with \m = n * s\ \n > 0\ that have \q = s\ by simp with \m = n * s\ show ?case by (simp add: ac_simps) next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed lemma mod_nat_eqI: \m mod n = r\ if \r < n\ and \r \ m\ and \n dvd m - r\ for m n r :: nat proof - have \(m div n, m mod n) = ((m - r) div n, r)\ - proof (cases n \(m - r) div n\ r m rule: euclidean_relation_natI) + proof (induction rule: euclidean_relation_natI) case by0 with that show ?case by simp next case divides from that dvd_minus_add [of r \m\ 1 n] have \n dvd m + (n - r)\ by simp with divides have \n dvd n - r\ by (simp add: dvd_add_right_iff) then have \n \ n - r\ by (rule dvd_imp_le) (use \r < n\ in simp) with \n > 0\ have \r = 0\ by simp with \n > 0\ that show ?case by simp next case euclidean_relation with that show ?case by (simp add: ac_simps) qed then show ?thesis by simp qed text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_nat_eqI mod_nat_eqI) lemma split_div: \P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))\ (is ?div) and split_mod: \Q (m mod n) \ (n = 0 \ Q m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ Q j))\ (is ?mod) for m n :: nat proof - have *: \R (m div n) (m mod n) \ (n = 0 \ R 0 m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ R i j))\ for R by (cases \n = 0\) auto from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed declare split_div [of _ _ \numeral n\, linarith_split] for n declare split_mod [of _ _ \numeral n\, linarith_split] for n lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all lemma div_mult2_eq: \m div (n * q) = (m div n) div q\ (is ?Q) and mod_mult2_eq: \m mod (n * q) = n * (m div n mod q) + m mod n\ (is ?R) for m n q :: nat proof - have \(m div (n * q), m mod (n * q)) = ((m div n) div q, n * (m div n mod q) + m mod n)\ - proof (cases \n * q\ \(m div n) div q\ \n * (m div n mod q) + m mod n\ m rule: euclidean_relation_natI) + proof (induction rule: euclidean_relation_natI) case by0 then show ?case by auto next case divides from \n * q dvd m\ obtain t where \m = n * q * t\ .. with \n * q > 0\ show ?case by (simp add: algebra_simps) next case euclidean_relation then have \n > 0\ \q > 0\ by simp_all from \n > 0\ have \m mod n < n\ by (rule mod_less_divisor) from \q > 0\ have \m div n mod q < q\ by (rule mod_less_divisor) then obtain s where \q = Suc (m div n mod q + s)\ by (blast dest: less_imp_Suc_add) moreover have \m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)\ using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately have \m mod n + n * (m div n mod q) < n * q\ by simp then show ?case by (simp add: algebra_simps flip: add_mult_distrib2) qed then show ?Q and ?R by simp_all qed lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed lemma div_less_iff_less_mult: \m div q < n \ m < n * q\ (is \?P \ ?Q\) if \q > 0\ for m n q :: nat proof assume ?Q then show ?P by (rule less_mult_imp_div_less) next assume ?P then obtain h where \n = Suc (m div q + h)\ using less_natE by blast moreover have \m < m + (Suc h * q - m mod q)\ using that by (simp add: trans_less_add1) ultimately show ?Q by (simp add: algebra_simps flip: minus_mod_eq_mult_div) qed lemma less_eq_div_iff_mult_less_eq: \m \ n div q \ m * q \ n\ if \q > 0\ for m n q :: nat using div_less_iff_less_mult [of q n m] that by auto lemma div_Suc: \Suc m div n = (if Suc m mod n = 0 then Suc (m div n) else m div n)\ (is "_ = ?rhs") proof (cases \n = 0 \ n = 1\) case True then show ?thesis by auto next case False then have \n > 1\ by simp then have *: \Suc 0 div n = 0\ by (simp add: div_eq_0_iff) have \(m + 1) div n = ?rhs\ proof (cases \n dvd Suc m\) case True then obtain q where \Suc m = n * q\ .. then have m: \m = n * q - 1\ by simp have \q > 0\ by (rule ccontr) (use \Suc m = n * q\ in simp) from m have \m mod n = (n * q - 1) mod n\ by simp also have \\ = (n * q - 1 + n) mod n\ by simp also have \n * q - 1 + n = n * q + (n - 1)\ using \n > 1\ \q > 0\ by (simp add: algebra_simps) finally have \m mod n = (n - 1) mod n\ by simp with \n > 1\ have \m mod n = n - 1\ by simp with True \n > 1\ show ?thesis by (subst div_add1_eq) auto next case False have \Suc (m mod n) \ n\ proof (rule ccontr) assume \\ Suc (m mod n) \ n\ then have \m mod n = n - 1\ by simp with \n > 1\ have \(m + 1) mod n = 0\ by (subst mod_add_left_eq [symmetric]) simp then have \n dvd Suc m\ by auto with False show False .. qed moreover have \Suc (m mod n) \ n\ using \n > 1\ by (simp add: Suc_le_eq) ultimately have \Suc (m mod n) < n\ by simp with False \n > 1\ show ?thesis by (subst div_add1_eq) (auto simp add: div_eq_0_iff mod_greater_zero_iff_not_dvd) qed then show ?thesis by simp qed lemma mod_Suc: \Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))\ (is "_ = ?rhs") proof (cases "n = 0") case True then show ?thesis by simp next case False have "Suc m mod n = Suc (m mod n) mod n" by (simp add: mod_simps) also have "\ = ?rhs" using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq) finally show ?thesis . qed lemma Suc_times_mod_eq: "Suc (m * n) mod m = 1" if "Suc 0 < m" using that by (simp add: mod_Suc) lemma Suc_times_numeral_mod_eq [simp]: "Suc (numeral k * n) mod numeral k = 1" if "numeral k \ (1::nat)" by (rule Suc_times_mod_eq) (use that in simp) lemma Suc_div_le_mono [simp]: "m div n \ Suc m div n" by (simp add: div_le_mono) text \These lemmas collapse some needless occurrences of Suc: at least three Sucs, since two and fewer are rewritten back to Suc again! We already have some rules to simplify operands smaller than 3.\ lemma div_Suc_eq_div_add3 [simp]: "m div Suc (Suc (Suc n)) = m div (3 + n)" by (simp add: Suc3_eq_add_3) lemma mod_Suc_eq_mod_add3 [simp]: "m mod Suc (Suc (Suc n)) = m mod (3 + n)" by (simp add: Suc3_eq_add_3) lemma Suc_div_eq_add3_div: "Suc (Suc (Suc m)) div n = (3 + m) div n" by (simp add: Suc3_eq_add_3) lemma Suc_mod_eq_add3_mod: "Suc (Suc (Suc m)) mod n = (3 + m) mod n" by (simp add: Suc3_eq_add_3) lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v lemma (in field_char_0) of_nat_div: "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)" proof - have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)" unfolding of_nat_add by (cases "n = 0") simp_all then show ?thesis by simp qed text \An ``induction'' law for modulus arithmetic.\ lemma mod_induct [consumes 3, case_names step]: "P m" if "P n" and "n < p" and "m < p" and step: "\n. n < p \ P n \ P (Suc n mod p)" using \m < p\ proof (induct m) case 0 show ?case proof (rule ccontr) assume "\ P 0" from \n < p\ have "0 < p" by simp from \n < p\ obtain m where "0 < m" and "p = n + m" by (blast dest: less_imp_add_positive) with \P n\ have "P (p - m)" by simp moreover have "\ P (p - m)" using \0 < m\ proof (induct m) case 0 then show ?case by simp next case (Suc m) show ?case proof assume P: "P (p - Suc m)" with \\ P 0\ have "Suc m < p" by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma funpow_mod_eq: \<^marker>\contributor \Lars Noschinski\\ \(f ^^ (m mod n)) x = (f ^^ m) x\ if \(f ^^ n) x = x\ proof - have \(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\ by simp also have \\ = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\ by (simp only: funpow_add funpow_mult ac_simps) simp also have \((f ^^ n) ^^ q) x = x\ for q by (induction q) (use \(f ^^ n) x = x\ in simp_all) finally show ?thesis by simp qed lemma mod_eq_dvd_iff_nat: \m mod q = n mod q \ q dvd m - n\ (is \?P \ ?Q\) if \m \ n\ for m n q :: nat proof assume ?Q then obtain s where \m - n = q * s\ .. with that have \m = q * s + n\ by simp then show ?P by simp next assume ?P have \m - n = m div q * q + m mod q - (n div q * q + n mod q)\ by simp also have \\ = q * (m div q - n div q)\ by (simp only: algebra_simps \?P\) finally show ?Q .. qed lemma mod_eq_iff_dvd_symdiff_nat: \m mod q = n mod q \ q dvd nat \int m - int n\\ by (auto simp add: abs_if mod_eq_dvd_iff_nat nat_diff_distrib dest: sym intro: sym) lemma mod_eq_nat1E: fixes m n q :: nat assumes "m mod q = n mod q" and "m \ n" obtains s where "m = n + q * s" proof - from assms have "q dvd m - n" by (simp add: mod_eq_dvd_iff_nat) then obtain s where "m - n = q * s" .. with \m \ n\ have "m = n + q * s" by simp with that show thesis . qed lemma mod_eq_nat2E: fixes m n q :: nat assumes "m mod q = n mod q" and "n \ m" obtains s where "n = m + q * s" using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" (is "?lhs = ?rhs") proof assume H: "x mod n = y mod n" { assume xy: "x \ y" from H have th: "y mod n = x mod n" by simp from mod_eq_nat1E [OF th xy] obtain q where "y = x + n * q" . then have "x + n * q = y + n * 0" by simp then have "\q1 q2. x + n * q1 = y + n * q2" by blast } moreover { assume xy: "y \ x" from mod_eq_nat1E [OF H xy] obtain q where "x = y + n * q" . then have "x + n * 0 = y + n * q" by simp then have "\q1 q2. x + n * q1 = y + n * q2" by blast } ultimately show ?rhs using linear[of x y] by blast next assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp thus ?lhs by simp qed subsection \Elementary euclidean division on \<^typ>\int\\ subsubsection \Basic instantiation\ instantiation int :: "{normalization_semidom, idom_modulo}" begin definition normalize_int :: \int \ int\ where [simp]: \normalize = (abs :: int \ int)\ definition unit_factor_int :: \int \ int\ where [simp]: \unit_factor = (sgn :: int \ int)\ definition divide_int :: \int \ int \ int\ where \k div l = (sgn k * sgn l * int (nat \k\ div nat \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k))\ lemma divide_int_unfold: \(sgn k * int m) div (sgn l * int n) = (sgn k * sgn l * int (m div n) - of_bool ((k = 0 \ m = 0) \ l \ 0 \ n \ 0 \ sgn k \ sgn l \ \ n dvd m))\ by (simp add: divide_int_def sgn_mult nat_mult_distrib abs_mult sgn_eq_0_iff ac_simps) definition modulo_int :: \int \ int \ int\ where \k mod l = sgn k * int (nat \k\ mod nat \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ lemma modulo_int_unfold: \(sgn k * int m) mod (sgn l * int n) = sgn k * int (m mod (of_bool (l \ 0) * n)) + (sgn l * int n) * of_bool ((k = 0 \ m = 0) \ sgn k \ sgn l \ \ n dvd m)\ by (auto simp add: modulo_int_def sgn_mult abs_mult) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) next fix k l :: int obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (simp add: divide_int_unfold modulo_int_unfold algebra_simps modulo_nat_def of_nat_diff) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end subsubsection \Algebraic foundations\ lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all subsubsection \Basic conversions\ lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (auto simp add: divide_int_def) lemma div_eq_div_abs: \k div l = sgn k * sgn l * (\k\ div \l\) - of_bool (l \ 0 \ sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: divide_int_def [of k l] div_abs_eq_div_nat) lemma div_abs_eq: \\k\ div \l\ = sgn k * sgn l * (k div l + of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (simp add: div_eq_div_abs [of k l] ac_simps) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma mod_eq_mod_abs: \k mod l = sgn k * (\k\ mod \l\) + l * of_bool (sgn k \ sgn l \ \ l dvd k)\ for k l :: int by (simp add: modulo_int_def [of k l] mod_abs_eq_div_nat) lemma mod_abs_eq: \\k\ mod \l\ = sgn k * (k mod l - l * of_bool (sgn k \ sgn l \ \ l dvd k))\ for k l :: int by (auto simp: mod_eq_mod_abs [of k l]) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" using assms by (simp add: sgn_mult abs_mult sgn_0_0 divide_int_def [of "sgn v * \k\" "sgn v * \l\"] flip: div_abs_eq_div_nat) lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" using assms by (auto simp add: div_abs_eq) lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" using assms by (auto simp add: div_abs_eq ac_simps) lemma div_noneq_sgn_abs: fixes k l :: int assumes "l \ 0" assumes "sgn k \ sgn l" shows "k div l = - (\k\ div \l\) - of_bool (\ l dvd k)" using assms by (auto simp add: div_abs_eq ac_simps sgn_0_0 dest!: sgn_not_eq_imp) subsubsection \Euclidean division\ instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold abs_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) (simp add: sgn_0_0) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (auto simp add: modulo_int_unfold sgn_mult mod_greater_zero_iff_not_dvd simp flip: right_diff_distrib dest!: sgn_not_eq_imp) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: * q l divide_int_unfold) (auto simp add: sgn_mult ac_simps) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) end lemma euclidean_relation_intI [case_names by0 divides euclidean_relation]: \(k div l, k mod l) = (q, r)\ if by0': \l = 0 \ q = 0 \ r = k\ and divides': \l \ 0 \ l dvd k \ r = 0 \ k = q * l\ and euclidean_relation': \l \ 0 \ \ l dvd k \ sgn r = sgn l \ \r\ < \l\ \ k = q * l + r\ for k l :: int -proof (cases l q r k rule: euclidean_relationI) +proof (induction rule: euclidean_relationI) case by0 then show ?case by (rule by0') next case divides then show ?case by (rule divides') next case euclidean_relation with euclidean_relation' have \sgn r = sgn l\ \\r\ < \l\\ \k = q * l + r\ by simp_all from \sgn r = sgn l\ \l \ 0\ have \division_segment r = division_segment l\ by (simp add: division_segment_int_def sgn_if split: if_splits) with \\r\ < \l\\ \k = q * l + r\ show ?case by simp qed subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ proof (cases \m = 0 \ n = 0\) case True then show ?thesis by auto next case False then have \m > 0\ \n > 0\ by simp_all show ?thesis proof (cases \of_nat m * of_nat n dvd a\) case True then obtain b where \a = (of_nat m * of_nat n) * b\ .. then have \a = of_nat m * (of_nat n * b)\ by (simp add: ac_simps) then show ?thesis by simp next case False define q where \q = a div (of_nat m * of_nat n)\ define r where \r = a mod (of_nat m * of_nat n)\ from \m > 0\ \n > 0\ \\ of_nat m * of_nat n dvd a\ r_def have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod) with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with \m > 0\ \n > 0\ r_def have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with \m > 0\ \n > 0\ q_def have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp moreover have \a = q * (of_nat m * of_nat n) + r\ by (simp add: q_def r_def div_mult_mod_eq) ultimately show \a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\ using q_def [symmetric] div_plus_div_distrib_dvd_right [of \of_nat m\ \q * (of_nat m * of_nat n)\ r] by (simp add: ac_simps) qed qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np) subsection \More on euclidean division on \<^typ>\int\\ subsubsection \Trivial reduction steps\ lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_pos_neg_trivial: \k div l = - 1\ (is ?Q) and mod_pos_neg_trivial: \k mod l = k + l\ (is ?R) if \0 < k\ and \k + l \ 0\ for k l :: int proof - from that have \l < 0\ by simp have \(k div l, k mod l) = (- 1, k + l)\ - proof (cases l \- 1 :: int\ \k + l\ k rule: euclidean_relation_intI) + proof (induction rule: euclidean_relation_intI) case by0 with \l < 0\ show ?case by simp next case divides from \l dvd k\ obtain j where \k = l * j\ .. with \l < 0\ \0 < k\ have \j < 0\ by (simp add: zero_less_mult_iff) moreover from \k + l \ 0\ \k = l * j\ have \l * (j + 1) \ 0\ by (simp add: algebra_simps) with \l < 0\ have \j + 1 \ 0\ by (simp add: mult_le_0_iff) with \j < 0\ have \j = - 1\ by simp with \k = l * j\ show ?case by simp next case euclidean_relation with \k + l \ 0\ have \k + l < 0\ by (auto simp add: less_le add_eq_0_iff) with \0 < k\ show ?case by simp qed then show ?Q and ?R by simp_all qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ subsubsection \More uniqueness rules\ lemma fixes a b q r :: int assumes \a = b * q + r\ \0 \ r\ \r < b\ shows int_div_pos_eq: \a div b = q\ (is ?Q) and int_mod_pos_eq: \a mod b = r\ (is ?R) proof - - from assms have \(a div b, a mod b) = (q, r)\ - by (cases b q r a rule: euclidean_relation_intI) - (auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le) + have \(a div b, a mod b) = (q, r)\ + by (induction rule: euclidean_relation_intI) + (use assms in \auto simp add: ac_simps dvd_add_left_iff sgn_1_pos le_less dest: zdvd_imp_le\) then show ?Q and ?R by simp_all qed lemma int_div_neg_eq: \a div b = q\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int using that int_div_pos_eq [of a \- b\ \- q\ \- r\] by simp_all lemma int_mod_neg_eq: \a mod b = r\ if \a = b * q + r\ \r \ 0\ \b < r\ for a b q r :: int using that int_div_neg_eq [of a b q r] by simp subsubsection \Laws for unary minus\ lemma zmod_zminus1_not_zero: fixes k l :: int shows "- k mod l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zmod_zminus2_not_zero: fixes k l :: int shows "k mod - l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zdiv_zminus1_eq_if: \(- a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that sgn_not_eq_imp [of b \- a\] by (cases \a = 0\) (auto simp add: div_eq_div_abs [of \- a\ b] div_eq_div_abs [of a b] sgn_eq_0_iff) lemma zdiv_zminus2_eq_if: \a div (- b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)\ if \b \ 0\ for a b :: int using that by (auto simp add: zdiv_zminus1_eq_if div_minus_right) lemma zmod_zminus1_eq_if: \(- a) mod b = (if a mod b = 0 then 0 else b - (a mod b))\ for a b :: int by (cases \b = 0\) (auto simp flip: minus_div_mult_eq_mod simp add: zdiv_zminus1_eq_if algebra_simps) lemma zmod_zminus2_eq_if: \a mod (- b) = (if a mod b = 0 then 0 else (a mod b) - b)\ for a b :: int by (auto simp add: zmod_zminus1_eq_if mod_minus_right) subsubsection \Borders\ lemma pos_mod_bound [simp]: "k mod l < l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_pos) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: mod_greater_zero_iff_not_dvd sgn_1_neg) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (auto simp add: sgn_1_pos) qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp moreover have \int (m mod n) \ int n\ using \Suc q = n\ by simp then have \sgn s * int (m mod n) \ int n\ by (cases s \0::int\ rule: linorder_cases) simp_all ultimately show ?thesis by (simp only: modulo_int_unfold) auto qed subsubsection \Splitting Rules for div and mod\ lemma split_zdiv: \P (n div k) \ (k = 0 \ P 0) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P i)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P i))\ (is ?div) and split_zmod: \Q (n mod k) \ (k = 0 \ Q n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ Q j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ Q j))\ (is ?mod) for n k :: int proof - have *: \R (n div k) (n mod k) \ (k = 0 \ R 0 n) \ (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ R i j)) \ (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ R i j))\ for R by (cases \k = 0\) (auto simp add: linorder_class.neq_iff) from * [of \\q _. P q\] show ?div . from * [of \\_ r. Q r\] show ?mod . qed text \Enable (lin)arith to deal with \<^const>\divide\ and \<^const>\modulo\ when these are applied to some constant that is of the form \<^term>\numeral k\:\ declare split_zdiv [of _ _ \numeral n\, linarith_split] for n declare split_zdiv [of _ _ \- numeral n\, linarith_split] for n declare split_zmod [of _ _ \numeral n\, linarith_split] for n declare split_zmod [of _ _ \- numeral n\, linarith_split] for n lemma zdiv_eq_0_iff: "i div k = 0 \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" (is "?L = ?R") for i k :: int proof assume ?L moreover have "?L \ ?R" by (rule split_zdiv [THEN iffD2]) simp ultimately show ?R by blast next assume ?R then show ?L by auto qed lemma zmod_trivial_iff: fixes i k :: int shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" proof - have "i mod k = i \ i div k = 0" using div_mult_mod_eq [of i k] by safe auto with zdiv_eq_0_iff show ?thesis by simp qed subsubsection \Algebraic rewrites\ lemma zdiv_zmult2_eq: \a div (b * c) = (a div b) div c\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using div_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using div_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma zdiv_zmult2_eq': \k div (l * j) = ((sgn j * k) div l) div \j\\ for k l j :: int proof - have \k div (l * j) = (sgn j * k) div (sgn j * (l * j))\ by (simp add: sgn_0_0) also have \sgn j * (l * j) = l * \j\\ by (simp add: mult.left_commute [of _ l] abs_sgn) (simp add: ac_simps) also have \(sgn j * k) div (l * \j\) = ((sgn j * k) div l) div \j\\ by (simp add: zdiv_zmult2_eq) finally show ?thesis . qed lemma zmod_zmult2_eq: \a mod (b * c) = b * (a div b mod c) + a mod b\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using mod_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using mod_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma half_nonnegative_int_iff [simp]: \k div 2 \ 0 \ k \ 0\ for k :: int by auto lemma half_negative_int_iff [simp]: \k div 2 < 0 \ k < 0\ for k :: int by auto subsubsection \Distributive laws for conversions.\ lemma zdiv_int: "int (a div b) = int a div int b" by (fact of_nat_div) lemma zmod_int: "int (a mod b) = int a mod int b" by (fact of_nat_mod) lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) subsubsection \Monotonicity in the First Argument (Dividend)\ lemma zdiv_mono1: \a div b \ a' div b\ if \a \ a'\ \0 < b\ for a b b' :: int proof - from \a \ a'\ have \b * (a div b) + a mod b \ b * (a' div b) + a' mod b\ by simp then have \b * (a div b) \ (a' mod b - a mod b) + b * (a' div b)\ by (simp add: algebra_simps) moreover have \a' mod b < b + a mod b\ by (rule less_le_trans [of _ b]) (use \0 < b\ in simp_all) ultimately have \b * (a div b) < b * (1 + a' div b)\ by (simp add: distrib_left) with \0 < b\ have \a div b < 1 + a' div b\ by (simp add: mult_less_cancel_left) then show ?thesis by simp qed lemma zdiv_mono1_neg: \a' div b \ a div b\ if \a \ a'\ \b < 0\ for a a' b :: int using that zdiv_mono1 [of \- a'\ \- a\ \- b\] by simp subsubsection \Monotonicity in the Second Argument (Divisor)\ lemma zdiv_mono2: \a div b \ a div b'\ if \0 \ a\ \0 < b'\ \b' \ b\ for a b b' :: int proof - define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ then have *: \b * q + r = b' * q' + r'\ \0 \ b' * q' + r'\ \r' < b'\ \0 \ r\ \0 < b'\ \b' \ b\ using that by simp_all have \0 < b' * (q' + 1)\ using * by (simp add: distrib_left) with * have \0 \ q'\ by (simp add: zero_less_mult_iff) moreover have \b * q = r' - r + b' * q'\ using * by linarith ultimately have \b * q < b * (q' + 1)\ using mult_right_mono * unfolding distrib_left by fastforce with * have \q \ q'\ by (simp add: mult_less_cancel_left_pos) with ** show ?thesis by simp qed lemma zdiv_mono2_neg: \a div b' \ a div b\ if \a < 0\ \0 < b'\ \b' \ b\ for a b b' :: int proof - define q q' r r' where **: \q = a div b\ \q' = a div b'\ \r = a mod b\ \r' = a mod b'\ then have *: \b * q + r = b' * q' + r'\ \b' * q' + r' < 0\ \r < b\ \0 \ r'\ \0 < b'\ \b' \ b\ using that by simp_all have \b' * q' < 0\ using * by linarith with * have \q' \ 0\ by (simp add: mult_less_0_iff) have \b * q' \ b' * q'\ by (simp add: \q' \ 0\ * mult_right_mono_neg) then have "b * q' < b * (q + 1)" using * by (simp add: distrib_left) then have \q' \ q\ using * by (simp add: mult_less_cancel_left) then show ?thesis by (simp add: **) qed subsubsection \Quotients of Signs\ lemma div_eq_minus1: \0 < b \ - 1 div b = - 1\ for b :: int by (simp add: divide_int_def) lemma zmod_minus1: \0 < b \ - 1 mod b = b - 1\ for b :: int by (auto simp add: modulo_int_def) lemma minus_mod_int_eq: \- k mod l = l - 1 - (k - 1) mod l\ if \l \ 0\ for k l :: int proof (cases \l = 0\) case True then show ?thesis by simp next case False with that have \l > 0\ by simp then show ?thesis proof (cases \l dvd k\) case True then obtain j where \k = l * j\ .. moreover have \(l * j mod l - 1) mod l = l - 1\ using \l > 0\ by (simp add: zmod_minus1) then have \(l * j - 1) mod l = l - 1\ by (simp only: mod_simps) ultimately show ?thesis by simp next case False moreover have 1: \0 < k mod l\ using \0 < l\ False le_less by fastforce moreover have 2: \k mod l < 1 + l\ using \0 < l\ pos_mod_bound[of l k] by linarith from 1 2 \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ by (simp add: zmod_trivial_iff) ultimately show ?thesis by (simp only: zmod_zminus1_eq_if) (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) qed qed lemma div_neg_pos_less0: \a div b < 0\ if \a < 0\ \0 < b\ for a b :: int proof - have "a div b \ - 1 div b" using zdiv_mono1 that by auto also have "... \ -1" by (simp add: that(2) div_eq_minus1) finally show ?thesis by force qed lemma div_nonneg_neg_le0: \a div b \ 0\ if \0 \ a\ \b < 0\ for a b :: int using that by (auto dest: zdiv_mono1_neg) lemma div_nonpos_pos_le0: \a div b \ 0\ if \a \ 0\ \0 < b\ for a b :: int using that by (auto dest: zdiv_mono1) text\Now for some equivalences of the form \a div b >=< 0 \ \\ conditional upon the sign of \a\ or \b\. There are many more. They should all be simp rules unless that causes too much search.\ lemma pos_imp_zdiv_nonneg_iff: \0 \ a div b \ 0 \ a\ if \0 < b\ for a b :: int proof assume \0 \ a div b\ show \0 \ a\ proof (rule ccontr) assume \\ 0 \ a\ then have \a < 0\ by simp then have \a div b < 0\ using that by (rule div_neg_pos_less0) with \0 \ a div b\ show False by simp qed next assume "0 \ a" then have "0 div b \ a div b" using zdiv_mono1 that by blast then show "0 \ a div b" by auto qed lemma neg_imp_zdiv_nonneg_iff: \0 \ a div b \ a \ 0\ if \b < 0\ for a b :: int using that pos_imp_zdiv_nonneg_iff [of \- b\ \- a\] by simp lemma pos_imp_zdiv_pos_iff: \0 < (i::int) div k \ k \ i\ if \0 < k\ for i k :: int using that pos_imp_zdiv_nonneg_iff [of k i] zdiv_eq_0_iff [of i k] by arith lemma pos_imp_zdiv_neg_iff: \a div b < 0 \ a < 0\ if \0 < b\ for a b :: int \ \But not \<^prop>\a div b \ 0 \ a \ 0\; consider \<^prop>\a = 1\, \<^prop>\b = 2\ when \<^prop>\a div b = 0\.\ using that by (simp add: pos_imp_zdiv_nonneg_iff flip: linorder_not_le) lemma neg_imp_zdiv_neg_iff: \ \But not \<^prop>\a div b \ 0 \ 0 \ a\; consider \<^prop>\a = - 1\, \<^prop>\b = - 2\ when \<^prop>\a div b = 0\.\ \a div b < 0 \ 0 < a\ if \b < 0\ for a b :: int using that by (simp add: neg_imp_zdiv_nonneg_iff flip: linorder_not_le) lemma nonneg1_imp_zdiv_pos_iff: \a div b > 0 \ a \ b \ b > 0\ if \0 \ a\ for a b :: int proof - have "0 < a div b \ b \ a" using div_pos_pos_trivial[of a b] that by arith moreover have "0 < a div b \ b > 0" using that div_nonneg_neg_le0[of a b] by (cases "b=0"; force) moreover have "b \ a \ 0 < b \ 0 < a div b" using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp ultimately show ?thesis by blast qed lemma zmod_le_nonneg_dividend: \m mod k \ m\ if \(m::int) \ 0\ for m k :: int proof - from that have \m > 0 \ m = 0\ by auto then show ?thesis proof assume \m = 0\ then show ?thesis by simp next assume \m > 0\ then show ?thesis proof (cases k \0::int\ rule: linorder_cases) case less moreover define l where \l = - k\ ultimately have \l > 0\ by simp with \m > 0\ have \int (nat m mod nat l) \ m\ by (simp flip: le_nat_iff) then have \int (nat m mod nat l) - l \ m\ using \l > 0\ by simp with \m > 0\ \l > 0\ show ?thesis by (simp add: modulo_int_def l_def flip: le_nat_iff) qed (simp_all add: modulo_int_def flip: le_nat_iff) qed qed lemma sgn_div_eq_sgn_mult: \sgn (k div l) = of_bool (k div l \ 0) * sgn (k * l)\ for k l :: int proof (cases \k div l = 0\) case True then show ?thesis by simp next case False have \0 \ \k\ div \l\\ by (cases \l = 0\) (simp_all add: pos_imp_zdiv_nonneg_iff) then have \\k\ div \l\ \ 0 \ 0 < \k\ div \l\\ by (simp add: less_le) also have \\ \ \k\ \ \l\\ using False nonneg1_imp_zdiv_pos_iff by auto finally have *: \\k\ div \l\ \ 0 \ \l\ \ \k\\ . show ?thesis using \0 \ \k\ div \l\\ False by (auto simp add: div_eq_div_abs [of k l] div_eq_sgn_abs [of k l] sgn_mult sgn_1_pos sgn_1_neg sgn_eq_0_iff nonneg1_imp_zdiv_pos_iff * dest: sgn_not_eq_imp) qed subsubsection \Further properties\ lemma div_int_pos_iff: "k div l \ 0 \ k = 0 \ l = 0 \ k \ 0 \ l \ 0 \ k < 0 \ l < 0" for k l :: int proof (cases "k = 0 \ l = 0") case False then have *: "k \ 0" "l \ 0" by auto then have "0 \ k div l \ \ k < 0 \ 0 \ l" by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) then show ?thesis using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) qed auto lemma mod_int_pos_iff: \k mod l \ 0 \ l dvd k \ l = 0 \ k \ 0 \ l > 0\ for k l :: int proof (cases "l > 0") case False then show ?thesis by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \auto simp add: le_less not_less\) qed auto lemma abs_div: \\x div y\ = \x\ div \y\\ if \y dvd x\ for x y :: int using that by (cases \y = 0\) (auto simp add: abs_mult) lemma int_power_div_base: \<^marker>\contributor \Matthias Daum\\ \k ^ m div k = k ^ (m - Suc 0)\ if \0 < m\ \0 < k\ for k :: int using that by (cases m) simp_all lemma int_div_less_self: \<^marker>\contributor \Matthias Daum\\ \x div k < x\ if \0 < x\ \1 < k\ for x k :: int proof - from that have \nat (x div k) = nat x div nat k\ by (simp add: nat_div_distrib) also from that have \nat x div nat k < nat x\ by simp finally show ?thesis by simp qed subsubsection \Computing \div\ and \mod\ by shifting\ lemma div_pos_geq: \k div l = (k - l) div l + 1\ if \0 < l\ \l \ k\ for k l :: int proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with that show ?thesis by (simp add: div_add_self2) qed lemma mod_pos_geq: \k mod l = (k - l) mod l\ if \0 < l\ \l \ k\ for k l :: int proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with that show ?thesis by simp qed lemma pos_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = b div a\ (is ?Q) and pos_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)\ (is ?R) if \0 \ a\ for a b :: int proof - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = (b div a, 1 + 2 * (b mod a))\ - proof (cases \2 * a\ \b div a\ \1 + 2 * (b mod a)\ \1 + 2 * b\ rule: euclidean_relation_intI) + proof (induction rule: euclidean_relation_intI) case by0 then show ?case by simp next case divides have \2 dvd (2 * a)\ by simp then have \2 dvd (1 + 2 * b)\ using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) then have \2 dvd (1 + b * 2)\ by (simp add: ac_simps) then have \is_unit (2 :: int)\ by simp then show ?case by simp next case euclidean_relation with that have \a > 0\ by simp moreover have \b mod a < a\ using \a > 0\ by simp then have \1 + 2 * (b mod a) < 2 * a\ by simp moreover have \2 * (b mod a) + a * (2 * (b div a)) = 2 * (b div a * a + b mod a)\ by (simp only: algebra_simps) moreover have \0 \ 2 * (b mod a)\ using \a > 0\ by simp ultimately show ?case by (simp add: algebra_simps) qed then show ?Q and ?R by simp_all qed lemma neg_zdiv_mult_2: \(1 + 2 * b) div (2 * a) = (b + 1) div a\ (is ?Q) and neg_zmod_mult_2: \(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1\ (is ?R) if \a \ 0\ for a b :: int proof - have \((1 + 2 * b) div (2 * a), (1 + 2 * b) mod (2 * a)) = ((b + 1) div a, 2 * ((b + 1) mod a) - 1)\ - proof (cases \2 * a\ \(b + 1) div a\ \2 * ((b + 1) mod a) - 1\ \1 + 2 * b\ rule: euclidean_relation_intI) + proof (induction rule: euclidean_relation_intI) case by0 then show ?case by simp next case divides have \2 dvd (2 * a)\ by simp then have \2 dvd (1 + 2 * b)\ using \2 * a dvd 1 + 2 * b\ by (rule dvd_trans) then have \2 dvd (1 + b * 2)\ by (simp add: ac_simps) then have \is_unit (2 :: int)\ by simp then show ?case by simp next case euclidean_relation with that have \a < 0\ by simp moreover have \(b + 1) mod a > a\ using \a < 0\ by simp then have \2 * ((b + 1) mod a) > 1 + 2 * a\ by simp moreover have \((1 + b) mod a) \ 0\ using \a < 0\ by simp then have \2 * ((1 + b) mod a) \ 0\ by simp moreover have \2 * ((1 + b) mod a) + a * (2 * ((1 + b) div a)) = 2 * ((1 + b) div a * a + (1 + b) mod a)\ by (simp only: algebra_simps) ultimately show ?case by (simp add: algebra_simps sgn_mult abs_mult) qed then show ?Q and ?R by simp_all qed lemma zdiv_numeral_Bit0 [simp]: \numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = numeral v div (numeral w :: int)\ unfolding numeral.simps unfolding mult_2 [symmetric] by (rule div_mult_mult1) simp lemma zdiv_numeral_Bit1 [simp]: \numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = (numeral v div (numeral w :: int))\ unfolding numeral.simps unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zdiv_mult_2) simp lemma zmod_numeral_Bit0 [simp]: \numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = (2::int) * (numeral v mod numeral w)\ unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] by (rule mod_mult_mult1) lemma zmod_numeral_Bit1 [simp]: \numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = 2 * (numeral v mod numeral w) + (1::int)\ unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zmod_mult_2) simp subsection \Generic symbolic computations\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. \ class unique_euclidean_semiring_with_nat_division = unique_euclidean_semiring_with_nat + fixes divmod :: \num \ num \ 'a \ 'a\ and divmod_step :: \'a \ 'a \ 'a \ 'a \ 'a\ \ \ These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ assumes divmod_def: \divmod m n = (numeral m div numeral n, numeral m mod numeral n)\ and divmod_step_def [simp]: \divmod_step l (q, r) = (if euclidean_size l \ euclidean_size r then (2 * q + 1, r - l) else (2 * q, r))\ \ \ This is a formulation of one step (referring to one digit position) in school-method division: compare the dividend at the current digit position with the remainder from previous division steps and evaluate accordingly.\ begin lemma fst_divmod: \fst (divmod m n) = numeral m div numeral n\ by (simp add: divmod_def) lemma snd_divmod: \snd (divmod m n) = numeral m mod numeral n\ by (simp add: divmod_def) text \ Following a formulation of school-method division. If the divisor is smaller than the dividend, terminate. If not, shift the dividend to the right until termination occurs and then reiterate single division steps in the opposite direction. \ lemma divmod_divmod_step: \divmod m n = (if m < n then (0, numeral m) else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\ proof (cases \m < n\) case True then show ?thesis by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod) next case False define r s t where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ \t = 2 * s\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 n) = of_nat t\ and \\ s \ r mod s\ by (simp_all add: not_le) have t: \2 * (r div t) = r div s - r div s mod 2\ \r mod t = s * (r div s mod 2) + r mod s\ by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Division.div_mult2_eq \t = 2 * s\) (use mod_mult2_eq [of r s 2] in \simp add: ac_simps \t = 2 * s\\) have rs: \r div s mod 2 = 0 \ r div s mod 2 = Suc 0\ by auto from \\ s \ r mod s\ have \s \ r mod t \ r div s = Suc (2 * (r div t)) \ r mod s = r mod t - s\ using rs by (auto simp add: t) moreover have \r mod t < s \ r div s = 2 * (r div t) \ r mod s = r mod t\ using rs by (auto simp add: t) ultimately show ?thesis by (simp add: divmod_def prod_eq_iff split_def Let_def not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *) (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff) qed text \The division rewrite proper -- first, trivial results involving \1\\ lemma divmod_trivial [simp]: "divmod m Num.One = (numeral m, 0)" "divmod num.One (num.Bit0 n) = (0, Numeral1)" "divmod num.One (num.Bit1 n) = (0, Numeral1)" using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) text \Division by an even number is a right-shift\ lemma divmod_cancel [simp]: \divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r))\ (is ?P) \divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r + 1))\ (is ?Q) proof - define r s where \r = (numeral m :: nat)\ \s = (numeral n :: nat)\ then have *: \numeral m = of_nat r\ \numeral n = of_nat s\ \numeral (num.Bit0 m) = of_nat (2 * r)\ \numeral (num.Bit0 n) = of_nat (2 * s)\ \numeral (num.Bit1 m) = of_nat (Suc (2 * r))\ by simp_all have **: \Suc (2 * r) div 2 = r\ by simp show ?P and ?Q by (simp_all add: divmod_def *) (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc add: Euclidean_Division.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **) qed text \The really hard work\ lemma divmod_steps [simp]: "divmod (num.Bit0 m) (num.Bit1 n) = (if m \ n then (0, numeral (num.Bit0 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit0 m) (num.Bit0 (num.Bit1 n))))" "divmod (num.Bit1 m) (num.Bit1 n) = (if m < n then (0, numeral (num.Bit1 m)) else divmod_step (numeral (num.Bit1 n)) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps text \Special case: divisibility\ definition divides_aux :: "'a \ 'a \ bool" where "divides_aux qr \ snd qr = 0" lemma divides_aux_eq [simp]: "divides_aux (q, r) \ r = 0" by (simp add: divides_aux_def) lemma dvd_numeral_simp [simp]: "numeral m dvd numeral n \ divides_aux (divmod n m)" by (simp add: divmod_def mod_eq_0_iff_dvd) text \Generic computation of quotient and remainder\ lemma numeral_div_numeral [simp]: "numeral k div numeral l = fst (divmod k l)" by (simp add: fst_divmod) lemma numeral_mod_numeral [simp]: "numeral k mod numeral l = snd (divmod k l)" by (simp add: snd_divmod) lemma one_div_numeral [simp]: "1 div numeral n = fst (divmod num.One n)" by (simp add: fst_divmod) lemma one_mod_numeral [simp]: "1 mod numeral n = snd (divmod num.One n)" by (simp add: snd_divmod) end instantiation nat :: unique_euclidean_semiring_with_nat_division begin definition divmod_nat :: "num \ num \ nat \ nat" where divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_nat :: "nat \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ l then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (simp_all add: divmod'_nat_def divmod_step_nat_def) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: \Suc 0 div numeral Num.One = 1\ \Suc 0 div numeral (Num.Bit0 n) = 0\ \Suc 0 div numeral (Num.Bit1 n) = 0\ by simp_all lemma Suc_0_mod_numeral [simp]: \Suc 0 mod numeral Num.One = 0\ \Suc 0 mod numeral (Num.Bit0 n) = 1\ \Suc 0 mod numeral (Num.Bit1 n) = 1\ by simp_all instantiation int :: unique_euclidean_semiring_with_nat_division begin definition divmod_int :: "num \ num \ int \ int" where "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_int :: "int \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if \l\ \ \r\ then (2 * q + 1, r - l) else (2 * q, r))" instance by standard (auto simp add: divmod_int_def divmod_step_int_def) end declare divmod_algorithm_code [where ?'a = int, code] context begin qualified definition adjust_div :: "int \ int \ int" where "adjust_div qr = (let (q, r) = qr in q + of_bool (r \ 0))" qualified lemma adjust_div_eq [simp, code]: "adjust_div (q, r) = q + of_bool (r \ 0)" by (simp add: adjust_div_def) qualified definition adjust_mod :: "num \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)" lemma minus_numeral_div_numeral [simp]: "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma minus_numeral_mod_numeral [simp]: "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma numeral_div_minus_numeral [simp]: "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma numeral_mod_minus_numeral [simp]: "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma minus_one_div_numeral [simp]: "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" using minus_numeral_div_numeral [of Num.One n] by simp lemma minus_one_mod_numeral [simp]: "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)" using minus_numeral_mod_numeral [of Num.One n] by simp lemma one_div_minus_numeral [simp]: "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" using numeral_div_minus_numeral [of Num.One n] by simp lemma one_mod_minus_numeral [simp]: "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp lemma [code]: fixes k :: int shows "k div 0 = 0" "k mod 0 = k" "0 div k = 0" "0 mod k = 0" "k div Int.Pos Num.One = k" "k mod Int.Pos Num.One = 0" "k div Int.Neg Num.One = - k" "k mod Int.Neg Num.One = 0" "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)" "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" by simp_all end lemma divmod_BitM_2_eq [simp]: \divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\ by (cases m) simp_all subsubsection \Computation by simplification\ lemma euclidean_size_nat_less_eq_iff: \euclidean_size m \ euclidean_size n \ m \ n\ for m n :: nat by simp lemma euclidean_size_int_less_eq_iff: \euclidean_size k \ euclidean_size l \ \k\ \ \l\\ for k l :: int by auto simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | "numeral a div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \ let val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\If\); fun successful_rewrite ctxt ct = let val thm = Simplifier.rewrite ctxt ct in if Thm.is_reflexive thm then NONE else SOME thm end; in fn phi => let val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral one_div_minus_numeral one_mod_minus_numeral numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral numeral_div_minus_numeral numeral_mod_minus_numeral div_minus_minus mod_minus_minus Euclidean_Division.adjust_div_eq of_bool_eq one_neq_zero numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One case_prod_beta rel_simps Euclidean_Division.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral} @ [@{lemma "0 = 0 \ True" by simp}]); fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end end \ \ \ There is space for improvement here: the calculation itself could be carried out outside the logic, and a generic simproc (simplifier setup) for generic calculation would be helpful. \ subsubsection \Code generation\ context begin qualified definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" qualified lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = divmod_nat (m - n) n in (Suc q, r))" by (simp add: divmod_nat_def prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) qualified lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by (simp_all add: divmod_nat_def) end code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end