# HG changeset patch # User haftmann # Date 1664181713 0 # Node ID 14dd8b46307f45e73262a3a6a13306f636d7b97f # Parent 8fcbce9f317cc22c443214ea1f88c4c7177e6cd3 streamlined division on polynomials diff -r 8fcbce9f317c -r 14dd8b46307f src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy Sun Sep 25 19:10:43 2022 +0000 +++ b/src/HOL/Computational_Algebra/Polynomial.thy Mon Sep 26 08:41:53 2022 +0000 @@ -3860,99 +3860,6 @@ by (auto simp add: mod_poly_def intro: *) qed -inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" - where - eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" - | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" - | eucl_rel_poly_remainderI: - "y \ 0 \ degree r < degree y \ x = q * y + r \ eucl_rel_poly x y (q, r)" - -lemma eucl_rel_poly_iff: - "eucl_rel_poly x y (q, r) \ - x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" - by (auto elim: eucl_rel_poly.cases - intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) - -lemma eucl_rel_poly: - "eucl_rel_poly x y (x div y, x mod y)" - using degree_mod_less_degree [of y x] - by (auto simp add: eucl_rel_poly_iff mod_eq_0_iff_dvd) - -lemma - assumes 1: "eucl_rel_poly x y (q1, r1)" - assumes 2: "eucl_rel_poly x y (q2, r2)" - shows eucl_rel_poly_unique_div: "q1 = q2" (is ?Q) - and eucl_rel_poly_unique_mod: "r1 = r2" (is ?R) -proof - - have \?Q \ ?R\ - proof (cases "y = 0") - assume "y = 0" - with assms show ?thesis - by (simp add: eucl_rel_poly_iff) - next - assume [simp]: "y \ 0" - from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" - unfolding eucl_rel_poly_iff by simp_all - from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" - unfolding eucl_rel_poly_iff by simp_all - from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" - by (simp add: algebra_simps) - from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" - by (auto intro: degree_diff_less) - show "q1 = q2 \ r1 = r2" - proof (rule classical) - assume "\ ?thesis" - with q3 have "q1 \ q2" and "r1 \ r2" by auto - with r3 have "degree (r2 - r1) < degree y" by simp - also have "degree y \ degree (q1 - q2) + degree y" by simp - also from \q1 \ q2\ have "\ = degree ((q1 - q2) * y)" - by (simp add: degree_mult_eq) - also from q3 have "\ = degree (r2 - r1)" - by simp - finally have "degree (r2 - r1) < degree (r2 - r1)" . - then show ?thesis by simp - qed - qed - then show ?Q and ?R - by simp_all -qed - -lemma pdivmod_pdivmodrel: - \eucl_rel_poly x y (q, r) \ (x div y, x mod y) = (q, r)\ - using eucl_rel_poly_unique_div [of x y _ _ q r, OF eucl_rel_poly] - eucl_rel_poly_unique_mod [of x y _ _ q r, OF eucl_rel_poly] - eucl_rel_poly [of x y] - by auto - -lemma div_poly_eq: "eucl_rel_poly x y (q, r) \ x div y = q" - for x :: "'a::field poly" - by (simp add: pdivmod_pdivmodrel) - -lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \ x mod y = r" - for x :: "'a::field poly" - by (simp add: pdivmod_pdivmodrel) - -lemma div_poly_less: - fixes x :: "'a::field poly" - assumes "degree x < degree y" - shows "x div y = 0" -proof - - from assms have "eucl_rel_poly x y (0, x)" - by (simp add: eucl_rel_poly_iff) - then show "x div y = 0" - by (rule div_poly_eq) -qed - -lemma mod_poly_less: - assumes "degree x < degree y" - shows "x mod y = x" -proof - - from assms have "eucl_rel_poly x y (0, x)" - by (simp add: eucl_rel_poly_iff) - then show "x mod y = x" - by (rule mod_poly_eq) -qed - instantiation poly :: (field) unique_euclidean_ring begin @@ -3974,12 +3881,36 @@ with \p \ 0\ \euclidean_size r < euclidean_size p\ have \degree r < degree p\ by (simp add: euclidean_size_poly_def) - show \(q * p + r) div p = q\ - by (rule div_poly_eq) - (use \degree r < degree p\ in \auto simp add: eucl_rel_poly_iff\) + 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 div_mult_mod_eq div_poly_less degree_mult_eq power_add - intro!: degree_mod_less_degree split: if_splits) +qed (auto simp: euclidean_size_poly_def degree_mult_eq power_add intro: degree_mod_less_degree) end @@ -3995,74 +3926,53 @@ \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: - fixes x y::"'a::field poly" - assumes "degree x\0" "degree y\0" - shows "degree (x div y) < degree x" -proof - - have "x\0" "y\0" using assms by auto - define q r where "q=x div y" and "r=x mod y" - have *:"eucl_rel_poly x y (q, r)" unfolding q_def r_def - by (simp add: eucl_rel_poly) - then have "r = 0 \ degree r < degree y" using \y\0\ unfolding eucl_rel_poly_iff by auto - moreover have ?thesis when "r=0" - proof - - have "x = q * y" using * that unfolding eucl_rel_poly_iff by auto - then have "q\0" using \x\0\ \y\0\ by auto - from degree_mult_eq[OF this \y\0\] \x = q * y\ - have "degree x = degree q +degree y" by auto - then show ?thesis unfolding q_def using assms by auto - qed - moreover have ?thesis when "degree rdegree x") + \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 have "q=0" unfolding q_def using div_poly_less by auto - then show ?thesis unfolding q_def using assms(1) by auto + 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 - then have "degree x>degree r" using that by auto - then have "degree x = degree (x-r)" using degree_add_eq_right[of "-r" x] by auto - have "x-r = q*y" using * unfolding eucl_rel_poly_iff by auto - then have "q\0" using \degree r < degree x\ by auto - have "degree x = degree q +degree y" - using degree_mult_eq[OF \q\0\ \y\0\] \x-r = q*y\ \degree x = degree (x-r)\ by auto - then show ?thesis unfolding q_def using assms by auto + 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 - ultimately show ?thesis by auto -qed - -lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" - using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp +qed lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" - using degree_mod_less [of b a] by auto - -lemma eucl_rel_poly_pCons: - assumes rel: "eucl_rel_poly x y (q, r)" - assumes y: "y \ 0" - assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" - shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" - (is "eucl_rel_poly ?x y (?q, ?r)") -proof - - from assms have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" - by (simp_all add: eucl_rel_poly_iff) - from b x have "?x = ?q * y + ?r" by simp - moreover - have "?r = 0 \ degree ?r < degree y" - proof (rule eq_zero_or_degree_less) - show "degree ?r \ degree y" - proof (rule degree_diff_le) - from r show "degree (pCons a r) \ degree y" - by auto - show "degree (smult b y) \ degree y" - by (rule degree_smult_le) - qed - from \y \ 0\ show "coeff ?r (degree y) = 0" - by (simp add: b) - qed - ultimately show ?thesis - unfolding eucl_rel_poly_iff using \y \ 0\ by simp -qed + 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) @@ -4191,19 +4101,107 @@ 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))" - using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] - by (auto intro: div_poly_eq) - -lemma mod_pCons_eq: - "pCons a p mod q = - (if q = 0 then pCons a p - else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" - using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] - by (auto intro: mod_poly_eq) + \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) + 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) =