# HG changeset patch # User haftmann # Date 1371309563 -7200 # Node ID 3cc46b8cca5efb5f68311b5ac1b2740c9c7c3d7d # Parent 7f864f2219a97d4e766ad4b81bc7677a882d51f4 lifting for primitive definitions; explicit conversions from and to lists of coefficients, used for generated code; replaced recursion operator poly_rec by fold_coeffs, preferring function definitions for non-trivial recursions; prefer pre-existing gcd operation for gcd diff -r 7f864f2219a9 -r 3cc46b8cca5e NEWS --- a/NEWS Sat Jun 15 17:19:23 2013 +0200 +++ b/NEWS Sat Jun 15 17:19:23 2013 +0200 @@ -61,6 +61,17 @@ *** HOL *** +* Library/Polynomial.thy: + * Use lifting for primitive definitions. + * Explicit conversions from and to lists of coefficients, used for generated code. + * Replaced recursion operator poly_rec by fold_coeffs. + * Prefer pre-existing gcd operation for gcd. + * Fact renames: + poly_eq_iff ~> poly_eq_poly_eq_iff + poly_ext ~> poly_eqI + expand_poly_eq ~> poly_eq_iff +IMCOMPATIBILTIY. + * Reification and reflection: * Reification is now directly available in HOL-Main in structure "Reification". * Reflection now handles multiple lists with variables also. diff -r 7f864f2219a9 -r 3cc46b8cca5e src/HOL/Library/Fundamental_Theorem_Algebra.thy --- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy Sat Jun 15 17:19:23 2013 +0200 +++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy Sat Jun 15 17:19:23 2013 +0200 @@ -93,16 +93,17 @@ text{* Offsetting the variable in a polynomial gives another of same degree *} -definition - "offset_poly p h = poly_rec 0 (\a p q. smult h q + pCons a q) p" +definition offset_poly :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" +where + "offset_poly p h = fold_coeffs (\a q. smult h q + pCons a q) p 0" lemma offset_poly_0: "offset_poly 0 h = 0" - unfolding offset_poly_def by (simp add: poly_rec_0) + by (simp add: offset_poly_def) lemma offset_poly_pCons: "offset_poly (pCons a p) h = smult h (offset_poly p h) + pCons a (offset_poly p h)" - unfolding offset_poly_def by (simp add: poly_rec_pCons) + by (cases "p = 0 \ a = 0") (auto simp add: offset_poly_def) lemma offset_poly_single: "offset_poly [:a:] h = [:a:]" by (simp add: offset_poly_pCons offset_poly_0) @@ -644,7 +645,7 @@ let ?r = "smult (inverse ?a0) q" have lgqr: "psize q = psize ?r" using a00 unfolding psize_def degree_def - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) {assume h: "\x y. poly ?r x = poly ?r y" {fix x y from qr[rule_format, of x] @@ -887,9 +888,7 @@ proof- {assume pe: "p = 0" hence eq: "(\x. poly p x = (0::complex) \ poly q x = 0) \ q = 0" - apply auto - apply (rule poly_zero [THEN iffD1]) - by (rule ext, simp) + by (auto simp add: poly_all_0_iff_0) {assume "p dvd (q ^ (degree p))" then obtain r where r: "q ^ (degree p) = p * r" .. from r pe have False by simp} @@ -927,7 +926,7 @@ assume l: ?lhs from l[unfolded constant_def, rule_format, of _ "0"] have th: "poly p = poly [:poly p 0:]" apply - by (rule ext, simp) - then have "p = [:poly p 0:]" by (simp add: poly_eq_iff) + then have "p = [:poly p 0:]" by (simp add: poly_eq_poly_eq_iff) then have "degree p = degree [:poly p 0:]" by simp then show ?rhs by simp next @@ -1050,7 +1049,7 @@ lemma basic_cqe_conv_2b: "(\x. poly p x \ (0::complex)) \ (p \ 0)" proof- have "p = 0 \ poly p = poly 0" - by (simp add: poly_zero) + by (simp add: poly_eq_poly_eq_iff) also have "\ \ (\ (\x. poly p x \ 0))" by auto finally show "(\x. poly p x \ (0::complex)) \ p \ 0" by - (atomize (full), blast) @@ -1074,7 +1073,7 @@ shows "p dvd (q ^ n) \ p dvd r" proof- from h have "poly (q ^ n) = poly r" by auto - then have "(q ^ n) = r" by (simp add: poly_eq_iff) + then have "(q ^ n) = r" by (simp add: poly_eq_poly_eq_iff) thus "p dvd (q ^ n) \ p dvd r" by simp qed diff -r 7f864f2219a9 -r 3cc46b8cca5e src/HOL/Library/Poly_Deriv.thy --- a/src/HOL/Library/Poly_Deriv.thy Sat Jun 15 17:19:23 2013 +0200 +++ b/src/HOL/Library/Poly_Deriv.thy Sat Jun 15 17:19:23 2013 +0200 @@ -11,26 +11,41 @@ subsection {* Derivatives of univariate polynomials *} -definition - pderiv :: "'a::real_normed_field poly \ 'a poly" where - "pderiv = poly_rec 0 (\a p p'. p + pCons 0 p')" +function pderiv :: "'a::real_normed_field poly \ 'a poly" +where + [simp del]: "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 -lemma pderiv_0 [simp]: "pderiv 0 = 0" - unfolding pderiv_def by (simp add: poly_rec_0) +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)" - unfolding pderiv_def by (simp add: poly_rec_pCons) +lemma pderiv_pCons: + "pderiv (pCons a p) = p + pCons 0 (pderiv p)" + by (simp add: pderiv.simps) lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" apply (induct p arbitrary: n, simp) apply (simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) done +primrec pderiv_coeffs :: "'a::comm_monoid_add list \ 'a list" +where + "pderiv_coeffs [] = []" +| "pderiv_coeffs (x # xs) = plus_coeffs xs (cCons 0 (pderiv_coeffs xs))" + +lemma coeffs_pderiv [code abstract]: + "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" + by (rule sym, induct p) (simp_all add: pderiv_pCons coeffs_plus_eq_plus_coeffs cCons_def) + lemma pderiv_eq_0_iff: "pderiv p = 0 \ degree p = 0" apply (rule iffI) apply (cases p, simp) - apply (simp add: expand_poly_eq coeff_pderiv del: of_nat_Suc) - apply (simp add: expand_poly_eq coeff_pderiv coeff_eq_0) + apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) + apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) done lemma degree_pderiv: "degree (pderiv p) = degree p - 1" @@ -47,16 +62,16 @@ by (simp add: pderiv_pCons) lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" -by (rule poly_ext, simp add: coeff_pderiv algebra_simps) +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) lemma pderiv_minus: "pderiv (- p) = - pderiv p" -by (rule poly_ext, simp add: coeff_pderiv) +by (rule poly_eqI, simp add: coeff_pderiv) lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" -by (rule poly_ext, simp add: coeff_pderiv algebra_simps) +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" -by (rule poly_ext, simp add: coeff_pderiv algebra_simps) +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" apply (induct p) @@ -75,6 +90,7 @@ apply (simp add: algebra_simps) (* FIXME *) done + lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" by (simp add: DERIV_cmult mult_commute [of _ c]) diff -r 7f864f2219a9 -r 3cc46b8cca5e src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Sat Jun 15 17:19:23 2013 +0200 +++ b/src/HOL/Library/Polynomial.thy Sat Jun 15 17:19:23 2013 +0200 @@ -1,47 +1,201 @@ (* Title: HOL/Library/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin + Author: Florian Haftmann *) -header {* Univariate Polynomials *} +header {* Polynomials as type over a ring structure *} theory Polynomial -imports Main +imports Main GCD begin +subsection {* Auxiliary: operations for lists (later) representing coefficients *} + +definition strip_while :: "('a \ bool) \ 'a list \ 'a list" +where + "strip_while P = rev \ dropWhile P \ rev" + +lemma strip_while_Nil [simp]: + "strip_while P [] = []" + by (simp add: strip_while_def) + +lemma strip_while_append [simp]: + "\ P x \ strip_while P (xs @ [x]) = xs @ [x]" + by (simp add: strip_while_def) + +lemma strip_while_append_rec [simp]: + "P x \ strip_while P (xs @ [x]) = strip_while P xs" + by (simp add: strip_while_def) + +lemma strip_while_Cons [simp]: + "\ P x \ strip_while P (x # xs) = x # strip_while P xs" + by (induct xs rule: rev_induct) (simp_all add: strip_while_def) + +lemma strip_while_eq_Nil [simp]: + "strip_while P xs = [] \ (\x\set xs. P x)" + by (simp add: strip_while_def) + +lemma strip_while_eq_Cons_rec: + "strip_while P (x # xs) = x # strip_while P xs \ \ (P x \ (\x\set xs. P x))" + by (induct xs rule: rev_induct) (simp_all add: strip_while_def) + +lemma strip_while_not_last [simp]: + "\ P (last xs) \ strip_while P xs = xs" + by (cases xs rule: rev_cases) simp_all + +lemma split_strip_while_append: + fixes xs :: "'a list" + obtains ys zs :: "'a list" + where "strip_while P xs = ys" and "\x\set zs. P x" and "xs = ys @ zs" +proof (rule that) + show "strip_while P xs = strip_while P xs" .. + show "\x\set (rev (takeWhile P (rev xs))). P x" by (simp add: takeWhile_eq_all_conv [symmetric]) + have "rev xs = rev (strip_while P xs @ rev (takeWhile P (rev xs)))" + by (simp add: strip_while_def) + then show "xs = strip_while P xs @ rev (takeWhile P (rev xs))" + by (simp only: rev_is_rev_conv) +qed + + +definition nth_default :: "'a \ 'a list \ nat \ 'a" +where + "nth_default x xs n = (if n < length xs then xs ! n else x)" + +lemma nth_default_Nil [simp]: + "nth_default y [] n = y" + by (simp add: nth_default_def) + +lemma nth_default_Cons_0 [simp]: + "nth_default y (x # xs) 0 = x" + by (simp add: nth_default_def) + +lemma nth_default_Cons_Suc [simp]: + "nth_default y (x # xs) (Suc n) = nth_default y xs n" + by (simp add: nth_default_def) + +lemma nth_default_map_eq: + "f y = x \ nth_default x (map f xs) n = f (nth_default y xs n)" + by (simp add: nth_default_def) + +lemma nth_default_strip_while_eq [simp]: + "nth_default x (strip_while (HOL.eq x) xs) n = nth_default x xs n" +proof - + from split_strip_while_append obtain ys zs + where "strip_while (HOL.eq x) xs = ys" and "\z\set zs. x = z" and "xs = ys @ zs" by blast + then show ?thesis by (simp add: nth_default_def not_less nth_append) +qed + + +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 {* Almost everywhere zero functions *} + +definition almost_everywhere_zero :: "(nat \ 'a::zero) \ bool" +where + "almost_everywhere_zero f \ (\n. \i>n. f i = 0)" + +lemma almost_everywhere_zeroI: + "(\i. i > n \ f i = 0) \ almost_everywhere_zero f" + by (auto simp add: almost_everywhere_zero_def) + +lemma almost_everywhere_zeroE: + assumes "almost_everywhere_zero f" + obtains n where "\i. i > n \ f i = 0" +proof - + from assms have "\n. \i>n. f i = 0" by (simp add: almost_everywhere_zero_def) + then obtain n where "\i. i > n \ f i = 0" by blast + with that show thesis . +qed + +lemma almost_everywhere_zero_nat_case: + assumes "almost_everywhere_zero f" + shows "almost_everywhere_zero (nat_case a f)" + using assms + by (auto intro!: almost_everywhere_zeroI elim!: almost_everywhere_zeroE split: nat.split) + blast + +lemma almost_everywhere_zero_Suc: + assumes "almost_everywhere_zero f" + shows "almost_everywhere_zero (\n. f (Suc n))" +proof - + from assms obtain n where "\i. i > n \ f i = 0" by (erule almost_everywhere_zeroE) + then have "\i. i > n \ f (Suc i) = 0" by auto + then show ?thesis by (rule almost_everywhere_zeroI) +qed + + subsection {* Definition of type @{text poly} *} -definition "Poly = {f::nat \ 'a::zero. \n. \i>n. f i = 0}" - -typedef 'a poly = "Poly :: (nat => 'a::zero) set" +typedef 'a poly = "{f :: nat \ 'a::zero. almost_everywhere_zero f}" morphisms coeff Abs_poly - unfolding Poly_def by auto + unfolding almost_everywhere_zero_def by auto -(* FIXME should be named poly_eq_iff *) -lemma expand_poly_eq: "p = q \ (\n. coeff p n = coeff q n)" +setup_lifting (no_code) 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) -(* FIXME should be named poly_eqI *) -lemma poly_ext: "(\n. coeff p n = coeff q n) \ p = q" - by (simp add: expand_poly_eq) +lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" + by (simp add: poly_eq_iff) + +lemma coeff_almost_everywhere_zero: + "almost_everywhere_zero (coeff p)" + using coeff [of p] by simp subsection {* Degree of a polynomial *} -definition - degree :: "'a::zero poly \ nat" where +definition degree :: "'a::zero poly \ nat" +where "degree p = (LEAST n. \i>n. coeff p i = 0)" -lemma coeff_eq_0: "degree p < n \ coeff p n = 0" +lemma coeff_eq_0: + assumes "degree p < n" + shows "coeff p n = 0" proof - - have "coeff p \ Poly" - by (rule coeff) - hence "\n. \i>n. coeff p i = 0" - unfolding Poly_def by simp - hence "\i>degree p. coeff p i = 0" + from coeff_almost_everywhere_zero + have "\n. \i>n. coeff p i = 0" by (blast intro: almost_everywhere_zeroE) + then have "\i>degree p. coeff p i = 0" unfolding degree_def by (rule LeastI_ex) - moreover assume "degree p < n" - ultimately show ?thesis by simp + with assms show ?thesis by simp qed lemma le_degree: "coeff p n \ 0 \ n \ degree p" @@ -59,25 +213,28 @@ instantiation poly :: (zero) zero begin -definition - zero_poly_def: "0 = Abs_poly (\n. 0)" +lift_definition zero_poly :: "'a poly" + is "\_. 0" by (rule almost_everywhere_zeroI) simp instance .. + end -lemma coeff_0 [simp]: "coeff 0 n = 0" - unfolding zero_poly_def - by (simp add: Abs_poly_inverse Poly_def) +lemma coeff_0 [simp]: + "coeff 0 n = 0" + by transfer rule -lemma degree_0 [simp]: "degree 0 = 0" +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" + assumes "p \ 0" + shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from `p \ 0` have "\n. coeff p n \ 0" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) then obtain n where "coeff p n \ 0" .. hence "n \ degree p" by (rule le_degree) with `coeff p n \ 0` and `degree p = 0` @@ -93,68 +250,59 @@ 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" +lemma leading_coeff_0_iff [simp]: + "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0", simp, simp add: leading_coeff_neq_0) subsection {* List-style constructor for polynomials *} -definition - pCons :: "'a::zero \ 'a poly \ 'a poly" -where - "pCons a p = Abs_poly (nat_case a (coeff p))" +lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" + is "\a p. nat_case a (coeff p)" + using coeff_almost_everywhere_zero by (rule almost_everywhere_zero_nat_case) -syntax - "_poly" :: "args \ 'a poly" ("[:(_):]") +lemmas coeff_pCons = pCons.rep_eq -translations - "[:x, xs:]" == "CONST pCons x [:xs:]" - "[:x:]" == "CONST pCons x 0" - "[:x:]" <= "CONST pCons x (_constrain 0 t)" +lemma coeff_pCons_0 [simp]: + "coeff (pCons a p) 0 = a" + by transfer simp -lemma Poly_nat_case: "f \ Poly \ nat_case a f \ Poly" - unfolding Poly_def by (auto split: nat.split) - -lemma coeff_pCons: - "coeff (pCons a p) = nat_case a (coeff p)" - unfolding pCons_def - by (simp add: Abs_poly_inverse Poly_nat_case coeff) - -lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" +lemma coeff_pCons_Suc [simp]: + "coeff (pCons a p) (Suc n) = coeff p n" by (simp add: coeff_pCons) -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_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)" -apply (rule order_antisym [OF degree_pCons_le]) -apply (rule le_degree, simp) -done + apply (rule order_antisym [OF degree_pCons_le]) + apply (rule le_degree, simp) + done -lemma degree_pCons_0: "degree (pCons a 0) = 0" -apply (rule order_antisym [OF _ le0]) -apply (rule degree_le, simp add: coeff_pCons split: nat.split) -done +lemma degree_pCons_0: + "degree (pCons a 0) = 0" + apply (rule order_antisym [OF _ le0]) + apply (rule degree_le, simp add: coeff_pCons split: nat.split) + done lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" -apply (cases "p = 0", simp_all) -apply (rule order_antisym [OF _ le0]) -apply (rule degree_le, simp add: coeff_pCons split: nat.split) -apply (rule order_antisym [OF degree_pCons_le]) -apply (rule le_degree, simp) -done + apply (cases "p = 0", simp_all) + apply (rule order_antisym [OF _ le0]) + apply (rule degree_le, simp add: coeff_pCons split: nat.split) + apply (rule order_antisym [OF degree_pCons_le]) + apply (rule le_degree, simp) + done -lemma pCons_0_0 [simp, code_post]: "pCons 0 0 = 0" -by (rule poly_ext, simp add: coeff_pCons split: nat.split) +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) +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 @@ -162,23 +310,19 @@ assume "pCons a p = pCons b q" then have "\n. coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" by simp - then show "p = q" by (simp add: expand_poly_eq) + 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" +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 Poly_Suc: "f \ Poly \ (\n. f (Suc n)) \ Poly" - unfolding Poly_def - by (clarify, rule_tac x=n in exI, 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 (rule poly_ext) - (simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons - split: nat.split) + by transfer + (simp add: Abs_poly_inverse almost_everywhere_zero_Suc fun_eq_iff split: nat.split) qed lemma pCons_induct [case_names 0 pCons, induct type: poly]: @@ -208,52 +352,227 @@ qed -subsection {* Recursion combinator for polynomials *} +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 + "Poly [] = 0" +| "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) + +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 -function - poly_rec :: "'b \ ('a::zero \ 'a poly \ 'b \ 'b) \ 'a poly \ 'b" -where - poly_rec_pCons_eq_if [simp del]: - "poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)" -by (case_tac x, rename_tac q, case_tac q, auto) +lemma coeffs_pCons_eq_cCons [simp]: + "coeffs (pCons a p) = a ## coeffs p" +proof - + { fix ms :: "nat list" and f :: "nat \ 'a" and x :: "'a" + assume "\m\set ms. m > 0" + then have "map (nat_case x f) ms = map f (map (\n. n - 1) ms)" + by (induct ms) (auto, metis Suc_pred' nat_case_Suc) } + note * = this + show ?thesis + by (simp add: coeffs_def * upt_conv_Cons coeff_pCons map_decr_upt One_nat_def del: upt_Suc) +qed + +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) (simp_all add: cCons_def) + +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) + have "(\n. as \ replicate n 0) \ (\a\set as. a \ 0)" + using replicate_length_same [of as 0] by (auto dest: sym [of _ as]) + with Cons show ?case by auto +qed + +lemma last_coeffs_not_0: + "p \ 0 \ last (coeffs p) \ 0" + by (induct p) (auto simp add: cCons_def) + +lemma strip_while_coeffs [simp]: + "strip_while (HOL.eq 0) (coeffs p) = coeffs p" + by (cases "p = 0") (auto dest: last_coeffs_not_0 intro: strip_while_not_last) + +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 coeff_Poly_eq: + "coeff (Poly xs) n = nth_default 0 xs n" + apply (induct xs arbitrary: n) apply simp_all + by (metis nat_case_0 nat_case_Suc not0_implies_Suc nth_default_Cons_0 nth_default_Cons_Suc pCons.rep_eq) -termination poly_rec -by (relation "measure (degree \ snd \ snd)", simp) - (simp add: degree_pCons_eq) +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: "xs \ [] \ last xs \ 0" + shows "coeffs p = xs" +proof - + from coeff have "p = Poly xs" by (simp add: poly_eq_iff coeff_Poly_eq) + with zero show ?thesis by simp (cases xs, simp_all) +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 add: cCons_def) + +lemma [code abstract]: + "coeffs 0 = []" + by (fact coeffs_0_eq_Nil) + +lemma [code abstract]: + "coeffs (pCons a p) = a ## coeffs p" + by (fact coeffs_pCons_eq_cCons) + +instantiation poly :: ("{zero, equal}") equal +begin + +definition + [code]: "HOL.equal (p::'a poly) q \ HOL.equal (coeffs p) (coeffs q)" + +instance proof +qed (simp add: equal equal_poly_def coeffs_eq_iff) + +end + +lemma [code nbe]: + "HOL.equal (p :: _ poly) p \ True" + by (fact equal_refl) -lemma poly_rec_0: - "f 0 0 z = z \ poly_rec z f 0 = z" - using poly_rec_pCons_eq_if [of z f 0 0] by simp +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) + + +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 poly_rec_pCons: - "f 0 0 z = z \ poly_rec z f (pCons a p) = f a p (poly_rec z f p)" - by (simp add: poly_rec_pCons_eq_if poly_rec_0) +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 = fold_coeffs (\a f x. a + x * f x) p (\x. 0)" -- {* The Horner Schema *} + +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) subsection {* Monomials *} -definition - monom :: "'a \ nat \ 'a::zero poly" where - "monom a m = Abs_poly (\n. if m = n then a else 0)" +lift_definition monom :: "'a \ nat \ 'a::zero poly" + is "\a m n. if m = n then a else 0" + by (auto intro!: almost_everywhere_zeroI) + +lemma coeff_monom [simp]: + "coeff (monom a m) n = (if m = n then a else 0)" + by transfer rule -lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)" - unfolding monom_def - by (subst Abs_poly_inverse, auto simp add: Poly_def) +lemma monom_0: + "monom a 0 = pCons a 0" + by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) -lemma monom_0: "monom a 0 = pCons a 0" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" - by (rule poly_ext, 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_ext) simp + by (rule poly_eqI) simp lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) @@ -263,37 +582,47 @@ apply (rule le_degree, simp) done +lemma coeffs_monom [code abstract]: + "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" + by (induct n) (simp_all add: monom_0 monom_Suc) + +lemma fold_coeffs_monom [simp]: + "a \ 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: + fixes a x :: "'a::{comm_semiring_1}" + shows "poly (monom a n) x = a * x ^ n" + by (cases "a = 0", simp_all) + (induct n, simp_all add: mult.left_commute poly_def) + subsection {* Addition and subtraction *} instantiation poly :: (comm_monoid_add) comm_monoid_add begin -definition - plus_poly_def: - "p + q = Abs_poly (\n. coeff p n + coeff q n)" - -lemma Poly_add: - fixes f g :: "nat \ 'a" - shows "\f \ Poly; g \ Poly\ \ (\n. f n + g n) \ Poly" - unfolding Poly_def - apply (clarify, rename_tac m n) - apply (rule_tac x="max m n" in exI, simp) - done +lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" + is "\p q n. coeff p n + coeff q n" +proof (rule almost_everywhere_zeroI) + fix q p :: "'a poly" and i + assume "max (degree q) (degree p) < i" + then show "coeff p i + coeff q i = 0" + by (simp add: coeff_eq_0) +qed lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" - unfolding plus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_add) + 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: expand_poly_eq add_assoc) + by (simp add: poly_eq_iff add_assoc) show "p + q = q + p" - by (simp add: expand_poly_eq add_commute) + by (simp add: poly_eq_iff add_commute) show "0 + p = p" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) qed end @@ -302,60 +631,58 @@ proof fix p q r :: "'a poly" assume "p + q = p + r" thus "q = r" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) qed instantiation poly :: (ab_group_add) ab_group_add begin -definition - uminus_poly_def: - "- p = Abs_poly (\n. - coeff p n)" - -definition - minus_poly_def: - "p - q = Abs_poly (\n. coeff p n - coeff q n)" +lift_definition uminus_poly :: "'a poly \ 'a poly" + is "\p n. - coeff p n" +proof (rule almost_everywhere_zeroI) + fix p :: "'a poly" and i + assume "degree p < i" + then show "- coeff p i = 0" + by (simp add: coeff_eq_0) +qed -lemma Poly_minus: - fixes f :: "nat \ 'a" - shows "f \ Poly \ (\n. - f n) \ Poly" - unfolding Poly_def by simp - -lemma Poly_diff: - fixes f g :: "nat \ 'a" - shows "\f \ Poly; g \ Poly\ \ (\n. f n - g n) \ Poly" - unfolding diff_minus by (simp add: Poly_add Poly_minus) +lift_definition minus_poly :: "'a poly \ 'a poly \ 'a poly" + is "\p q n. coeff p n - coeff q n" +proof (rule almost_everywhere_zeroI) + fix q p :: "'a poly" and i + assume "max (degree q) (degree p) < i" + then show "coeff p i - coeff q i = 0" + by (simp add: coeff_eq_0) +qed lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" - unfolding uminus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_minus) + by (simp add: uminus_poly.rep_eq) lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" - unfolding minus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_diff) + by (simp add: minus_poly.rep_eq) instance proof fix p q :: "'a poly" show "- p + p = 0" - by (simp add: expand_poly_eq) + by (simp add: poly_eq_iff) show "p - q = p + - q" - by (simp add: expand_poly_eq diff_minus) + by (simp add: poly_eq_iff diff_minus) qed end lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) + by (rule poly_eqI, simp add: coeff_pCons split: nat.split) lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) + 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_ext, simp add: coeff_pCons split: nat.split) + 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) @@ -398,75 +725,133 @@ by (simp add: diff_minus degree_add_less) lemma add_monom: "monom a n + monom b n = monom (a + b) n" - by (rule poly_ext) simp + by (rule poly_eqI) simp lemma diff_monom: "monom a n - monom b n = monom (a - b) n" - by (rule poly_ext) simp + by (rule poly_eqI) simp lemma minus_monom: "- monom a n = monom (-a) n" - by (rule poly_ext) simp + by (rule poly_eqI) simp lemma coeff_setsum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" by (cases "finite A", induct set: finite, simp_all) lemma monom_setsum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" - by (rule poly_ext) (simp add: coeff_setsum) + by (rule poly_eqI) (simp add: coeff_setsum) + +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 - + { fix xs ys :: "'a list" and n + have "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys 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 } + note * = this + { fix xs ys :: "'a list" + assume "xs \ [] \ last xs \ 0" and "ys \ [] \ last ys \ 0" + moreover assume "plus_coeffs xs ys \ []" + ultimately have "last (plus_coeffs xs ys) \ 0" + proof (induct xs ys rule: plus_coeffs.induct) + case (3 x xs y ys) then show ?case by (auto simp add: cCons_def) metis + qed simp_all } + note ** = this + show ?thesis + apply (rule coeffs_eqI) + apply (simp add: * nth_default_coeffs_eq) + apply (rule **) + apply (auto dest: last_coeffs_not_0) + done +qed + +lemma coeffs_uminus [code abstract]: + "coeffs (- p) = map (\a. - a) (coeffs p)" + by (rule coeffs_eqI) + (simp_all add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) + +lemma [code]: + fixes p q :: "'a::ab_group_add poly" + shows "p - q = p + - q" + by simp + +lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" + apply (induct p arbitrary: q, simp) + apply (case_tac q, simp, simp add: algebra_simps) + done + +lemma poly_minus [simp]: + fixes x :: "'a::comm_ring" + shows "poly (- p) x = - poly p x" + by (induct p) simp_all + +lemma poly_diff [simp]: + fixes x :: "'a::comm_ring" + shows "poly (p - q) x = poly p x - poly q x" + by (simp add: diff_minus) + +lemma poly_setsum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" + by (induct A rule: infinite_finite_induct) simp_all -subsection {* Multiplication by a constant *} - -definition - smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" where - "smult a p = Abs_poly (\n. a * coeff p n)" +subsection {* Multiplication by a constant, polynomial multiplication and the unit polynomial *} -lemma Poly_smult: - fixes f :: "nat \ 'a::comm_semiring_0" - shows "f \ Poly \ (\n. a * f n) \ Poly" - unfolding Poly_def - by (clarify, rule_tac x=n in exI, simp) +lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" + is "\a p n. a * coeff p n" +proof (rule almost_everywhere_zeroI) + fix a :: 'a and p :: "'a poly" and i + assume "degree p < i" + then show "a * coeff p i = 0" + by (simp add: coeff_eq_0) +qed -lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" - unfolding smult_def - by (simp add: Abs_poly_inverse Poly_smult coeff) +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_ext, simp add: mult_assoc) + by (rule poly_eqI, simp add: mult_assoc) lemma smult_0_right [simp]: "smult a 0 = 0" - by (rule poly_ext, simp) + by (rule poly_eqI, simp) lemma smult_0_left [simp]: "smult 0 p = 0" - by (rule poly_ext, simp) + by (rule poly_eqI, simp) lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" - by (rule poly_ext, simp) + by (rule poly_eqI, simp) lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" - by (rule poly_ext, simp add: algebra_simps) + 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_ext, simp add: algebra_simps) + by (rule poly_eqI, simp add: algebra_simps) lemma smult_minus_right [simp]: "smult (a::'a::comm_ring) (- p) = - smult a p" - by (rule poly_ext, simp) + by (rule poly_eqI, simp) lemma smult_minus_left [simp]: "smult (- a::'a::comm_ring) p = - smult a p" - by (rule poly_ext, simp) + by (rule poly_eqI, simp) lemma smult_diff_right: "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" - by (rule poly_ext, simp add: algebra_simps) + by (rule poly_eqI, simp add: algebra_simps) lemma smult_diff_left: "smult (a - b::'a::comm_ring) p = smult a p - smult b p" - by (rule poly_ext, simp add: algebra_simps) + by (rule poly_eqI, simp add: algebra_simps) lemmas smult_distribs = smult_add_left smult_add_right @@ -474,7 +859,7 @@ lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) + 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 add: monom_0, simp add: monom_Suc) @@ -487,65 +872,48 @@ lemma smult_eq_0_iff [simp]: fixes a :: "'a::idom" shows "smult a p = 0 \ a = 0 \ p = 0" - by (simp add: expand_poly_eq) - - -subsection {* Multiplication of polynomials *} + by (simp add: poly_eq_iff) -(* TODO: move to Set_Interval.thy *) -lemma setsum_atMost_Suc_shift: - fixes f :: "nat \ 'a::comm_monoid_add" - shows "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" -proof (induct n) - case 0 show ?case by simp -next - case (Suc n) note IH = this - have "(\i\Suc (Suc n). f i) = (\i\Suc n. f i) + f (Suc (Suc n))" - by (rule setsum_atMost_Suc) - also have "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" - by (rule IH) - also have "f 0 + (\i\n. f (Suc i)) + f (Suc (Suc n)) = - f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" - by (rule add_assoc) - also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" - by (rule setsum_atMost_Suc [symmetric]) - finally show ?case . -qed +lemma coeffs_smult [code abstract]: + fixes p :: "'a::idom poly" + shows "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" + by (rule coeffs_eqI) + (auto simp add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) instantiation poly :: (comm_semiring_0) comm_semiring_0 begin definition - times_poly_def: - "p * q = poly_rec 0 (\a p pq. smult a q + pCons 0 pq) p" + "p * q = fold_coeffs (\a p. smult a q + pCons 0 p) p 0" lemma mult_poly_0_left: "(0::'a poly) * q = 0" - unfolding times_poly_def by (simp add: poly_rec_0) + by (simp add: times_poly_def) lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" - unfolding times_poly_def by (simp add: poly_rec_pCons) + 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 add: mult_poly_0_left, simp) + by (induct p) (simp add: mult_poly_0_left, simp) lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" - by (induct p, simp add: mult_poly_0_left, simp add: algebra_simps) + by (induct p) (simp add: mult_poly_0_left, simp add: 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 add: mult_poly_0, simp add: smult_add_right) +lemma mult_smult_left [simp]: + "smult a p * q = smult a (p * q)" + by (induct p) (simp add: mult_poly_0, simp add: smult_add_right) -lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" - by (induct q, simp add: mult_poly_0, simp add: smult_add_right) +lemma mult_smult_right [simp]: + "p * smult a q = smult a (p * q)" + by (induct q) (simp add: mult_poly_0, simp add: smult_add_right) lemma mult_poly_add_left: fixes p q r :: "'a poly" shows "(p + q) * r = p * r + q * r" - by (induct r, simp add: mult_poly_0, - simp add: smult_distribs algebra_simps) + by (induct r) (simp add: mult_poly_0, simp add: smult_distribs algebra_simps) instance proof fix p q r :: "'a poly" @@ -585,20 +953,15 @@ 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) - -subsection {* The unit polynomial and exponentiation *} - instantiation poly :: (comm_semiring_1) comm_semiring_1 begin -definition - one_poly_def: - "1 = pCons 1 0" +definition one_poly_def: + "1 = pCons 1 0" instance proof fix p :: "'a poly" show "1 * p = p" - unfolding one_poly_def - by simp + unfolding one_poly_def by simp next show "0 \ (1::'a poly)" unfolding one_poly_def by simp @@ -608,6 +971,10 @@ instance poly :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. +instance poly :: (comm_ring) comm_ring .. + +instance poly :: (comm_ring_1) comm_ring_1 .. + lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" unfolding one_poly_def by (simp add: coeff_pCons split: nat.split) @@ -616,7 +983,33 @@ unfolding one_poly_def by (rule degree_pCons_0) -text {* Lemmas about divisibility *} +lemma coeffs_1_eq [simp, code abstract]: + "coeffs 1 = [1]" + by (simp add: one_poly_def) + +lemma degree_power_le: + "degree (p ^ n) \ degree p * n" + by (induct n) (auto intro: order_trans degree_mult_le) + +lemma poly_smult [simp]: + "poly (smult a p) x = a * poly p x" + by (induct p, simp, simp add: algebra_simps) + +lemma poly_mult [simp]: + "poly (p * q) x = poly p x * poly q x" + by (induct p, simp_all, simp add: algebra_simps) + +lemma poly_1 [simp]: + "poly 1 x = 1" + by (simp add: one_poly_def) + +lemma poly_power [simp]: + fixes p :: "'a::{comm_semiring_1} poly" + shows "poly (p ^ n) x = poly p x ^ n" + by (induct n) simp_all + + +subsection {* Lemmas about divisibility *} lemma dvd_smult: "p dvd q \ p dvd smult a q" proof - @@ -655,13 +1048,6 @@ shows "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" by (auto elim: smult_dvd smult_dvd_cancel) -lemma degree_power_le: "degree (p ^ n) \ degree p * n" -by (induct n, simp, auto intro: order_trans degree_mult_le) - -instance poly :: (comm_ring) comm_ring .. - -instance poly :: (comm_ring_1) comm_ring_1 .. - subsection {* Polynomials form an integral domain *} @@ -680,7 +1066,7 @@ also have "coeff p (degree p) * coeff q (degree q) \ 0" using `p \ 0` and `q \ 0` by simp finally have "\n. coeff (p * q) n \ 0" .. - thus "p * q \ 0" by (simp add: expand_poly_eq) + thus "p * q \ 0" by (simp add: poly_eq_iff) qed lemma degree_mult_eq: @@ -698,8 +1084,7 @@ subsection {* Polynomials form an ordered integral domain *} -definition - pos_poly :: "'a::linordered_idom poly \ bool" +definition pos_poly :: "'a::linordered_idom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" @@ -725,6 +1110,20 @@ lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" by (induct p) (auto simp add: pos_poly_pCons) +lemma last_coeffs_eq_coeff_degree: + "p \ 0 \ last (coeffs p) = coeff p (degree p)" + by (simp add: coeffs_def) + +lemma pos_poly_coeffs [code]: + "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?P \ ?Q") +proof + assume ?Q then show ?P by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) +next + assume ?P then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) + then have "p \ 0" by auto + with * show ?Q by (simp add: last_coeffs_eq_coeff_degree) +qed + instantiation poly :: (linordered_idom) linordered_idom begin @@ -802,10 +1201,145 @@ text {* TODO: Simplification rules for comparisons *} +subsection {* Synthetic division and polynomial roots *} + +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" + unfolding synthetic_div_def by simp + +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, simp add: split_def) + +lemma synthetic_div_pCons [simp]: + "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" + unfolding synthetic_div_def + by (simp add: split_def snd_synthetic_divmod) + +lemma synthetic_div_eq_0_iff: + "synthetic_div p c = 0 \ degree p = 0" + by (induct p, simp, case_tac p, simp) + +lemma degree_synthetic_div: + "degree (synthetic_div p c) = degree p - 1" + by (induct p, simp, simp 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" +apply (induct p arbitrary: q r) +apply (simp, frule synthetic_div_unique_lemma, simp) +apply (case_tac q, force) +done + +lemma synthetic_div_correct': + fixes c :: "'a::comm_ring_1" + shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" + using synthetic_div_correct [of p c] + by (simp add: algebra_simps) + +lemma poly_eq_0_iff_dvd: + fixes c :: "'a::idom" + shows "poly p c = 0 \ [:-c, 1:] dvd p" +proof + assume "poly p c = 0" + with synthetic_div_correct' [of c p] + have "p = [:-c, 1:] * synthetic_div p c" by simp + then show "[:-c, 1:] dvd p" .. +next + assume "[:-c, 1:] dvd p" + then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) + then show "poly p c = 0" by simp +qed + +lemma dvd_iff_poly_eq_0: + fixes c :: "'a::idom" + shows "[:c, 1:] dvd p \ poly p (-c) = 0" + by (simp add: poly_eq_0_iff_dvd) + +lemma poly_roots_finite: + fixes p :: "'a::idom poly" + shows "p \ 0 \ finite {x. poly p x = 0}" +proof (induct n \ "degree p" arbitrary: p) + case (0 p) + 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 p) + 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 + then have "finite {x. poly k x = 0}" using `k \ 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 uminus_add_conv_diff Collect_disj_eq + del: mult_pCons_left) + qed +qed + +lemma poly_eq_poly_eq_iff: + fixes p q :: "'a::{idom,ring_char_0} poly" + shows "poly p = poly q \ p = q" (is "?P \ ?Q") +proof + assume ?Q then show ?P by simp +next + { fix p :: "'a::{idom,ring_char_0} poly" + have "poly p = poly 0 \ p = 0" + apply (cases "p = 0", simp_all) + apply (drule poly_roots_finite) + apply (auto simp add: infinite_UNIV_char_0) + done + } note this [of "p - q"] + moreover assume ?P + ultimately show ?Q by auto +qed + +lemma poly_all_0_iff_0: + fixes p :: "'a::{ring_char_0, idom} poly" + shows "(\x. poly p x = 0) \ p = 0" + by (auto simp add: poly_eq_poly_eq_iff [symmetric]) + + subsection {* Long division of polynomials *} -definition - pdivmod_rel :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" +definition pdivmod_rel :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" where "pdivmod_rel x y q r \ x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" @@ -1106,327 +1640,54 @@ apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) done - -subsection {* GCD of polynomials *} - -function - poly_gcd :: "'a::field poly \ 'a poly \ 'a poly" where - "poly_gcd x 0 = smult (inverse (coeff x (degree x))) x" -| "y \ 0 \ poly_gcd x y = poly_gcd y (x mod y)" -by auto - -termination poly_gcd -by (relation "measure (\(x, y). if y = 0 then 0 else Suc (degree y))") - (auto dest: degree_mod_less) - -declare poly_gcd.simps [simp del] - -lemma poly_gcd_dvd1 [iff]: "poly_gcd x y dvd x" - and poly_gcd_dvd2 [iff]: "poly_gcd x y dvd y" - apply (induct x y rule: poly_gcd.induct) - apply (simp_all add: poly_gcd.simps) - apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) - apply (blast dest: dvd_mod_imp_dvd) - done - -lemma poly_gcd_greatest: "k dvd x \ k dvd y \ k dvd poly_gcd x y" - by (induct x y rule: poly_gcd.induct) - (simp_all add: poly_gcd.simps dvd_mod dvd_smult) - -lemma dvd_poly_gcd_iff [iff]: - "k dvd poly_gcd x y \ k dvd x \ k dvd y" - by (blast intro!: poly_gcd_greatest intro: dvd_trans) +definition pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" +where + "pdivmod p q = (p div q, p mod q)" -lemma poly_gcd_monic: - "coeff (poly_gcd x y) (degree (poly_gcd x y)) = - (if x = 0 \ y = 0 then 0 else 1)" - by (induct x y rule: poly_gcd.induct) - (simp_all add: poly_gcd.simps nonzero_imp_inverse_nonzero) - -lemma poly_gcd_zero_iff [simp]: - "poly_gcd x y = 0 \ x = 0 \ y = 0" - by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) - -lemma poly_gcd_0_0 [simp]: "poly_gcd 0 0 = 0" - by simp +lemma div_poly_code [code]: + "p div q = fst (pdivmod p q)" + by (simp add: pdivmod_def) -lemma poly_dvd_antisym: - fixes p q :: "'a::idom poly" - assumes coeff: "coeff p (degree p) = coeff q (degree q)" - assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" -proof (cases "p = 0") - case True with coeff show "p = q" by simp -next - case False with coeff have "q \ 0" by auto - have degree: "degree p = degree q" - using `p dvd q` `q dvd p` `p \ 0` `q \ 0` - by (intro order_antisym dvd_imp_degree_le) - - from `p dvd q` obtain a where a: "q = p * a" .. - with `q \ 0` have "a \ 0" by auto - with degree a `p \ 0` have "degree a = 0" - by (simp add: degree_mult_eq) - with coeff a show "p = q" - by (cases a, auto split: if_splits) -qed +lemma mod_poly_code [code]: + "p mod q = snd (pdivmod p q)" + by (simp add: pdivmod_def) -lemma poly_gcd_unique: - assumes dvd1: "d dvd x" and dvd2: "d dvd y" - and greatest: "\k. k dvd x \ k dvd y \ k dvd d" - and monic: "coeff d (degree d) = (if x = 0 \ y = 0 then 0 else 1)" - shows "poly_gcd x y = d" -proof - - have "coeff (poly_gcd x y) (degree (poly_gcd x y)) = coeff d (degree d)" - by (simp_all add: poly_gcd_monic monic) - moreover have "poly_gcd x y dvd d" - using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) - moreover have "d dvd poly_gcd x y" - using dvd1 dvd2 by (rule poly_gcd_greatest) - ultimately show ?thesis - by (rule poly_dvd_antisym) -qed - -interpretation poly_gcd: abel_semigroup poly_gcd -proof - fix x y z :: "'a poly" - show "poly_gcd (poly_gcd x y) z = poly_gcd x (poly_gcd y z)" - by (rule poly_gcd_unique) (auto intro: dvd_trans simp add: poly_gcd_monic) - show "poly_gcd x y = poly_gcd y x" - by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) -qed - -lemmas poly_gcd_assoc = poly_gcd.assoc -lemmas poly_gcd_commute = poly_gcd.commute -lemmas poly_gcd_left_commute = poly_gcd.left_commute - -lemmas poly_gcd_ac = poly_gcd_assoc poly_gcd_commute poly_gcd_left_commute - -lemma poly_gcd_1_left [simp]: "poly_gcd 1 y = 1" -by (rule poly_gcd_unique) simp_all +lemma pdivmod_0: + "pdivmod 0 q = (0, 0)" + by (simp add: pdivmod_def) -lemma poly_gcd_1_right [simp]: "poly_gcd x 1 = 1" -by (rule poly_gcd_unique) simp_all - -lemma poly_gcd_minus_left [simp]: "poly_gcd (- x) y = poly_gcd x y" -by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) - -lemma poly_gcd_minus_right [simp]: "poly_gcd x (- y) = poly_gcd x y" -by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) - - -subsection {* Evaluation of polynomials *} - -definition - poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" where - "poly = poly_rec (\x. 0) (\a p f x. a + x * f x)" - -lemma poly_0 [simp]: "poly 0 x = 0" - unfolding poly_def by (simp add: poly_rec_0) - -lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" - unfolding poly_def by (simp add: poly_rec_pCons) - -lemma poly_1 [simp]: "poly 1 x = 1" - unfolding one_poly_def by simp - -lemma poly_monom: - fixes a x :: "'a::{comm_semiring_1}" - shows "poly (monom a n) x = a * x ^ n" - by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac) - -lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" - apply (induct p arbitrary: q, simp) - apply (case_tac q, simp, simp add: algebra_simps) +lemma pdivmod_pCons: + "pdivmod (pCons a p) q = + (if q = 0 then (0, pCons a p) else + (let (s, r) = pdivmod p q; + b = coeff (pCons a r) (degree q) / coeff q (degree q) + in (pCons b s, pCons a r - smult b q)))" + apply (simp add: pdivmod_def Let_def, safe) + apply (rule div_poly_eq) + apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) + apply (rule mod_poly_eq) + apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) done -lemma poly_minus [simp]: - fixes x :: "'a::comm_ring" - shows "poly (- p) x = - poly p x" - by (induct p, simp_all) - -lemma poly_diff [simp]: - fixes x :: "'a::comm_ring" - shows "poly (p - q) x = poly p x - poly q x" - by (simp add: diff_minus) - -lemma poly_setsum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" - by (cases "finite A", induct set: finite, simp_all) - -lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" - by (induct p, simp, simp add: algebra_simps) - -lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" - by (induct p, simp_all, simp add: algebra_simps) - -lemma poly_power [simp]: - fixes p :: "'a::{comm_semiring_1} poly" - shows "poly (p ^ n) x = poly p x ^ n" - by (induct n, simp, simp add: power_Suc) - - -subsection {* 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 = - poly_rec (0, 0) (\a p (q, r). (pCons r q, a + c * r)) p" - -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)" - unfolding synthetic_divmod_def - by (simp add: poly_rec_0) - -lemma synthetic_divmod_pCons [simp]: - "synthetic_divmod (pCons a p) c = - (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" - unfolding synthetic_divmod_def - by (simp add: poly_rec_pCons) - -lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" - by (induct p, simp, simp add: split_def) - -lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" - unfolding synthetic_div_def by simp - -lemma synthetic_div_pCons [simp]: - "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" - unfolding synthetic_div_def - by (simp add: split_def snd_synthetic_divmod) - -lemma synthetic_div_eq_0_iff: - "synthetic_div p c = 0 \ degree p = 0" - by (induct p, simp, case_tac p, simp) - -lemma degree_synthetic_div: - "degree (synthetic_div p c) = degree p - 1" - by (induct p, simp, simp 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_lemma: "smult c p = pCons a p \ p = 0" -by (induct p arbitrary: a) simp_all - -lemma synthetic_div_unique: - "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" -apply (induct p arbitrary: q r) -apply (simp, frule synthetic_div_unique_lemma, simp) -apply (case_tac q, force) -done - -lemma synthetic_div_correct': - fixes c :: "'a::comm_ring_1" - shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" - using synthetic_div_correct [of p c] - by (simp add: algebra_simps) - -lemma poly_eq_0_iff_dvd: - fixes c :: "'a::idom" - shows "poly p c = 0 \ [:-c, 1:] dvd p" -proof - assume "poly p c = 0" - with synthetic_div_correct' [of c p] - have "p = [:-c, 1:] * synthetic_div p c" by simp - then show "[:-c, 1:] dvd p" .. -next - assume "[:-c, 1:] dvd p" - then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) - then show "poly p c = 0" by simp -qed - -lemma dvd_iff_poly_eq_0: - fixes c :: "'a::idom" - shows "[:c, 1:] dvd p \ poly p (-c) = 0" - by (simp add: poly_eq_0_iff_dvd) - -lemma poly_roots_finite: - fixes p :: "'a::idom poly" - shows "p \ 0 \ finite {x. poly p x = 0}" -proof (induct n \ "degree p" arbitrary: p) - case (0 p) - 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 p) - 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 - then have "finite {x. poly k x = 0}" using `k \ 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 uminus_add_conv_diff Collect_disj_eq - del: mult_pCons_left) - qed -qed - -lemma poly_zero: - fixes p :: "'a::{idom,ring_char_0} poly" - shows "poly p = poly 0 \ p = 0" -apply (cases "p = 0", simp_all) -apply (drule poly_roots_finite) -apply (auto simp add: infinite_UNIV_char_0) -done - -lemma poly_eq_iff: - fixes p q :: "'a::{idom,ring_char_0} poly" - shows "poly p = poly q \ p = q" - using poly_zero [of "p - q"] - by (simp add: fun_eq_iff) - - -subsection {* Composition of polynomials *} - -definition - pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" -where - "pcompose p q = poly_rec 0 (\a _ c. [:a:] + q * c) p" - -lemma pcompose_0 [simp]: "pcompose 0 q = 0" - unfolding pcompose_def by (simp add: poly_rec_0) - -lemma pcompose_pCons: - "pcompose (pCons a p) q = [:a:] + q * pcompose p q" - unfolding pcompose_def by (simp add: poly_rec_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" -apply (induct p, simp) -apply (simp add: pcompose_pCons, clarify) -apply (rule degree_add_le, simp) -apply (rule order_trans [OF degree_mult_le], simp) -done +lemma pdivmod_fold_coeffs [code]: + "pdivmod p 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))" + apply (cases "q = 0") + apply (simp add: pdivmod_def) + apply (rule sym) + apply (induct p) + apply (simp_all add: pdivmod_0 pdivmod_pCons) + apply (case_tac "a = 0 \ p = 0") + apply (auto simp add: pdivmod_def) + done subsection {* Order of polynomial roots *} -definition - order :: "'a::idom \ 'a poly \ nat" +definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" @@ -1490,107 +1751,161 @@ done -subsection {* Configuration of the code generator *} - -code_datatype "0::'a::zero poly" pCons +subsection {* GCD of polynomials *} -quickcheck_generator poly constructors: "0::'a::zero poly", pCons - -instantiation poly :: ("{zero, equal}") equal +instantiation poly :: (field) gcd begin -definition - "HOL.equal (p::'a poly) q \ p = q" +function gcd_poly :: "'a::field poly \ 'a poly \ 'a poly" +where + "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x" +| "y \ 0 \ gcd (x::'a poly) y = gcd y (x mod y)" +by auto -instance proof -qed (rule equal_poly_def) +termination "gcd :: _ poly \ _" +by (relation "measure (\(x, y). if y = 0 then 0 else Suc (degree y))") + (auto dest: degree_mod_less) + +declare gcd_poly.simps [simp del] + +instance .. end -lemma eq_poly_code [code]: - "HOL.equal (0::_ poly) (0::_ poly) \ True" - "HOL.equal (0::_ poly) (pCons b q) \ HOL.equal 0 b \ HOL.equal 0 q" - "HOL.equal (pCons a p) (0::_ poly) \ HOL.equal a 0 \ HOL.equal p 0" - "HOL.equal (pCons a p) (pCons b q) \ HOL.equal a b \ HOL.equal p q" - by (simp_all add: equal) +lemma + fixes x y :: "_ poly" + shows poly_gcd_dvd1 [iff]: "gcd x y dvd x" + and poly_gcd_dvd2 [iff]: "gcd x y dvd y" + apply (induct x y rule: gcd_poly.induct) + apply (simp_all add: gcd_poly.simps) + apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) + apply (blast dest: dvd_mod_imp_dvd) + done -lemma [code nbe]: - "HOL.equal (p :: _ poly) p \ True" - by (fact equal_refl) +lemma poly_gcd_greatest: + fixes k x y :: "_ poly" + shows "k dvd x \ k dvd y \ k dvd gcd x y" + by (induct x y rule: gcd_poly.induct) + (simp_all add: gcd_poly.simps dvd_mod dvd_smult) -lemmas coeff_code [code] = - coeff_0 coeff_pCons_0 coeff_pCons_Suc +lemma dvd_poly_gcd_iff [iff]: + fixes k x y :: "_ poly" + shows "k dvd gcd x y \ k dvd x \ k dvd y" + by (blast intro!: poly_gcd_greatest intro: dvd_trans) -lemmas degree_code [code] = - degree_0 degree_pCons_eq_if +lemma poly_gcd_monic: + fixes x y :: "_ poly" + shows "coeff (gcd x y) (degree (gcd x y)) = + (if x = 0 \ y = 0 then 0 else 1)" + by (induct x y rule: gcd_poly.induct) + (simp_all add: gcd_poly.simps nonzero_imp_inverse_nonzero) -lemmas monom_poly_code [code] = - monom_0 monom_Suc +lemma poly_gcd_zero_iff [simp]: + fixes x y :: "_ poly" + shows "gcd x y = 0 \ x = 0 \ y = 0" + by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) -lemma add_poly_code [code]: - "0 + q = (q :: _ poly)" - "p + 0 = (p :: _ poly)" - "pCons a p + pCons b q = pCons (a + b) (p + q)" -by simp_all +lemma poly_gcd_0_0 [simp]: + "gcd (0::_ poly) 0 = 0" + by simp -lemma minus_poly_code [code]: - "- 0 = (0 :: _ poly)" - "- pCons a p = pCons (- a) (- p)" -by simp_all +lemma poly_dvd_antisym: + fixes p q :: "'a::idom poly" + assumes coeff: "coeff p (degree p) = coeff q (degree q)" + assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" +proof (cases "p = 0") + case True with coeff show "p = q" by simp +next + case False with coeff have "q \ 0" by auto + have degree: "degree p = degree q" + using `p dvd q` `q dvd p` `p \ 0` `q \ 0` + by (intro order_antisym dvd_imp_degree_le) -lemma diff_poly_code [code]: - "0 - q = (- q :: _ poly)" - "p - 0 = (p :: _ poly)" - "pCons a p - pCons b q = pCons (a - b) (p - q)" -by simp_all + from `p dvd q` obtain a where a: "q = p * a" .. + with `q \ 0` have "a \ 0" by auto + with degree a `p \ 0` have "degree a = 0" + by (simp add: degree_mult_eq) + with coeff a show "p = q" + by (cases a, auto split: if_splits) +qed -lemmas smult_poly_code [code] = - smult_0_right smult_pCons - -lemma mult_poly_code [code]: - "0 * q = (0 :: _ poly)" - "pCons a p * q = smult a q + pCons 0 (p * q)" -by simp_all +lemma poly_gcd_unique: + fixes d x y :: "_ poly" + assumes dvd1: "d dvd x" and dvd2: "d dvd y" + and greatest: "\k. k dvd x \ k dvd y \ k dvd d" + and monic: "coeff d (degree d) = (if x = 0 \ y = 0 then 0 else 1)" + shows "gcd x y = d" +proof - + have "coeff (gcd x y) (degree (gcd x y)) = coeff d (degree d)" + by (simp_all add: poly_gcd_monic monic) + moreover have "gcd x y dvd d" + using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) + moreover have "d dvd gcd x y" + using dvd1 dvd2 by (rule poly_gcd_greatest) + ultimately show ?thesis + by (rule poly_dvd_antisym) +qed -lemmas poly_code [code] = - poly_0 poly_pCons - -lemmas synthetic_divmod_code [code] = - synthetic_divmod_0 synthetic_divmod_pCons +interpretation gcd_poly!: abel_semigroup "gcd :: _ poly \ _" +proof + fix x y z :: "'a poly" + show "gcd (gcd x y) z = gcd x (gcd y z)" + by (rule poly_gcd_unique) (auto intro: dvd_trans simp add: poly_gcd_monic) + show "gcd x y = gcd y x" + by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) +qed -text {* code generator setup for div and mod *} +lemmas poly_gcd_assoc = gcd_poly.assoc +lemmas poly_gcd_commute = gcd_poly.commute +lemmas poly_gcd_left_commute = gcd_poly.left_commute -definition - pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" -where - "pdivmod x y = (x div y, x mod y)" +lemmas poly_gcd_ac = poly_gcd_assoc poly_gcd_commute poly_gcd_left_commute + +lemma poly_gcd_1_left [simp]: "gcd 1 y = (1 :: _ poly)" +by (rule poly_gcd_unique) simp_all -lemma div_poly_code [code]: "x div y = fst (pdivmod x y)" - unfolding pdivmod_def by simp +lemma poly_gcd_1_right [simp]: "gcd x 1 = (1 :: _ poly)" +by (rule poly_gcd_unique) simp_all + +lemma poly_gcd_minus_left [simp]: "gcd (- x) y = gcd x (y :: _ poly)" +by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) -lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)" - unfolding pdivmod_def by simp +lemma poly_gcd_minus_right [simp]: "gcd x (- y) = gcd x (y :: _ poly)" +by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) -lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)" - unfolding pdivmod_def by simp +lemma poly_gcd_code [code]: + "gcd x y = (if y = 0 then smult (inverse (coeff x (degree x))) x else gcd y (x mod (y :: _ poly)))" + by (simp add: gcd_poly.simps) + + +subsection {* Composition of polynomials *} -lemma pdivmod_pCons [code]: - "pdivmod (pCons a x) y = - (if y = 0 then (0, pCons a x) else - (let (q, r) = pdivmod x y; - b = coeff (pCons a r) (degree y) / coeff y (degree y) - in (pCons b q, pCons a r - smult b y)))" -apply (simp add: pdivmod_def Let_def, safe) -apply (rule div_poly_eq) -apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) -apply (rule mod_poly_eq) -apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) +definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" +where + "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" + +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 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" +apply (induct p, simp) +apply (simp add: pcompose_pCons, clarify) +apply (rule degree_add_le, simp) +apply (rule order_trans [OF degree_mult_le], simp) done -lemma poly_gcd_code [code]: - "poly_gcd x y = - (if y = 0 then smult (inverse (coeff x (degree x))) x - else poly_gcd y (x mod y))" - by (simp add: poly_gcd.simps) + +no_notation cCons (infixr "##" 65) end + diff -r 7f864f2219a9 -r 3cc46b8cca5e src/HOL/List.thy --- a/src/HOL/List.thy Sat Jun 15 17:19:23 2013 +0200 +++ b/src/HOL/List.thy Sat Jun 15 17:19:23 2013 +0200 @@ -2308,6 +2308,14 @@ ==> dropWhile P l = dropWhile Q k" by (induct k arbitrary: l, simp_all) +lemma takeWhile_idem [simp]: + "takeWhile P (takeWhile P xs) = takeWhile P xs" + by (induct xs) auto + +lemma dropWhile_idem [simp]: + "dropWhile P (dropWhile P xs) = dropWhile P xs" + by (induct xs) auto + subsubsection {* @{const zip} *} @@ -2947,6 +2955,10 @@ apply (auto simp add: less_diff_conv) done +lemma map_decr_upt: + "map (\n. n - Suc 0) [Suc m.. k <= length ys ==> (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys" @@ -3703,7 +3715,6 @@ apply clarsimp done - lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}" by (induct n) auto @@ -3819,6 +3830,22 @@ then show ?thesis by blast qed +lemma Cons_replicate_eq: + "x # xs = replicate n y \ x = y \ n > 0 \ xs = replicate (n - 1) x" + by (induct n) auto + +lemma replicate_length_same: + "(\y\set xs. y = x) \ replicate (length xs) x = xs" + by (induct xs) simp_all + +lemma foldr_replicate [simp]: + "foldr f (replicate n x) = f x ^^ n" + by (induct n) (simp_all) + +lemma fold_replicate [simp]: + "fold f (replicate n x) = f x ^^ n" + by (subst foldr_fold [symmetric]) simp_all + subsubsection {* @{const enumerate} *} diff -r 7f864f2219a9 -r 3cc46b8cca5e src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Sat Jun 15 17:19:23 2013 +0200 +++ b/src/HOL/Set_Interval.thy Sat Jun 15 17:19:23 2013 +0200 @@ -1438,6 +1438,26 @@ apply(simp add:setsum_head_upt_Suc) done +lemma setsum_atMost_Suc_shift: + fixes f :: "nat \ 'a::comm_monoid_add" + shows "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) note IH = this + have "(\i\Suc (Suc n). f i) = (\i\Suc n. f i) + f (Suc (Suc n))" + by (rule setsum_atMost_Suc) + also have "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" + by (rule IH) + also have "f 0 + (\i\n. f (Suc i)) + f (Suc (Suc n)) = + f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" + by (rule add_assoc) + also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" + by (rule setsum_atMost_Suc [symmetric]) + finally show ?case . +qed + + subsection {* The formula for geometric sums *} lemma geometric_sum: