diff --git a/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy b/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy --- a/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy +++ b/thys/Algebraic_Numbers/Algebraic_Numbers_Prelim.thy @@ -1,1364 +1,1364 @@ (* Author: RenĂ© Thiemann Sebastiaan Joosten Akihisa Yamada License: BSD *) section \Algebraic Numbers -- Excluding Addition and Multiplication\ text \This theory contains basic definition and results on algebraic numbers, namely that algebraic numbers are closed under negation, inversion, $n$-th roots, and that every rational number is algebraic. For all of these closure properties, corresponding polynomial witnesses are available. Moreover, this theory contains the uniqueness result, that for every algebraic number there is exactly one content-free irreducible polynomial with positive leading coefficient for it. This result is stronger than similar ones which you find in many textbooks. The reason is that here we do not require a least degree construction. This is essential, since given some content-free irreducible polynomial for x, how should we check whether the degree is optimal. In the formalized result, this is not required. The result is proven via GCDs, and that the GCD does not change when executed on the rational numbers or on the reals or complex numbers, and that the GCD of a rational polynomial can be expressed via the GCD of integer polynomials.\ text \Many results are taken from the textbook \cite[pages 317ff]{AlgNumbers}.\ theory Algebraic_Numbers_Prelim imports "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Polynomial_Factorization.Rational_Factorization Berlekamp_Zassenhaus.Factorize_Int_Poly Berlekamp_Zassenhaus.Unique_Factorization_Poly begin text \For algebraic numbers, it turned out that @{const gcd_int_poly} is not preferable to the default implementation of @{const gcd}, which just implements Collin's primitive remainder sequence.\ declare gcd_int_poly_code[code_unfold del] lemma primitive_imp_unit_iff: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes pr: "primitive p" shows "p dvd 1 \ degree p = 0" proof assume "degree p = 0" from degree0_coeffs[OF this] obtain p0 where p: "p = [:p0:]" by auto then have "\c \ set (coeffs p). p0 dvd c" by (simp add: cCons_def) with pr have "p0 dvd 1" by (auto dest: primitiveD) with p show "p dvd 1" by auto next assume "p dvd 1" then show "degree p = 0" by (auto simp: poly_dvd_1) qed lemma dvd_all_coeffs_imp_dvd: assumes "\a \ set (coeffs p). c dvd a" shows "[:c:] dvd p" proof (insert assms, induct p) case 0 then show ?case by simp next case (pCons a p) have "pCons a p = [:a:] + pCons 0 p" by simp also have "[:c:] dvd ..." proof (rule dvd_add) from pCons show "[:c:] dvd [:a:]" by (auto simp: cCons_def) from pCons have "[:c:] dvd p" by auto from Rings.dvd_mult[OF this] show "[:c:] dvd pCons 0 p" by (subst pCons_0_as_mult) qed finally show ?case. qed lemma irreducible_content: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "irreducible p" shows "degree p = 0 \ primitive p" proof(rule ccontr) assume not: "\?thesis" then obtain c where c1: "\c dvd 1" and "\a \ set (coeffs p). c dvd a" by (auto elim: not_primitiveE) from dvd_all_coeffs_imp_dvd[OF this(2)] obtain r where p: "p = r * [:c:]" by (elim dvdE, auto) from irreducibleD[OF assms this] have "r dvd 1 \ [:c:] dvd 1" by auto with c1 have "r dvd 1" unfolding const_poly_dvd_1 by auto then have "degree r = 0" unfolding poly_dvd_1 by auto with p have "degree p = 0" by auto with not show False by auto qed (* TODO: move *) lemma linear_irreducible_field: fixes p :: "'a :: field poly" assumes deg: "degree p = 1" shows "irreducible p" proof (intro irreducibleI) from deg show p0: "p \ 0" by auto from deg show "\ p dvd 1" by (auto simp: poly_dvd_1) fix a b assume p: "p = a * b" with p0 have a0: "a \ 0" and b0: "b \ 0" by auto from degree_mult_eq[OF this, folded p] assms consider "degree a = 1" "degree b = 0" | "degree a = 0" "degree b = 1" by force then show "a dvd 1 \ b dvd 1" by (cases; insert a0 b0, auto simp: primitive_imp_unit_iff) qed (* TODO: move *) lemma linear_irreducible_int: fixes p :: "int poly" assumes deg: "degree p = 1" and cp: "content p dvd 1" shows "irreducible p" proof (intro irreducibleI) from deg show p0: "p \ 0" by auto from deg show "\ p dvd 1" by (auto simp: poly_dvd_1) fix a b assume p: "p = a * b" note * = cp[unfolded p is_unit_content_iff, unfolded content_mult] have a1: "content a dvd 1" and b1: "content b dvd 1" using content_ge_0_int[of a] pos_zmult_eq_1_iff_lemma[OF *] * by (auto simp: abs_mult) with p0 have a0: "a \ 0" and b0: "b \ 0" by auto from degree_mult_eq[OF this, folded p] assms consider "degree a = 1" "degree b = 0" | "degree a = 0" "degree b = 1" by force then show "a dvd 1 \ b dvd 1" by (cases; insert a1 b1, auto simp: primitive_imp_unit_iff) qed lemma irreducible_connect_rev: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes irr: "irreducible p" and deg: "degree p > 0" shows "irreducible\<^sub>d p" proof(intro irreducible\<^sub>dI deg) fix q r assume degq: "degree q > 0" and diff: "degree q < degree p" and p: "p = q * r" from degq have nu: "\ q dvd 1" by (auto simp: poly_dvd_1) from irreducibleD[OF irr p] nu have "r dvd 1" by auto then have "degree r = 0" by (auto simp: poly_dvd_1) with degq diff show False unfolding p using degree_mult_le[of q r] by auto qed subsection \Polynomial Evaluation of Integer and Rational Polynomials in Fields.\ abbreviation ipoly where "ipoly f x \ poly (of_int_poly f) x" lemma poly_map_poly_code[code_unfold]: "poly (map_poly h p) x = fold_coeffs (\ a b. h a + x * b) p 0" by (induct p, auto) abbreviation real_of_int_poly :: "int poly \ real poly" where "real_of_int_poly \ of_int_poly" abbreviation real_of_rat_poly :: "rat poly \ real poly" where "real_of_rat_poly \ map_poly of_rat" lemma of_rat_of_int[simp]: "of_rat \ of_int = of_int" by auto lemma ipoly_of_rat[simp]: "ipoly p (of_rat y) = of_rat (ipoly p y)" proof- have id: "of_int = of_rat o of_int" unfolding comp_def by auto show ?thesis by (subst id, subst map_poly_map_poly[symmetric], auto) qed lemma ipoly_of_real[simp]: "ipoly p (of_real x :: 'a :: {field,real_algebra_1}) = of_real (ipoly p x)" proof - have id: "of_int = of_real o of_int" unfolding comp_def by auto show ?thesis by (subst id, subst map_poly_map_poly[symmetric], auto) qed lemma finite_ipoly_roots: assumes "p \ 0" shows "finite {x :: real. ipoly p x = 0}" proof - let ?p = "real_of_int_poly p" from assms have "?p \ 0" by auto thus ?thesis by (rule poly_roots_finite) qed subsection \Algebraic Numbers -- Definition, Inverse, and Roots\ text \A number @{term "x :: 'a :: field"} is algebraic iff it is the root of an integer polynomial. Whereas the Isabelle distribution this is defined via the embedding of integers in an field via @{const Ints}, we work with integer polynomials of type @{type int} and then use @{const ipoly} for evaluating the polynomial at a real or complex point.\ lemma algebraic_altdef_ipoly: shows "algebraic x \ (\p. ipoly p x = 0 \ p \ 0)" unfolding algebraic_def proof (safe, goal_cases) case (1 p) 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 \ \" using of_int_the_int[OF that] by 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 1 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 ?case proof (intro exI conjI notI) from 1 show "ipoly p' x = 0" by (simp add: p_p') next assume "p' = 0" hence "p = 0" by (simp add: p_p' [symmetric]) with \p \ 0\ show False by contradiction qed next case (2 p) thus ?case by (intro exI[of _ "map_poly of_int p"], auto) qed text \Definition of being algebraic with explicit witness polynomial.\ definition represents :: "int poly \ 'a :: field_char_0 \ bool" (infix "represents" 51) where "p represents x = (ipoly p x = 0 \ p \ 0)" lemma representsI[intro]: "ipoly p x = 0 \ p \ 0 \ p represents x" unfolding represents_def by auto lemma representsD: assumes "p represents x" shows "p \ 0" and "ipoly p x = 0" using assms unfolding represents_def by auto lemma representsE: assumes "p represents x" and "p \ 0 \ ipoly p x = 0 \ thesis" shows thesis using assms unfolding represents_def by auto lemma represents_imp_degree: fixes x :: "'a :: field_char_0" assumes "p represents x" shows "degree p \ 0" proof- from assms have "p \ 0" and px: "ipoly p x = 0" by (auto dest:representsD) then have "(of_int_poly p :: 'a poly) \ 0" by auto then have "degree (of_int_poly p :: 'a poly) \ 0" by (fold poly_zero[OF px]) then show ?thesis by auto qed lemma representsE_full[elim]: assumes rep: "p represents x" and main: "p \ 0 \ ipoly p x = 0 \ degree p \ 0 \ thesis" shows thesis by (rule main, insert represents_imp_degree[OF rep] rep, auto elim: representsE) lemma represents_of_rat[simp]: "p represents (of_rat x) = p represents x" by (auto elim!:representsE) lemma represents_of_real[simp]: "p represents (of_real x) = p represents x" by (auto elim!:representsE) lemma algebraic_iff_represents: "algebraic x \ (\ p. p represents x)" unfolding algebraic_altdef_ipoly represents_def .. lemma represents_irr_non_0: assumes irr: "irreducible p" and ap: "p represents x" and x0: "x \ 0" shows "poly p 0 \ 0" proof have nu: "\ [:0,1::int:] dvd 1" by (auto simp: poly_dvd_1) assume "poly p 0 = 0" hence dvd: "[: 0, 1 :] dvd p" by (unfold dvd_iff_poly_eq_0, simp) then obtain q where pq: "p = [:0,1:] * q" by (elim dvdE) from irreducibleD[OF irr this] nu have "q dvd 1" by auto from this obtain r where "q = [:r:]" "r dvd 1" by (auto simp add: poly_dvd_1 dest: degree0_coeffs) with pq have "p = [:0,r:]" by auto with ap have "x = 0" by (auto simp: of_int_hom.map_poly_pCons_hom) with x0 show False by auto qed text \The polynomial encoding a rational number.\ definition poly_rat :: "rat \ int poly" where "poly_rat x = (case quotient_of x of (n,d) \ [:-n,d:])" definition abs_int_poly:: "int poly \ int poly" where "abs_int_poly p \ if lead_coeff p < 0 then -p else p" lemma pos_poly_abs_poly[simp]: shows "lead_coeff (abs_int_poly p) > 0 \ p \ 0" proof- have "p \ 0 \ lead_coeff p * sgn (lead_coeff p) > 0" by (fold abs_sgn, auto) then show ?thesis by (auto simp: abs_int_poly_def mult.commute) qed lemma abs_int_poly_0[simp]: "abs_int_poly 0 = 0" by (auto simp: abs_int_poly_def) lemma abs_int_poly_eq_0_iff[simp]: "abs_int_poly p = 0 \ p = 0" by (auto simp: abs_int_poly_def sgn_eq_0_iff) lemma degree_abs_int_poly[simp]: "degree (abs_int_poly p) = degree p" by (auto simp: abs_int_poly_def sgn_eq_0_iff) lemma abs_int_poly_dvd[simp]: "abs_int_poly p dvd q \ p dvd q" by (unfold abs_int_poly_def, auto) (*TODO: move & generalize *) lemma (in idom) irreducible_uminus[simp]: "irreducible (-x) \ irreducible x" proof- have "-x = -1 * x" by simp also have "irreducible ... \ irreducible x" by (rule irreducible_mult_unit_left, auto) finally show ?thesis. qed lemma irreducible_abs_int_poly[simp]: "irreducible (abs_int_poly p) \ irreducible p" by (unfold abs_int_poly_def, auto) lemma coeff_abs_int_poly[simp]: "coeff (abs_int_poly p) n = (if lead_coeff p < 0 then - coeff p n else coeff p n)" by (simp add: abs_int_poly_def) lemma lead_coeff_abs_int_poly[simp]: "lead_coeff (abs_int_poly p) = abs (lead_coeff p)" by auto lemma ipoly_abs_int_poly_eq_zero_iff[simp]: "ipoly (abs_int_poly p) (x :: 'a :: comm_ring_1) = 0 \ ipoly p x = 0" by (auto simp: abs_int_poly_def sgn_eq_0_iff of_int_poly_hom.hom_uminus) lemma abs_int_poly_represents[simp]: "abs_int_poly p represents x \ p represents x" by (auto elim!:representsE) (* TODO: Move *) lemma content_pCons[simp]: "content (pCons a p) = gcd a (content p)" by (unfold content_def coeffs_pCons_eq_cCons cCons_def, auto) lemma content_uminus[simp]: fixes p :: "'a :: ring_gcd poly" shows "content (-p) = content p" by (induct p, auto) lemma primitive_abs_int_poly[simp]: "primitive (abs_int_poly p) \ primitive p" by (auto simp: abs_int_poly_def) lemma abs_int_poly_inv[simp]: "smult (sgn (lead_coeff p)) (abs_int_poly p) = p" by (cases "lead_coeff p > 0", auto simp: abs_int_poly_def) definition cf_pos :: "int poly \ bool" where "cf_pos p = (content p = 1 \ lead_coeff p > 0)" definition cf_pos_poly :: "int poly \ int poly" where "cf_pos_poly f = (let c = content f; d = (sgn (lead_coeff f) * c) in sdiv_poly f d)" lemma sgn_is_unit[intro!]: fixes x :: "'a :: linordered_idom" (* find/make better class *) assumes "x \ 0" shows "sgn x dvd 1" using assms by(cases x "0::'a" rule:linorder_cases, auto) lemma cf_pos_poly_0[simp]: "cf_pos_poly 0 = 0" by (unfold cf_pos_poly_def sdiv_poly_def, auto) lemma cf_pos_poly_eq_0[simp]: "cf_pos_poly f = 0 \ f = 0" proof(cases "f = 0") case True thus ?thesis unfolding cf_pos_poly_def Let_def by (simp add: sdiv_poly_def) next case False then have lc0: "lead_coeff f \ 0" by auto then have s0: "sgn (lead_coeff f) \ 0" (is "?s \ 0") and "content f \ 0" (is "?c \ 0") by (auto simp: sgn_0_0) then have sc0: "?s * ?c \ 0" by auto { fix i from content_dvd_coeff sgn_is_unit[OF lc0] have "?s * ?c dvd coeff f i" by (auto simp: unit_dvd_iff) then have "coeff f i div (?s * ?c) = 0 \ coeff f i = 0" by (auto simp:dvd_div_eq_0_iff) } note * = this show ?thesis unfolding cf_pos_poly_def Let_def sdiv_poly_def poly_eq_iff by (auto simp: coeff_map_poly *) qed lemma shows cf_pos_poly_main: "smult (sgn (lead_coeff f) * content f) (cf_pos_poly f) = f" (is ?g1) and content_cf_pos_poly[simp]: "content (cf_pos_poly f) = (if f = 0 then 0 else 1)" (is ?g2) and lead_coeff_cf_pos_poly[simp]: "lead_coeff (cf_pos_poly f) > 0 \ f \ 0" (is ?g3) and cf_pos_poly_dvd[simp]: "cf_pos_poly f dvd f" (is ?g4) proof(atomize(full), (cases "f = 0"; intro conjI)) case True then show ?g1 ?g2 ?g3 ?g4 by simp_all next case f0: False let ?s = "sgn (lead_coeff f)" have s: "?s \ {-1,1}" using f0 unfolding sgn_if by auto define g where "g \ smult ?s f" define d where "d \ ?s * content f" have "content g = content ([:?s:] * f)" unfolding g_def by simp also have "\ = content [:?s:] * content f" unfolding content_mult by simp also have "content [:?s:] = 1" using s by (auto simp: content_def) finally have cg: "content g = content f" by simp from f0 have d: "cf_pos_poly f = sdiv_poly f d" by (auto simp: cf_pos_poly_def Let_def d_def) let ?g = "primitive_part g" define ng where "ng = primitive_part g" note d also have "sdiv_poly f d = sdiv_poly g (content g)" unfolding cg unfolding g_def d_def by (rule poly_eqI, unfold coeff_sdiv_poly coeff_smult, insert s, auto simp: div_minus_right) finally have fg: "cf_pos_poly f = primitive_part g" unfolding primitive_part_alt_def . have "lead_coeff f \ 0" using f0 by auto hence lg: "lead_coeff g > 0" unfolding g_def lead_coeff_smult by (meson linorder_neqE_linordered_idom sgn_greater sgn_less zero_less_mult_iff) hence g0: "g \ 0" by auto from f0 content_primitive_part[OF this] show ?g2 unfolding fg by auto from g0 have "content g \ 0" by simp with arg_cong[OF content_times_primitive_part[of g], of lead_coeff, unfolded lead_coeff_smult] lg content_ge_0_int[of g] have lg': "lead_coeff ng > 0" unfolding ng_def by (metis dual_order.antisym dual_order.strict_implies_order zero_less_mult_iff) with f0 show ?g3 unfolding fg ng_def by auto have d0: "d \ 0" using s f0 by (force simp add: d_def) have "smult d (cf_pos_poly f) = smult ?s (smult (content f) (sdiv_poly (smult ?s f) (content f)))" unfolding fg primitive_part_alt_def cg by (simp add: g_def d_def) also have "sdiv_poly (smult ?s f) (content f) = smult ?s (sdiv_poly f (content f))" using s by (metis cg g_def primitive_part_alt_def primitive_part_smult_int sgn_sgn) finally have "smult d (cf_pos_poly f) = smult (content f) (primitive_part f)" unfolding primitive_part_alt_def using s by auto also have "\ = f" by (rule content_times_primitive_part) finally have df: "smult d (cf_pos_poly f) = f" . with d0 show ?g1 by (auto simp: d_def) from df have *: "f = cf_pos_poly f * [:d:]" by simp from dvdI[OF this] show ?g4. qed (* TODO: remove *) lemma irreducible_connect_int: fixes p :: "int poly" assumes ir: "irreducible\<^sub>d p" and c: "content p = 1" shows "irreducible p" using c primitive_iff_content_eq_1 ir irreducible_primitive_connect by blast lemma fixes x :: "'a :: {idom,ring_char_0}" shows ipoly_cf_pos_poly_eq_0[simp]: "ipoly (cf_pos_poly p) x = 0 \ ipoly p x = 0" and degree_cf_pos_poly[simp]: "degree (cf_pos_poly p) = degree p" and cf_pos_cf_pos_poly[intro]: "p \ 0 \ cf_pos (cf_pos_poly p)" proof- show "degree (cf_pos_poly p) = degree p" by (subst(3) cf_pos_poly_main[symmetric], auto simp:sgn_eq_0_iff) { assume p: "p \ 0" show "cf_pos (cf_pos_poly p)" using cf_pos_poly_main p by (auto simp: cf_pos_def) have "(ipoly (cf_pos_poly p) x = 0) = (ipoly p x = 0)" apply (subst(3) cf_pos_poly_main[symmetric]) by (auto simp: sgn_eq_0_iff hom_distribs) } then show "(ipoly (cf_pos_poly p) x = 0) = (ipoly p x = 0)" by (cases "p = 0", auto) qed lemma cf_pos_poly_eq_1: "cf_pos_poly f = 1 \ degree f = 0 \ f \ 0" (is "?l \ ?r") proof(intro iffI conjI) assume ?r then have df0: "degree f = 0" and f0: "f \ 0" by auto from degree0_coeffs[OF df0] obtain f0 where f: "f = [:f0:]" by auto show "cf_pos_poly f = 1" using f0 unfolding f cf_pos_poly_def Let_def sdiv_poly_def by (auto simp: content_def mult_sgn_abs) next assume l: ?l then have "degree (cf_pos_poly f) = 0" by auto then show "degree f = 0" by simp from l have "cf_pos_poly f \ 0" by auto then show "f \ 0" by simp qed lemma irr_cf_root_free_poly_rat[simp]: "irreducible (poly_rat x)" "lead_coeff (poly_rat x) > 0" "primitive (poly_rat x)" "root_free (poly_rat x)" "square_free (poly_rat x)" proof - obtain n d where x: "quotient_of x = (n,d)" by force hence id: "poly_rat x = [:-n,d:]" by (auto simp: poly_rat_def) from quotient_of_denom_pos[OF x] have d: "d > 0" by auto show "root_free (poly_rat x)" unfolding id root_free_def using d by auto show "lead_coeff (poly_rat x) > 0" "primitive (poly_rat x)" unfolding id cf_pos_def using d quotient_of_coprime[OF x] by (auto simp: content_def) from this[unfolded cf_pos_def] show irr: "irreducible (poly_rat x)" unfolding id using d by (auto intro!: linear_irreducible_int) show "square_free (poly_rat x)" by (rule irreducible_imp_square_free[OF irr]) qed lemma poly_rat[simp]: "ipoly (poly_rat x) (of_rat x :: 'a :: field_char_0) = 0" "ipoly (poly_rat x) x = 0" "poly_rat x \ 0" "ipoly (poly_rat x) y = 0 \ y = (of_rat x :: 'a)" proof - from irr_cf_root_free_poly_rat(1)[of x] show "poly_rat x \ 0" unfolding Factorial_Ring.irreducible_def by auto obtain n d where x: "quotient_of x = (n,d)" by force hence id: "poly_rat x = [:-n,d:]" by (auto simp: poly_rat_def) from quotient_of_denom_pos[OF x] have d: "d \ 0" by auto have "y * of_int d = of_int n \ y = of_int n / of_int d" using d by (simp add: eq_divide_imp) with d id show "ipoly (poly_rat x) (of_rat x) = 0" "ipoly (poly_rat x) x = 0" "ipoly (poly_rat x) y = 0 \ y = (of_rat x :: 'a)" by (auto simp: of_rat_minus of_rat_divide simp: quotient_of_div[OF x]) qed lemma poly_rat_represents_of_rat: "(poly_rat x) represents (of_rat x)" by auto lemma ipoly_smult_0_iff: assumes c: "c \ 0" shows "(ipoly (smult c p) x = (0 :: real)) = (ipoly p x = 0)" using c by (simp add: hom_distribs) (* TODO *) lemma not_irreducibleD: assumes "\ irreducible x" and "x \ 0" and "\ x dvd 1" shows "\y z. x = y * z \ \ y dvd 1 \ \ z dvd 1" using assms apply (unfold Factorial_Ring.irreducible_def) by auto lemma cf_pos_poly_represents[simp]: "(cf_pos_poly p) represents x \ p represents x" unfolding represents_def by auto definition factors_of_int_poly :: "int poly \ int poly list" where "factors_of_int_poly p = map (abs_int_poly o fst) (snd (factorize_int_poly p))" lemma factors_of_int_poly_const: assumes "degree p = 0" shows "factors_of_int_poly p = []" proof - from degree0_coeffs[OF assms] obtain a where p: "p = [: a :]" by auto show ?thesis unfolding p factors_of_int_poly_def factorize_int_poly_generic_def x_split_def by (cases "a = 0", auto simp add: Let_def factorize_int_last_nz_poly_def) qed lemma coprime_prod: (* TODO: move *) "a \ 0 \ c \ 0 \ coprime (a * b) (c * d) \ coprime b (d::'a::{semiring_gcd})" by auto lemma smult_prod: (* TODO: move or find corresponding lemma *) "smult a b = monom a 0 * b" by (simp add: monom_0) lemma degree_map_poly_2: assumes "f (lead_coeff p) \ 0" shows "degree (map_poly f p) = degree p" proof (cases "p=0") case False thus ?thesis unfolding degree_eq_length_coeffs Polynomial.coeffs_map_poly using assms by (simp add:coeffs_def) qed auto lemma degree_of_gcd: "degree (gcd q r) \ 0 \ - degree (gcd (of_int_poly q :: 'a :: {field_char_0,euclidean_ring_gcd} poly) (of_int_poly r)) \ 0" + degree (gcd (of_int_poly q :: 'a :: {field_char_0, field_gcd} poly) (of_int_poly r)) \ 0" proof - let ?r = "of_rat :: rat \ 'a" interpret rpoly: field_hom' ?r by (unfold_locales, auto simp: of_rat_add of_rat_mult) { fix p have "of_int_poly p = map_poly (?r o of_int) p" unfolding o_def by auto also have "\ = map_poly ?r (map_poly of_int p)" by (subst map_poly_map_poly, auto) finally have "of_int_poly p = map_poly ?r (map_poly of_int p)" . } note id = this show ?thesis unfolding id by (fold hom_distribs, simp add: gcd_rat_to_gcd_int) qed lemma irreducible_cf_pos_poly: assumes irr: "irreducible p" and deg: "degree p \ 0" shows "irreducible (cf_pos_poly p)" (is "irreducible ?p") proof (unfold irreducible_altdef, intro conjI allI impI) from irr show "?p \ 0" by auto from deg have "degree ?p \ 0" by simp then show "\ ?p dvd 1" unfolding poly_dvd_1 by auto fix b assume "b dvd cf_pos_poly p" also note cf_pos_poly_dvd finally have "b dvd p". with irr[unfolded irreducible_altdef] have "p dvd b \ b dvd 1" by auto then show "?p dvd b \ b dvd 1" by (auto dest: dvd_trans[OF cf_pos_poly_dvd]) qed locale dvd_preserving_hom = comm_semiring_1_hom + assumes hom_eq_mult_hom_imp: "hom x = hom y * hz \ \z. hz = hom z \ x = y * z" begin lemma hom_dvd_hom_iff[simp]: "hom x dvd hom y \ x dvd y" proof assume "hom x dvd hom y" then obtain hz where "hom y = hom x * hz" by (elim dvdE) from hom_eq_mult_hom_imp[OF this] obtain z where "hz = hom z" and mult: "y = x * z" by auto then show "x dvd y" by auto qed auto sublocale unit_preserving_hom proof unfold_locales fix x assume "hom x dvd 1" then have "hom x dvd hom 1" by simp then show "x dvd 1" by (unfold hom_dvd_hom_iff) qed sublocale zero_hom_0 proof (unfold_locales) fix a :: 'a assume "hom a = 0" then have "hom 0 dvd hom a" by auto then have "0 dvd a" by (unfold hom_dvd_hom_iff) then show "a = 0" by auto qed end lemma factors_of_int_poly: - defines "rp \ ipoly :: int poly \ 'a :: {euclidean_ring_gcd,field_char_0} \ 'a" + defines "rp \ ipoly :: int poly \ 'a :: {field_gcd,field_char_0} \ 'a" assumes "factors_of_int_poly p = qs" shows "\ q. q \ set qs \ irreducible q \ lead_coeff q > 0 \ degree q \ degree p \ degree q \ 0" "p \ 0 \ rp p x = 0 \ (\ q \ set qs. rp q x = 0)" "p \ 0 \ rp p x = 0 \ \! q \ set qs. rp q x = 0" "distinct qs" proof - obtain c qis where factt: "factorize_int_poly p = (c,qis)" by force from assms[unfolded factors_of_int_poly_def factt] have qs: "qs = map (abs_int_poly \ fst) (snd (c, qis))" by auto note fact = factorize_int_poly(1)[OF factt] note fact_mem = factorize_int_poly(2,3)[OF factt] have sqf: "square_free_factorization p (c, qis)" by (rule fact(1)) note sff = square_free_factorizationD[OF sqf] have sff': "p = Polynomial.smult c (\(a, i)\ qis. a ^ Suc i)" unfolding sff(1) prod.distinct_set_conv_list[OF sff(5)] .. { fix q assume q: "q \ set qs" then obtain r i where qi: "(r,i) \ set qis" and qr: "q = abs_int_poly r" unfolding qs by auto from split_list[OF qi] obtain qis1 qis2 where qis: "qis = qis1 @ (r,i) # qis2" by auto have dvd: "r dvd p" unfolding sff' qis dvd_def by (intro exI[of _ "smult c (r ^ i * (\(a, i)\qis1 @ qis2. a ^ Suc i))"], auto) from fact_mem[OF qi] have r0: "r \ 0" by auto from qi factt have p: "p \ 0" by (cases p, auto) with dvd have deg: "degree r \ degree p" by (metis dvd_imp_degree_le) with fact_mem[OF qi] r0 show "irreducible q \ lead_coeff q > 0 \ degree q \ degree p \ degree q \ 0" unfolding qr lead_coeff_abs_int_poly by auto } note * = this show "distinct qs" unfolding distinct_conv_nth proof (intro allI impI) fix i j assume "i < length qs" "j < length qs" and diff: "i \ j" hence ij: "i < length qis" "j < length qis" and id: "qs ! i = abs_int_poly (fst (qis ! i))" "qs ! j = abs_int_poly (fst (qis ! j))" unfolding qs by auto obtain qi I where qi: "qis ! i = (qi, I)" by force obtain qj J where qj: "qis ! j = (qj, J)" by force from sff(5)[unfolded distinct_conv_nth, rule_format, OF ij diff] qi qj have diff: "(qi, I) \ (qj, J)" by auto from ij qi qj have "(qi, I) \ set qis" "(qj, J) \ set qis" unfolding set_conv_nth by force+ from sff(3)[OF this diff] sff(2) this have cop: "coprime qi qj" "degree qi \ 0" "degree qj \ 0" by auto note i = cf_pos_poly_main[of qi, unfolded smult_prod monom_0] note j = cf_pos_poly_main[of qj, unfolded smult_prod monom_0] from cop(2) i have deg: "degree (qs ! i) \ 0" by (auto simp: id qi) have cop: "coprime (qs ! i) (qs ! j)" unfolding id qi qj fst_conv apply (rule coprime_prod[of "[:sgn (lead_coeff qi):]" "[:sgn (lead_coeff qj):]"]) using cop unfolding i j by (auto simp: sgn_eq_0_iff) show "qs ! i \ qs ! j" proof assume id: "qs ! i = qs ! j" have "degree (gcd (qs ! i) (qs ! j)) = degree (qs ! i)" unfolding id by simp also have "\ \ 0" using deg by simp finally show False using cop by simp qed qed assume p: "p \ 0" from fact(1) p have c: "c \ 0" using sff(1) by auto let ?r = "of_int :: int \ 'a" let ?rp = "map_poly ?r" have rp: "\ x p. rp p x = 0 \ poly (?rp p) x = 0" unfolding rp_def .. have "rp p x = 0 \ rp (\(x, y)\qis. x ^ Suc y) x = 0" unfolding sff'(1) unfolding rp hom_distribs using c by simp also have "\ = (\ (q,i) \set qis. poly (?rp (q ^ Suc i)) x = 0)" unfolding qs rp of_int_poly_hom.hom_prod_list poly_prod_list_zero_iff set_map by fastforce also have "\ = (\ (q,i) \set qis. poly (?rp q) x = 0)" unfolding of_int_poly_hom.hom_power poly_power_zero_iff by auto also have "\ = (\ q \ fst ` set qis. poly (?rp q) x = 0)" by force also have "\ = (\ q \ set qs. rp q x = 0)" unfolding rp qs snd_conv o_def bex_simps set_map by simp finally show iff: "rp p x = 0 \ (\ q \ set qs. rp q x = 0)" by auto assume "rp p x = 0" with iff obtain q where q: "q \ set qs" and rtq: "rp q x = 0" by auto then obtain i q' where qi: "(q',i) \ set qis" and qq': "q = abs_int_poly q'" unfolding qs by auto show "\! q \ set qs. rp q x = 0" proof (intro ex1I, intro conjI, rule q, rule rtq, clarify) fix r assume "r \ set qs" and rtr: "rp r x = 0" then obtain j r' where rj: "(r',j) \ set qis" and rr': "r = abs_int_poly r'" unfolding qs by auto from rtr rtq have rtr: "rp r' x = 0" and rtq: "rp q' x = 0" unfolding rp rr' qq' by auto from rtr rtq have "[:-x,1:] dvd ?rp q'" "[:-x,1:] dvd ?rp r'" unfolding rp by (auto simp: poly_eq_0_iff_dvd) hence "[:-x,1:] dvd gcd (?rp q') (?rp r')" by simp hence "gcd (?rp q') (?rp r') = 0 \ degree (gcd (?rp q') (?rp r')) \ 0" by (metis is_unit_gcd_iff is_unit_iff_degree is_unit_pCons_iff one_poly_eq_simps(1)) hence "gcd q' r' = 0 \ degree (gcd q' r') \ 0" unfolding gcd_eq_0_iff degree_of_gcd[of q' r',symmetric] by auto hence "\ coprime q' r'" by auto with sff(3)[OF qi rj] have "q' = r'" by auto thus "r = q" unfolding rr' qq' by simp qed qed lemma factors_int_poly_represents: - fixes x :: "'a :: {field_char_0,euclidean_ring_gcd}" + fixes x :: "'a :: {field_char_0,field_gcd}" assumes p: "p represents x" shows "\ q \ set (factors_of_int_poly p). q represents x \ irreducible q \ lead_coeff q > 0 \ degree q \ degree p" proof - from representsD[OF p] have p: "p \ 0" and rt: "ipoly p x = 0" by auto note fact = factors_of_int_poly[OF refl] from fact(2)[OF p, of x] rt obtain q where q: "q \ set (factors_of_int_poly p)" and rt: "ipoly q x = 0" by auto from fact(1)[OF q] rt show ?thesis by (intro bexI[OF _ q], auto simp: represents_def irreducible_def) qed lemma smult_inverse_monom:"p \ 0 \ smult (inverse c) (p::rat poly) = 1 \ p = [: c :]" proof (cases "c=0") case True thus "p \ 0 \ ?thesis" by auto next case False thus ?thesis by (metis left_inverse right_inverse smult_1 smult_1_left smult_smult) qed lemma of_int_monom:"of_int_poly p = [:rat_of_int c:] \ p = [: c :]" by (induct p, auto) lemma degree_0_content: fixes p :: "int poly" assumes deg: "degree p = 0" shows "content p = abs (coeff p 0)" proof- from deg obtain a where p: "p = [:a:]" by (auto dest: degree0_coeffs) show ?thesis by (auto simp: p) qed lemma prime_elem_imp_gcd_eq: fixes x::"'a:: ring_gcd" shows "prime_elem x \ gcd x y = normalize x \ gcd x y = 1" using prime_elem_imp_coprime [of x y] by (auto simp add: gcd_proj1_iff intro: coprime_imp_gcd_eq_1) lemma irreducible_pos_gcd: fixes p :: "int poly" assumes ir: "irreducible p" and pos: "lead_coeff p > 0" shows "gcd p q \ {1,p}" proof- from pos have "[:sgn (lead_coeff p):] = 1" by auto with prime_elem_imp_gcd_eq[of p, unfolded prime_elem_iff_irreducible, OF ir, of q] show ?thesis by (auto simp: normalize_poly_def) qed lemma irreducible_pos_gcd_twice: fixes p q :: "int poly" assumes p: "irreducible p" "lead_coeff p > 0" and q: "irreducible q" "lead_coeff q > 0" shows "gcd p q = 1 \ p = q" proof (cases "gcd p q = 1") case False note pq = this have "p = gcd p q" using irreducible_pos_gcd [OF p, of q] pq by auto also have "\ = q" using irreducible_pos_gcd [OF q, of p] pq by (auto simp add: ac_simps) finally show ?thesis by auto qed simp interpretation of_rat_hom: field_hom_0' of_rat.. lemma poly_zero_imp_not_unit: assumes "poly p x = 0" shows "\ p dvd 1" proof (rule notI) assume "p dvd 1" from poly_hom.hom_dvd_1[OF this] have "poly p x dvd 1" by auto with assms show False by auto qed lemma poly_prod_mset_zero_iff: fixes x :: "'a :: idom" shows "poly (prod_mset F) x = 0 \ (\f \# F. poly f x = 0)" by (induct F, auto simp: poly_mult_zero_iff) lemma algebraic_imp_represents_irreducible: fixes x :: "'a :: field_char_0" assumes "algebraic x" shows "\p. p represents x \ irreducible p" proof - from assms obtain p where px0: "ipoly p x = 0" and p0: "p \ 0" unfolding algebraic_altdef_ipoly by auto from poly_zero_imp_not_unit[OF px0] have "\ p dvd 1" by (auto dest: of_int_poly_hom.hom_dvd_1[where 'a = 'a]) from mset_factors_exist[OF p0 this] obtain F where F: "mset_factors F p" by auto then have "p = prod_mset F" by auto also have "(of_int_poly ... :: 'a poly) = prod_mset (image_mset of_int_poly F)" by simp finally have "poly ... x = 0" using px0 by auto from this[unfolded poly_prod_mset_zero_iff] obtain f where "f \# F" and fx0: "ipoly f x = 0" by auto with F have "irreducible f" by auto with fx0 show ?thesis by auto qed lemma algebraic_imp_represents_irreducible_cf_pos: assumes "algebraic (x::'a::field_char_0)" shows "\p. p represents x \ irreducible p \ lead_coeff p > 0 \ primitive p" proof - from algebraic_imp_represents_irreducible[OF assms(1)] obtain p where px: "p represents x" and irr: "irreducible p" by auto let ?p = "cf_pos_poly p" from px irr represents_imp_degree have 1: "?p represents x" and 2: "irreducible ?p" and 3: "cf_pos ?p" by (auto intro: irreducible_cf_pos_poly) then show ?thesis by (auto intro: exI[of _ ?p] simp: cf_pos_def) qed -lemma gcd_of_int_poly: "gcd (of_int_poly f) (of_int_poly g :: 'a :: {field_char_0,euclidean_ring_gcd} poly) = +lemma gcd_of_int_poly: "gcd (of_int_poly f) (of_int_poly g :: 'a :: {field_char_0,field_gcd} poly) = smult (inverse (of_int (lead_coeff (gcd f g)))) (of_int_poly (gcd f g))" proof - let ?ia = "of_int_poly :: _ \ 'a poly" let ?ir = "of_int_poly :: _ \ rat poly" let ?ra = "map_poly of_rat :: _ \ 'a poly" have id: "?ia x = ?ra (?ir x)" for x by (subst map_poly_map_poly, auto) show ?thesis unfolding id unfolding of_rat_hom.map_poly_gcd[symmetric] unfolding gcd_rat_to_gcd_int by (auto simp: hom_distribs) qed lemma algebraic_imp_represents_unique: - fixes x :: "'a :: {field_char_0,euclidean_ring_gcd}" + fixes x :: "'a :: {field_char_0,field_gcd}" assumes "algebraic x" shows "\! p. p represents x \ irreducible p \ lead_coeff p > 0" (is "Ex1 ?p") proof - from assms obtain p where p: "?p p" and cfp: "cf_pos p" by (auto simp: cf_pos_def dest: algebraic_imp_represents_irreducible_cf_pos) show ?thesis proof (rule ex1I) show "?p p" by fact fix q assume q: "?p q" then have "q represents x" by auto from represents_imp_degree[OF this] q irreducible_content[of q] have cfq: "cf_pos q" by (auto simp: cf_pos_def) show "q = p" proof (rule ccontr) let ?ia = "map_poly of_int :: int poly \ 'a poly" assume "q \ p" with irreducible_pos_gcd_twice[of p q] p q cfp cfq have gcd: "gcd p q = 1" by auto from p q have rt: "ipoly p x = 0" "ipoly q x = 0" unfolding represents_def by auto define c :: 'a where "c = inverse (of_int (lead_coeff (gcd p q)))" have rt: "poly (?ia p) x = 0" "poly (?ia q) x = 0" using rt by auto hence "[:-x,1:] dvd ?ia p" "[:-x,1:] dvd ?ia q" unfolding poly_eq_0_iff_dvd by auto hence "[:-x,1:] dvd gcd (?ia p) (?ia q)" by (rule gcd_greatest) also have "\ = smult c (?ia (gcd p q))" unfolding gcd_of_int_poly c_def .. also have "?ia (gcd p q) = 1" by (simp add: gcd) also have "smult c 1 = [: c :]" by simp finally show False using c_def gcd by (simp add: dvd_iff_poly_eq_0) qed qed qed corollary irreducible_represents_imp_degree: - fixes x :: "'a :: {field_char_0,euclidean_ring_gcd}" + fixes x :: "'a :: {field_char_0,field_gcd}" assumes "irreducible f" and "f represents x" and "g represents x" shows "degree f \ degree g" proof - from factors_of_int_poly(1)[OF refl, of _ g] factors_of_int_poly(3)[OF refl, of g x] assms(3) obtain h where *: "h represents x" "degree h \ degree g" "irreducible h" by blast let ?af = "abs_int_poly f" let ?ah = "abs_int_poly h" from assms have af: "irreducible ?af" "?af represents x" "lead_coeff ?af > 0" by fastforce+ from * have ah: "irreducible ?ah" "?ah represents x" "lead_coeff ?ah > 0" by fastforce+ from algebraic_imp_represents_unique[of x] af ah have id: "?af = ?ah" unfolding algebraic_iff_represents by blast show ?thesis using arg_cong[OF id, of degree] \degree h \ degree g\ by simp qed lemma ipoly_poly_compose: fixes x :: "'a :: idom" shows "ipoly (p \\<^sub>p q) x = ipoly p (ipoly q x)" proof (induct p) case (pCons a p) have "ipoly ((pCons a p) \\<^sub>p q) x = of_int a + ipoly (q * p \\<^sub>p q) x" by (simp add: hom_distribs) also have "ipoly (q * p \\<^sub>p q) x = ipoly q x * ipoly (p \\<^sub>p q) x" by (simp add: hom_distribs) also have "ipoly (p \\<^sub>p q) x = ipoly p (ipoly q x)" unfolding pCons(2) .. also have "of_int a + ipoly q x * \ = ipoly (pCons a p) (ipoly q x)" unfolding map_poly_pCons[OF pCons(1)] by simp finally show ?case . qed simp text \Polynomial for unary minus.\ definition poly_uminus :: "'a :: ring_1 poly \ 'a poly" where [code del]: "poly_uminus p \ \i\degree p. monom ((-1)^i * coeff p i) i" lemma poly_uminus_pCons_pCons[simp]: "poly_uminus (pCons a (pCons b p)) = pCons a (pCons (-b) (poly_uminus p))" (is "?l = ?r") proof(cases "p = 0") case False then have deg: "degree (pCons a (pCons b p)) = Suc (Suc (degree p))" by simp show ?thesis by (unfold poly_uminus_def deg sum.atMost_Suc_shift monom_Suc monom_0 sum_pCons_0_commute, simp) next case True then show ?thesis by (auto simp add: poly_uminus_def monom_0 monom_Suc) qed fun poly_uminus_inner :: "'a :: ring_1 list \ 'a poly" where "poly_uminus_inner [] = 0" | "poly_uminus_inner [a] = [:a:]" | "poly_uminus_inner (a#b#cs) = pCons a (pCons (-b) (poly_uminus_inner cs))" lemma poly_uminus_code[code,simp]: "poly_uminus p = poly_uminus_inner (coeffs p)" proof- have "poly_uminus (Poly as) = poly_uminus_inner as" for as :: "'a list" proof (induct "length as" arbitrary:as rule: less_induct) case less show ?case proof(cases as) case Nil then show ?thesis by (simp add: poly_uminus_def) next case [simp]: (Cons a bs) show ?thesis proof (cases bs) case Nil then show ?thesis by (simp add: poly_uminus_def monom_0) next case [simp]: (Cons b cs) show ?thesis by (simp add: less) qed qed qed from this[of "coeffs p"] show ?thesis by simp qed lemma poly_uminus_inner_0[simp]: "poly_uminus_inner as = 0 \ Poly as = 0" by (induct as rule: poly_uminus_inner.induct, auto) lemma degree_poly_uminus_inner[simp]: "degree (poly_uminus_inner as) = degree (Poly as)" by (induct as rule: poly_uminus_inner.induct, auto) lemma ipoly_uminus_inner[simp]: "ipoly (poly_uminus_inner as) (x::'a::comm_ring_1) = ipoly (Poly as) (-x)" by (induct as rule: poly_uminus_inner.induct, auto simp: hom_distribs ring_distribs) lemma represents_uminus: assumes alg: "p represents x" shows "(poly_uminus p) represents (-x)" proof - from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto hence 0: "poly_uminus p \ 0" by simp show ?thesis by (rule representsI[OF _ 0], insert rp, auto) qed lemma content_poly_uminus_inner[simp]: fixes as :: "'a :: ring_gcd list" shows "content (poly_uminus_inner as) = content (Poly as)" by (induct as rule: poly_uminus_inner.induct, auto) text \Multiplicative inverse is represented by @{const reflect_poly}.\ lemma inverse_pow_minus: assumes "x \ (0 :: 'a :: field)" and "i \ n" shows "inverse x ^ n * x ^ i = inverse x ^ (n - i)" using assms by (simp add: field_class.field_divide_inverse power_diff power_inverse) lemma (in inj_idom_hom) reflect_poly_hom: "reflect_poly (map_poly hom p) = map_poly hom (reflect_poly p)" proof - obtain xs where xs: "rev (coeffs p) = xs" by auto show ?thesis unfolding reflect_poly_def coeffs_map_poly_hom rev_map xs by (induct xs, auto simp: hom_distribs) qed lemma ipoly_reflect_poly: assumes x: "(x :: 'a :: field_char_0) \ 0" shows "ipoly (reflect_poly p) x = x ^ (degree p) * ipoly p (inverse x)" (is "?l = ?r") proof - let ?or = "of_int :: int \ 'a" have hom: "inj_idom_hom ?or" .. show ?thesis using poly_reflect_poly_nz[OF x, of "map_poly ?or p"] by (simp add: inj_idom_hom.reflect_poly_hom[OF hom]) qed lemma represents_inverse: assumes x: "x \ 0" and alg: "p represents x" shows "(reflect_poly p) represents (inverse x)" proof (intro representsI) from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto then show "reflect_poly p \ 0" by simp show "ipoly (reflect_poly p) (inverse x) = 0" by (subst ipoly_reflect_poly, insert x, auto simp:rp) qed lemma inverse_roots: assumes x: "(x :: 'a :: field_char_0) \ 0" shows "ipoly (reflect_poly p) x = 0 \ ipoly p (inverse x) = 0" using x by (auto simp: ipoly_reflect_poly) context fixes n :: nat begin text \Polynomial for n-th root.\ definition poly_nth_root :: "'a :: idom poly \ 'a poly" where "poly_nth_root p = p \\<^sub>p monom 1 n" lemma ipoly_nth_root: fixes x :: "'a :: idom" shows "ipoly (poly_nth_root p) x = ipoly p (x ^ n)" unfolding poly_nth_root_def ipoly_poly_compose by (simp add: map_poly_monom poly_monom) context assumes n: "n \ 0" begin lemma poly_nth_root_0[simp]: "poly_nth_root p = 0 \ p = 0" unfolding poly_nth_root_def by (rule pcompose_eq_0, insert n, auto simp: degree_monom_eq) lemma represents_nth_root: assumes y: "y^n = x" and alg: "p represents x" shows "(poly_nth_root p) represents y" proof - from representsD[OF alg] have "p \ 0" and rp: "ipoly p x = 0" by auto hence 0: "poly_nth_root p \ 0" by simp show ?thesis by (rule representsI[OF _ 0], unfold ipoly_nth_root y rp, simp) qed lemma represents_nth_root_odd_real: assumes alg: "p represents x" and odd: "odd n" shows "(poly_nth_root p) represents (root n x)" by (rule represents_nth_root[OF odd_real_root_pow[OF odd] alg]) lemma represents_nth_root_pos_real: assumes alg: "p represents x" and pos: "x > 0" shows "(poly_nth_root p) represents (root n x)" proof - from n have id: "Suc (n - 1) = n" by auto show ?thesis proof (rule represents_nth_root[OF _ alg]) show "root n x ^ n = x" using id pos by auto qed qed lemma represents_nth_root_neg_real: assumes alg: "p represents x" and neg: "x < 0" shows "(poly_uminus (poly_nth_root (poly_uminus p))) represents (root n x)" proof - have rt: "root n x = - root n (-x)" unfolding real_root_minus by simp show ?thesis unfolding rt by (rule represents_uminus[OF represents_nth_root_pos_real[OF represents_uminus[OF alg]]], insert neg, auto) qed end end lemma represents_csqrt: assumes alg: "p represents x" shows "(poly_nth_root 2 p) represents (csqrt x)" by (rule represents_nth_root[OF _ _ alg], auto) lemma represents_sqrt: assumes alg: "p represents x" and pos: "x \ 0" shows "(poly_nth_root 2 p) represents (sqrt x)" by (rule represents_nth_root[OF _ _ alg], insert pos, auto) lemma represents_degree: assumes "p represents x" shows "degree p \ 0" proof assume "degree p = 0" from degree0_coeffs[OF this] obtain c where p: "p = [:c:]" by auto from assms[unfolded represents_def p] show False by auto qed text \Polynomial for multiplying a rational number with an algebraic number.\ definition poly_mult_rat_main where "poly_mult_rat_main n d (f :: 'a :: idom poly) = (let fs = coeffs f; k = length fs in poly_of_list (map (\ (fi, i). fi * d ^ i * n ^ (k - Suc i)) (zip fs [0 ..< k])))" definition poly_mult_rat :: "rat \ int poly \ int poly" where "poly_mult_rat r p \ case quotient_of r of (n,d) \ poly_mult_rat_main n d p" lemma coeff_poly_mult_rat_main: "coeff (poly_mult_rat_main n d f) i = coeff f i * n ^ (degree f - i) * d ^ i" proof - have id: "coeff (poly_mult_rat_main n d f) i = (coeff f i * d ^ i) * n ^ (length (coeffs f) - Suc i)" unfolding poly_mult_rat_main_def Let_def poly_of_list_def coeff_Poly unfolding nth_default_coeffs_eq[symmetric] unfolding nth_default_def by auto show ?thesis unfolding id by (simp add: degree_eq_length_coeffs) qed lemma degree_poly_mult_rat_main: "n \ 0 \ degree (poly_mult_rat_main n d f) = (if d = 0 then 0 else degree f)" proof (cases "d = 0") case True thus ?thesis unfolding degree_def unfolding coeff_poly_mult_rat_main by simp next case False hence id: "(d = 0) = False" by simp show "n \ 0 \ ?thesis" unfolding degree_def coeff_poly_mult_rat_main id by (simp add: id) qed lemma ipoly_mult_rat_main: fixes x :: "'a :: {field,ring_char_0}" assumes "d \ 0" and "n \ 0" shows "ipoly (poly_mult_rat_main n d p) x = of_int n ^ degree p * ipoly p (x * of_int d / of_int n)" proof - from assms have d: "(if d = 0 then t else f) = f" for t f :: 'b by simp show ?thesis unfolding poly_altdef of_int_hom.coeff_map_poly_hom mult.assoc[symmetric] of_int_mult[symmetric] sum_distrib_left unfolding of_int_hom.degree_map_poly_hom degree_poly_mult_rat_main[OF assms(2)] d proof (rule sum.cong[OF refl]) fix i assume "i \ {..degree p}" hence i: "i \ degree p" by auto hence id: "of_int n ^ (degree p - i) = (of_int n ^ degree p / of_int n ^ i :: 'a)" by (simp add: assms(2) power_diff) thus "of_int (coeff (poly_mult_rat_main n d p) i) * x ^ i = of_int n ^ degree p * of_int (coeff p i) * (x * of_int d / of_int n) ^ i" unfolding coeff_poly_mult_rat_main by (simp add: field_simps) qed qed lemma degree_poly_mult_rat[simp]: assumes "r \ 0" shows "degree (poly_mult_rat r p) = degree p" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto with assms r have n0: "n \ 0" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id degree_poly_mult_rat_main[OF n0] using d by simp qed lemma ipoly_mult_rat: assumes r0: "r \ 0" shows "ipoly (poly_mult_rat r p) x = of_int (fst (quotient_of r)) ^ degree p * ipoly p (x * inverse (of_rat r))" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto from r r0 have n: "n \ 0" by simp from r d n have inv: "of_int d / of_int n = inverse r" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id ipoly_mult_rat_main[OF d n] quot fst_conv of_rat_inverse[symmetric] inv[symmetric] by (simp add: of_rat_divide) qed lemma poly_mult_rat_main_0[simp]: assumes "n \ 0" "d \ 0" shows "poly_mult_rat_main n d p = 0 \ p = 0" proof assume "p = 0" thus "poly_mult_rat_main n d p = 0" by (simp add: poly_mult_rat_main_def) next assume 0: "poly_mult_rat_main n d p = 0" { fix i from 0 have "coeff (poly_mult_rat_main n d p) i = 0" by simp hence "coeff p i = 0" unfolding coeff_poly_mult_rat_main using assms by simp } thus "p = 0" by (intro poly_eqI, auto) qed lemma poly_mult_rat_0[simp]: assumes r0: "r \ 0" shows "poly_mult_rat r p = 0 \ p = 0" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" by auto from r r0 have n: "n \ 0" by simp from quot have id: "poly_mult_rat r p = poly_mult_rat_main n d p" unfolding poly_mult_rat_def by simp show ?thesis unfolding id using n d by simp qed lemma represents_mult_rat: assumes r: "r \ 0" and "p represents x" shows "(poly_mult_rat r p) represents (of_rat r * x)" using assms unfolding represents_def ipoly_mult_rat[OF r] by (simp add: field_simps) text \Polynomial for adding a rational number on an algebraic number. Again, we do not have to factor afterwards.\ definition poly_add_rat :: "rat \ int poly \ int poly" where "poly_add_rat r p \ case quotient_of r of (n,d) \ (poly_mult_rat_main d 1 p \\<^sub>p [:-n,d:])" lemma poly_add_rat_code[code]: "poly_add_rat r p \ case quotient_of r of (n,d) \ let p' = (let fs = coeffs p; k = length fs in poly_of_list (map (\(fi, i). fi * d ^ (k - Suc i)) (zip fs [0..\<^sub>p [:-n,d:] in p''" unfolding poly_add_rat_def poly_mult_rat_main_def Let_def by simp lemma degree_poly_add_rat[simp]: "degree (poly_add_rat r p) = degree p" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto show ?thesis unfolding poly_add_rat_def quot split by (simp add: degree_poly_mult_rat_main d) qed lemma ipoly_add_rat: "ipoly (poly_add_rat r p) x = (of_int (snd (quotient_of r)) ^ degree p) * ipoly p (x - of_rat r)" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto have id: "ipoly [:- n, 1:] (x / of_int d :: 'a) = - of_int n + x / of_int d" by simp show ?thesis unfolding poly_add_rat_def quot split by (simp add: ipoly_mult_rat_main ipoly_poly_compose d r degree_poly_mult_rat_main field_simps id of_rat_divide) qed lemma poly_add_rat_0[simp]: "poly_add_rat r p = 0 \ p = 0" proof - obtain n d where quot: "quotient_of r = (n,d)" by force from quotient_of_div[OF quot] have r: "r = of_int n / of_int d" by auto from quotient_of_denom_pos[OF quot] have d: "d \ 0" "d > 0" by auto show ?thesis unfolding poly_add_rat_def quot split by (simp add: d pcompose_eq_0) qed lemma add_rat_roots: "ipoly (poly_add_rat r p) x = 0 \ ipoly p (x - of_rat r) = 0" unfolding ipoly_add_rat using quotient_of_nonzero by auto lemma represents_add_rat: assumes "p represents x" shows "(poly_add_rat r p) represents (of_rat r + x)" using assms unfolding represents_def ipoly_add_rat by simp (* TODO: move? *) lemmas pos_mult[simplified,simp] = mult_less_cancel_left_pos[of _ 0] mult_less_cancel_left_pos[of _ _ 0] lemma ipoly_add_rat_pos_neg: "ipoly (poly_add_rat r p) (x::'a::linordered_field) < 0 \ ipoly p (x - of_rat r) < 0" "ipoly (poly_add_rat r p) (x::'a::linordered_field) > 0 \ ipoly p (x - of_rat r) > 0" using quotient_of_nonzero unfolding ipoly_add_rat by auto lemma sgn_ipoly_add_rat[simp]: "sgn (ipoly (poly_add_rat r p) (x::'a::linordered_field)) = sgn (ipoly p (x - of_rat r))" (is "sgn ?l = sgn ?r") using ipoly_add_rat_pos_neg[of r p x] by (cases ?r "0::'a" rule: linorder_cases,auto simp: sgn_1_pos sgn_1_neg sgn_eq_0_iff) lemma irreducible_preservation: - fixes x :: "'a :: {field_char_0,euclidean_ring_gcd}" + fixes x :: "'a :: {field_char_0,field_gcd}" assumes irr: "irreducible p" and x: "p represents x" and y: "q represents y" and deg: "degree p \ degree q" and f: "\ q. q represents y \ (f q) represents x \ degree (f q) \ degree q" and pr: "primitive q" shows "irreducible q" proof (rule ccontr) define pp where "pp = abs_int_poly p" have dp: "degree p \ 0" using x by (rule represents_degree) have dq: "degree q \ 0" using y by (rule represents_degree) from dp have p0: "p \ 0" by auto from x deg irr p0 have irr: "irreducible pp" and x: "pp represents x" and deg: "degree pp \ degree q" and cf_pos: "lead_coeff pp > 0" unfolding pp_def lead_coeff_abs_int_poly by (auto intro!: representsI) from x have ax: "algebraic x" unfolding algebraic_altdef_ipoly represents_def by blast assume "\ ?thesis" from this irreducible_connect_int[of q] pr have "\ irreducible\<^sub>d q" by auto from this dq obtain r where r: "degree r \ 0" "degree r < degree q" and "r dvd q" by auto then obtain rr where q: "q = r * rr" unfolding dvd_def by auto have "degree q = degree r + degree rr" using dq unfolding q by (subst degree_mult_eq, auto) with r have rr: "degree rr \ 0" "degree rr < degree q" by auto from representsD(2)[OF y, unfolded q hom_distribs] have "ipoly r y = 0 \ ipoly rr y = 0" by auto with r rr have "r represents y \ rr represents y" unfolding represents_def by auto with r rr obtain r where r: "r represents y" "degree r < degree q" by blast from f[OF r(1)] deg r(2) obtain r where r: "r represents x" "degree r < degree pp" by auto from factors_int_poly_represents[OF r(1)] r(2) obtain r where r: "r represents x" "irreducible r" "lead_coeff r > 0" and deg: "degree r < degree pp" by force from algebraic_imp_represents_unique[OF ax] r irr cf_pos x have "r = pp" by auto with deg show False by auto qed lemma deg_nonzero_represents: assumes deg: "degree p \ 0" shows "\ x :: complex. p represents x" proof - let ?p = "of_int_poly p :: complex poly" from fundamental_theorem_algebra_factorized[of ?p] obtain as c where id: "smult c (\a\as. [:- a, 1:]) = ?p" and len: "length as = degree ?p" by blast have "degree ?p = degree p" by simp with deg len obtain b bs where as: "as = b # bs" by (cases as, auto) have "p represents b" unfolding represents_def id[symmetric] as using deg by auto thus ?thesis by blast qed declare irreducible_const_poly_iff [simp] lemma poly_uminus_irreducible: assumes p: "irreducible (p :: int poly)" and deg: "degree p \ 0" shows "irreducible (poly_uminus p)" proof- from deg_nonzero_represents[OF deg] obtain x :: complex where x: "p represents x" by auto from represents_uminus[OF x] have y: "poly_uminus p represents (- x)" . show ?thesis proof (rule irreducible_preservation[OF p x y], force) from deg irreducible_imp_primitive[OF p] have "primitive p" by auto then show "primitive (poly_uminus p)" by simp fix q assume "q represents (- x)" from represents_uminus[OF this] have "(poly_uminus q) represents x" by simp thus "(poly_uminus q) represents x \ degree (poly_uminus q) \ degree q" by auto qed qed lemma reflect_poly_irreducible: - fixes x :: "'a :: {field_char_0,euclidean_ring_gcd}" + fixes x :: "'a :: {field_char_0,field_gcd}" assumes p: "irreducible p" and x: "p represents x" and x0: "x \ 0" shows "irreducible (reflect_poly p)" proof - from represents_inverse[OF x0 x] have y: "(reflect_poly p) represents (inverse x)" by simp from x0 have ix0: "inverse x \ 0" by auto show ?thesis proof (rule irreducible_preservation[OF p x y]) from x irreducible_imp_primitive[OF p] show "primitive (reflect_poly p)" by (auto simp: content_reflect_poly) fix q assume "q represents (inverse x)" from represents_inverse[OF ix0 this] have "(reflect_poly q) represents x" by simp with degree_reflect_poly_le show "(reflect_poly q) represents x \ degree (reflect_poly q) \ degree q" by auto qed (insert p, auto simp: degree_reflect_poly_le) qed lemma poly_add_rat_irreducible: assumes p: "irreducible p" and deg: "degree p \ 0" shows "irreducible (cf_pos_poly (poly_add_rat r p))" proof - from deg_nonzero_represents[OF deg] obtain x :: complex where x: "p represents x" by auto from represents_add_rat[OF x] have y: "cf_pos_poly (poly_add_rat r p) represents (of_rat r + x)" by simp show ?thesis proof (rule irreducible_preservation[OF p x y], force) fix q assume "q represents (of_rat r + x)" from represents_add_rat[OF this, of "- r"] have "(poly_add_rat (- r) q) represents x" by (simp add: of_rat_minus) thus "(poly_add_rat (- r) q) represents x \ degree (poly_add_rat (- r) q) \ degree q" by auto qed (insert p, auto) qed lemma poly_mult_rat_irreducible: assumes p: "irreducible p" and deg: "degree p \ 0" and r: "r \ 0" shows "irreducible (cf_pos_poly (poly_mult_rat r p))" proof - from deg_nonzero_represents[OF deg] obtain x :: complex where x: "p represents x" by auto from represents_mult_rat[OF r x] have y: "cf_pos_poly (poly_mult_rat r p) represents (of_rat r * x)" by simp show ?thesis proof (rule irreducible_preservation[OF p x y], force simp: r) fix q from r have r': "inverse r \ 0" by simp assume "q represents (of_rat r * x)" from represents_mult_rat[OF r' this] have "(poly_mult_rat (inverse r) q) represents x" using r by (simp add: of_rat_divide field_simps) thus "(poly_mult_rat (inverse r) q) represents x \ degree (poly_mult_rat (inverse r) q) \ degree q" using r by auto qed (insert p r, auto) qed end diff --git a/thys/Algebraic_Numbers/Sturm_Rat.thy b/thys/Algebraic_Numbers/Sturm_Rat.thy --- a/thys/Algebraic_Numbers/Sturm_Rat.thy +++ b/thys/Algebraic_Numbers/Sturm_Rat.thy @@ -1,319 +1,319 @@ section \Separation of Roots: Sturm\ text \We adapt the existing theory on Sturm's theorem to work on rational numbers instead of real numbers. The reason is that we want to implement real numbers as real algebraic numbers with the help of Sturm's theorem to separate the roots. To this end, we just copy the definitions of of the algorithms w.r.t. Sturm and let them be executed on rational numbers. We then prove that corresponds to a homomorphism and therefore can transfer the existing soundness results.\ theory Sturm_Rat imports Sturm_Sequences.Sturm_Theorem Algebraic_Numbers_Prelim begin (* TODO: Move *) lemma root_primitive_part [simp]: fixes p :: "'a :: {semiring_gcd, semiring_no_zero_divisors} poly" shows "poly (primitive_part p) x = 0 \ poly p x = 0" proof(cases "p = 0") case True then show ?thesis by auto next case False have "poly p x = content p * poly (primitive_part p) x" by (metis content_times_primitive_part poly_smult) also have "\ = 0 \ poly (primitive_part p) x = 0" by (simp add: False) finally show ?thesis by auto qed (*TODO: Move*) lemma irreducible_primitive_part: assumes "irreducible p" and "degree p > 0" shows "primitive_part p = p" using irreducible_content[OF assms(1), unfolded primitive_iff_content_eq_1] assms(2) by (auto simp: primitive_part_def abs_poly_def) subsection \Interface for Separating Roots\ text \For a given rational polynomial, we need to know how many real roots are in a given closed interval, and how many real roots are in an interval $(-\infty,r]$.\ datatype root_info = Root_Info (l_r: "rat \ rat \ nat") (number_root: "rat \ nat") hide_const (open) l_r hide_const (open) number_root definition count_roots_interval_sf :: "real poly \ (real \ real \ nat) \ (real \ nat)" where "count_roots_interval_sf p = (let ps = sturm_squarefree p in ((\ a b. sign_changes ps a - sign_changes ps b + (if poly p a = 0 then 1 else 0)), (\ a. sign_changes_neg_inf ps - sign_changes ps a)))" definition count_roots_interval :: "real poly \ (real \ real \ nat) \ (real \ nat)" where "count_roots_interval p = (let ps = sturm p in ((\ a b. sign_changes ps a - sign_changes ps b + (if poly p a = 0 then 1 else 0)), (\ a. sign_changes_neg_inf ps - sign_changes ps a)))" lemma count_roots_interval_iff: "square_free p \ count_roots_interval p = count_roots_interval_sf p" unfolding count_roots_interval_def count_roots_interval_sf_def sturm_squarefree_def square_free_iff_separable separable_def by (cases "p = 0", auto) lemma count_roots_interval_sf: assumes p: "p \ 0" and cr: "count_roots_interval_sf p = (cr,nr)" shows "a \ b \ cr a b = (card {x. a \ x \ x \ b \ poly p x = 0})" "nr a = card {x. x \ a \ poly p x = 0}" proof - have id: "a \ b \ { x. a \ x \ x \ b \ poly p x = 0} = { x. a < x \ x \ b \ poly p x = 0} \ (if poly p a = 0 then {a} else {})" (is "_ \ _ = ?R \ ?S") using not_less by force have RS: "finite ?R" "finite ?S" "?R \ ?S = {}" using p by (auto simp: poly_roots_finite) show "a \ b \ cr a b = (card {x. a \ x \ x \ b \ poly p x = 0})" "nr a = card {x. x \ a \ poly p x = 0}" using cr unfolding arg_cong[OF id, of card] card_Un_disjoint[OF RS] count_roots_interval_sf_def count_roots_between_correct[symmetric] count_roots_below_correct[symmetric] count_roots_below_def count_roots_between_def Let_def using p by auto qed lemma count_roots_interval: assumes cr: "count_roots_interval p = (cr,nr)" and sf: "square_free p" shows "a \ b \ cr a b = (card {x. a \ x \ x \ b \ poly p x = 0})" "nr a = card {x. x \ a \ poly p x = 0}" using count_roots_interval_sf[OF _ cr[unfolded count_roots_interval_iff[OF sf]]] sf[unfolded square_free_def] by blast+ definition root_cond :: "int poly \ rat \ rat \ real \ bool" where "root_cond plr x = (case plr of (p,l,r) \ of_rat l \ x \ x \ of_rat r \ ipoly p x = 0)" definition root_info_cond :: "root_info \ int poly \ bool" where "root_info_cond ri p \ (\ a b. a \ b \ root_info.l_r ri a b = card {x. root_cond (p,a,b) x}) \ (\ a. root_info.number_root ri a = card {x. x \ real_of_rat a \ ipoly p x = 0})" lemma root_info_condD: "root_info_cond ri p \ a \ b \ root_info.l_r ri a b = card {x. root_cond (p,a,b) x}" "root_info_cond ri p \ root_info.number_root ri a = card {x. x \ real_of_rat a \ ipoly p x = 0}" unfolding root_info_cond_def by auto definition count_roots_interval_sf_rat :: "int poly \ root_info" where "count_roots_interval_sf_rat p = (let pp = real_of_int_poly p; (cr,nr) = count_roots_interval_sf pp in Root_Info (\ a b. cr (of_rat a) (of_rat b)) (\ a. nr (of_rat a)))" definition count_roots_interval_rat :: "int poly \ root_info" where [code del]: "count_roots_interval_rat p = (let pp = real_of_int_poly p; (cr,nr) = count_roots_interval pp in Root_Info (\ a b. cr (of_rat a) (of_rat b)) (\ a. nr (of_rat a)))" definition count_roots_rat :: "int poly \ nat" where [code del]: "count_roots_rat p = (count_roots (real_of_int_poly p))" lemma count_roots_interval_sf_rat: assumes p: "p \ 0" shows "root_info_cond (count_roots_interval_sf_rat p) p" proof - let ?p = "real_of_int_poly p" let ?r = real_of_rat let ?ri = "count_roots_interval_sf_rat p" from p have p: "?p \ 0" by auto obtain cr nr where cr: "count_roots_interval_sf ?p = (cr,nr)" by force have "?ri = Root_Info (\a b. cr (?r a) (?r b)) (\a. nr (?r a))" unfolding count_roots_interval_sf_rat_def Let_def cr by auto hence id: "root_info.l_r ?ri = (\a b. cr (?r a) (?r b))" "root_info.number_root ?ri = (\a. nr (?r a))" by auto note cr = count_roots_interval_sf[OF p cr] show ?thesis unfolding root_info_cond_def id proof (intro conjI impI allI) fix a show "nr (?r a) = card {x. x \ (?r a) \ ipoly p x = 0}" using cr(2)[of "?r a"] by simp next fix a b :: rat assume ab: "a \ b" from ab have ab: "?r a \ ?r b" by (simp add: of_rat_less_eq) from cr(1)[OF this] show "cr (?r a) (?r b) = card (Collect (root_cond (p, a, b)))" unfolding root_cond_def[abs_def] split by simp qed qed lemma of_rat_of_int_poly: "map_poly of_rat (of_int_poly p) = of_int_poly p" by (subst map_poly_map_poly, auto simp: o_def) lemma square_free_of_int_poly: assumes "square_free p" - shows "square_free (of_int_poly p :: 'a :: {euclidean_ring_gcd, field_char_0} poly)" + shows "square_free (of_int_poly p :: 'a :: {field_gcd, field_char_0} poly)" proof - have "square_free (map_poly of_rat (of_int_poly p) :: 'a poly)" unfolding of_rat_hom.square_free_map_poly by (rule square_free_int_rat[OF assms]) thus ?thesis unfolding of_rat_of_int_poly . qed lemma count_roots_interval_rat: assumes sf: "square_free p" shows "root_info_cond (count_roots_interval_rat p) p" proof - from sf have sf: "square_free (real_of_int_poly p)" by (rule square_free_of_int_poly) from sf have p: "p \ 0" unfolding square_free_def by auto show ?thesis using count_roots_interval_sf_rat[OF p] unfolding count_roots_interval_rat_def count_roots_interval_sf_rat_def Let_def count_roots_interval_iff[OF sf] . qed lemma count_roots_rat: "count_roots_rat p = card {x. ipoly p x = (0 :: real)}" unfolding count_roots_rat_def count_roots_correct .. subsection \Implementing Sturm on Rational Polynomials\ function sturm_aux_rat where "sturm_aux_rat (p :: rat poly) q = (if degree q = 0 then [p,q] else p # sturm_aux_rat q (-(p mod q)))" by (pat_completeness, simp_all) termination by (relation "measure (degree \ snd)", simp_all add: o_def degree_mod_less') lemma sturm_aux_rat: "sturm_aux (real_of_rat_poly p) (real_of_rat_poly q) = map real_of_rat_poly (sturm_aux_rat p q)" proof (induct p q rule: sturm_aux_rat.induct) case (1 p q) interpret map_poly_inj_idom_hom of_rat.. note deg = of_int_hom.degree_map_poly_hom show ?case unfolding sturm_aux.simps[of "real_of_rat_poly p"] sturm_aux_rat.simps[of p] using 1 by (cases "degree q = 0"; simp add: hom_distribs) qed definition sturm_rat where "sturm_rat p = sturm_aux_rat p (pderiv p)" lemma sturm_rat: "sturm (real_of_rat_poly p) = map real_of_rat_poly (sturm_rat p)" unfolding sturm_rat_def sturm_def apply (fold of_rat_hom.map_poly_pderiv) unfolding sturm_aux_rat.. definition poly_number_rootat :: "rat poly \ rat" where "poly_number_rootat p \ sgn (coeff p (degree p))" definition poly_neg_number_rootat :: "rat poly \ rat" where "poly_neg_number_rootat p \ if even (degree p) then sgn (coeff p (degree p)) else -sgn (coeff p (degree p))" lemma poly_number_rootat: "poly_inf (real_of_rat_poly p) = real_of_rat (poly_number_rootat p)" unfolding poly_inf_def poly_number_rootat_def of_int_hom.degree_map_poly_hom of_rat_hom.coeff_map_poly_hom real_of_rat_sgn by simp lemma poly_neg_number_rootat: "poly_neg_inf (real_of_rat_poly p) = real_of_rat (poly_neg_number_rootat p)" unfolding poly_neg_inf_def poly_neg_number_rootat_def of_int_hom.degree_map_poly_hom of_rat_hom.coeff_map_poly_hom real_of_rat_sgn by (simp add:hom_distribs) definition sign_changes_rat where "sign_changes_rat ps (x::rat) = length (remdups_adj (filter (\x. x \ 0) (map (\p. sgn (poly p x)) ps))) - 1" definition sign_changes_number_rootat where "sign_changes_number_rootat ps = length (remdups_adj (filter (\x. x \ 0) (map poly_number_rootat ps))) - 1" definition sign_changes_neg_number_rootat where "sign_changes_neg_number_rootat ps = length (remdups_adj (filter (\x. x \ 0) (map poly_neg_number_rootat ps))) - 1" lemma real_of_rat_list_neq: "list_neq (map real_of_rat xs) 0 = map real_of_rat (list_neq xs 0)" by (induct xs, auto) lemma real_of_rat_remdups_adj: "remdups_adj (map real_of_rat xs) = map real_of_rat (remdups_adj xs)" by (induct xs rule: remdups_adj.induct, auto) lemma sign_changes_rat: "sign_changes (map real_of_rat_poly ps) (real_of_rat x) = sign_changes_rat ps x" (is "?l = ?r") proof - define xs where "xs = list_neq (map (\p. sgn (poly p x)) ps) 0" have "?l = length (remdups_adj (list_neq (map real_of_rat (map (\xa. (sgn (poly xa x))) ps)) 0)) - 1" by (simp add: sign_changes_def real_of_rat_sgn o_def) also have "\ = ?r" unfolding sign_changes_rat_def real_of_rat_list_neq unfolding real_of_rat_remdups_adj by simp finally show ?thesis . qed lemma sign_changes_neg_number_rootat: "sign_changes_neg_inf (map real_of_rat_poly ps) = sign_changes_neg_number_rootat ps" (is "?l = ?r") proof - have "?l = length (remdups_adj (list_neq (map real_of_rat (map poly_neg_number_rootat ps)) 0)) - 1" by (simp add: sign_changes_neg_inf_def o_def real_of_rat_sgn poly_neg_number_rootat) also have "\ = ?r" unfolding sign_changes_neg_number_rootat_def real_of_rat_list_neq unfolding real_of_rat_remdups_adj by simp finally show ?thesis . qed lemma sign_changes_number_rootat: "sign_changes_inf (map real_of_rat_poly ps) = sign_changes_number_rootat ps" (is "?l = ?r") proof - have "?l = length (remdups_adj (list_neq (map real_of_rat (map poly_number_rootat ps)) 0)) - 1" unfolding sign_changes_inf_def unfolding map_map o_def real_of_rat_sgn poly_number_rootat .. also have "\ = ?r" unfolding sign_changes_number_rootat_def real_of_rat_list_neq unfolding real_of_rat_remdups_adj by simp finally show ?thesis . qed lemma count_roots_interval_rat_code[code]: "count_roots_interval_rat p = (let rp = map_poly rat_of_int p; ps = sturm_rat rp in Root_Info (\ a b. sign_changes_rat ps a - sign_changes_rat ps b + (if poly rp a = 0 then 1 else 0)) (\ a. sign_changes_neg_number_rootat ps - sign_changes_rat ps a))" unfolding count_roots_interval_rat_def Let_def count_roots_interval_def split of_rat_of_int_poly[symmetric, where 'a = real] sturm_rat sign_changes_rat by (simp add: sign_changes_neg_number_rootat) lemma count_roots_rat_code[code]: "count_roots_rat p = (let rp = map_poly rat_of_int p in if p = 0 then 0 else let ps = sturm_rat rp in sign_changes_neg_number_rootat ps - sign_changes_number_rootat ps)" unfolding count_roots_rat_def Let_def sturm_rat count_roots_code of_rat_of_int_poly[symmetric, where 'a = real] sign_changes_neg_number_rootat sign_changes_number_rootat by simp hide_const (open) count_roots_interval_sf_rat text \Finally we provide an even more efficient implementation which avoids the "poly p x = 0" test, but it is restricted to irreducible polynomials.\ definition root_info :: "int poly \ root_info" where "root_info p = (if degree p = 1 then (let x = Rat.Fract (- coeff p 0) (coeff p 1) in Root_Info (\ l r. if l \ x \ x \ r then 1 else 0) (\ b. if x \ b then 1 else 0)) else (let rp = map_poly rat_of_int p; ps = sturm_rat rp in Root_Info (\ a b. sign_changes_rat ps a - sign_changes_rat ps b) (\ a. sign_changes_neg_number_rootat ps - sign_changes_rat ps a)))" lemma root_info: assumes irr: "irreducible p" and deg: "degree p > 0" shows "root_info_cond (root_info p) p" proof (cases "degree p = 1") case deg: True from degree1_coeffs[OF this] obtain a b where p: "p = [:b,a:]" and "a \ 0" by auto from deg have "degree (real_of_int_poly p) = 1" by simp from roots1[OF this, unfolded roots1_def] p have id: "(ipoly p x = 0) = ((x :: real) = - b / a)" for x by auto have idd: "{x. real_of_rat aa \ x \ x \ real_of_rat ba \ x = real_of_int (- b) / real_of_int a} = (if real_of_rat aa \ real_of_int (- b) / real_of_int a \ real_of_int (- b) / real_of_int a \ real_of_rat ba then {real_of_int (- b) / real_of_int a} else {})" for aa ba by auto have iddd: "{x. x \ real_of_rat aa \ x = real_of_int (- b) / real_of_int a} = (if real_of_int (- b) / real_of_int a \ real_of_rat aa then {real_of_int (- b) / real_of_int a} else {})" for aa by auto have id4: "real_of_int x = real_of_rat (rat_of_int x)" for x by simp show ?thesis unfolding root_info_def deg unfolding root_info_cond_def id root_cond_def split unfolding p Fract_of_int_quotient Let_def idd iddd unfolding id4 of_rat_divide[symmetric] of_rat_less_eq by auto next case False have irr_d: "irreducible\<^sub>d p" by (simp add: deg irr irreducible_connect_rev) from irreducible\<^sub>d_int_rat[OF this] have "irreducible (of_int_poly p :: rat poly)" by auto from irreducible_root_free[OF this] have idd: "(poly (of_int_poly p) a = 0) = False" for a :: rat unfolding root_free_def using False by auto have id: "root_info p = count_roots_interval_rat p" unfolding root_info_def if_False count_roots_interval_rat_code Let_def idd using False by auto show ?thesis unfolding id by (rule count_roots_interval_rat[OF irreducible\<^sub>d_square_free[OF irr_d]]) qed end diff --git a/thys/Berlekamp_Zassenhaus/Arithmetic_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Arithmetic_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Arithmetic_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Arithmetic_Record_Based.thy @@ -1,230 +1,230 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) section \Arithmetics via Records\ text \We create a locale for rings and fields based on a record that includes all the necessary operations.\ theory Arithmetic_Record_Based imports "HOL-Library.More_List" "HOL-Computational_Algebra.Euclidean_Algorithm" begin datatype 'a arith_ops_record = Arith_Ops_Record (zero : 'a) (one : 'a) (plus : "'a \ 'a \ 'a") (times : "'a \ 'a \ 'a") (minus : "'a \ 'a \ 'a") (uminus : "'a \ 'a") (divide : "'a \ 'a \ 'a") (inverse : "'a \ 'a") ("modulo" : "'a \ 'a \ 'a") (normalize : "'a \ 'a") (unit_factor : "'a \ 'a") (of_int : "int \ 'a") (to_int : "'a \ int") (DP : "'a \ bool") hide_const (open) zero one plus times minus uminus divide inverse modulo normalize unit_factor of_int to_int DP fun listprod_i :: "'i arith_ops_record \ 'i list \ 'i" where "listprod_i ops (x # xs) = arith_ops_record.times ops x (listprod_i ops xs)" | "listprod_i ops [] = arith_ops_record.one ops" locale arith_ops = fixes ops :: "'i arith_ops_record" (structure) begin abbreviation (input) zero where "zero \ arith_ops_record.zero ops" abbreviation (input) one where "one \ arith_ops_record.one ops" abbreviation (input) plus where "plus \ arith_ops_record.plus ops" abbreviation (input) times where "times \ arith_ops_record.times ops" abbreviation (input) minus where "minus \ arith_ops_record.minus ops" abbreviation (input) uminus where "uminus \ arith_ops_record.uminus ops" abbreviation (input) divide where "divide \ arith_ops_record.divide ops" abbreviation (input) inverse where "inverse \ arith_ops_record.inverse ops" abbreviation (input) modulo where "modulo \ arith_ops_record.modulo ops" abbreviation (input) normalize where "normalize \ arith_ops_record.normalize ops" abbreviation (input) unit_factor where "unit_factor \ arith_ops_record.unit_factor ops" abbreviation (input) DP where "DP \ arith_ops_record.DP ops" partial_function (tailrec) gcd_eucl_i :: "'i \ 'i \ 'i" where "gcd_eucl_i a b = (if b = zero then normalize a else gcd_eucl_i b (modulo a b))" partial_function (tailrec) euclid_ext_aux_i :: "'i \ 'i \ 'i \ 'i \ 'i \ 'i \ ('i \ 'i) \ 'i" where "euclid_ext_aux_i s' s t' t r' r = ( if r = zero then let c = divide one (unit_factor r') in ((times s' c, times t' c), normalize r') else let q = divide r' r in euclid_ext_aux_i s (minus s' (times q s)) t (minus t' (times q t)) r (modulo r' r))" abbreviation (input) euclid_ext_i :: "'i \ 'i \ ('i \ 'i) \ 'i" where "euclid_ext_i \ euclid_ext_aux_i one zero zero one" end declare arith_ops.gcd_eucl_i.simps[code] declare arith_ops.euclid_ext_aux_i.simps[code] unbundle lifting_syntax locale ring_ops = arith_ops ops for ops :: "'i arith_ops_record" + fixes R :: "'i \ 'a :: comm_ring_1 \ bool" assumes bi_unique[transfer_rule]: "bi_unique R" and right_total[transfer_rule]: "right_total R" and zero[transfer_rule]: "R zero 0" and one[transfer_rule]: "R one 1" and plus[transfer_rule]: "(R ===> R ===> R) plus (+)" and minus[transfer_rule]: "(R ===> R ===> R) minus (-)" and uminus[transfer_rule]: "(R ===> R) uminus Groups.uminus" and times[transfer_rule]: "(R ===> R ===> R) times ((*))" and eq[transfer_rule]: "(R ===> R ===> (=)) (=) (=)" and DPR[transfer_domain_rule]: "Domainp R = DP" begin lemma left_right_unique[transfer_rule]: "left_unique R" "right_unique R" using bi_unique unfolding bi_unique_def left_unique_def right_unique_def by auto lemma listprod_i[transfer_rule]: "(list_all2 R ===> R) (listprod_i ops) prod_list" proof (intro rel_funI, goal_cases) case (1 xs ys) thus ?case proof (induct xs ys rule: list_all2_induct) case (Cons x xs y ys) note [transfer_rule] = this show ?case by simp transfer_prover qed (simp add: one) qed end locale idom_ops = ring_ops ops R for ops :: "'i arith_ops_record" and R :: "'i \ 'a :: idom \ bool" locale idom_divide_ops = idom_ops ops R for ops :: "'i arith_ops_record" and R :: "'i \ 'a :: idom_divide \ bool" + assumes divide[transfer_rule]: "(R ===> R ===> R) divide Rings.divide" locale euclidean_semiring_ops = idom_ops ops R for ops :: "'i arith_ops_record" and R :: "'i \ 'a :: {idom,normalization_euclidean_semiring} \ bool" + assumes modulo[transfer_rule]: "(R ===> R ===> R) modulo (mod)" and normalize[transfer_rule]: "(R ===> R) normalize Rings.normalize" and unit_factor[transfer_rule]: "(R ===> R) unit_factor Rings.unit_factor" begin lemma gcd_eucl_i [transfer_rule]: "(R ===> R ===> R) gcd_eucl_i Euclidean_Algorithm.gcd" proof (intro rel_funI, goal_cases) case (1 x X y Y) thus ?case proof (induct X Y arbitrary: x y rule: Euclidean_Algorithm.gcd.induct) case (1 X Y x y) note [transfer_rule] = 1(2-) note simps = gcd_eucl_i.simps[of x y] Euclidean_Algorithm.gcd.simps[of X Y] have eq: "(y = zero) = (Y = 0)" by transfer_prover show ?case proof (cases "Y = 0") case True hence *: "y = zero" using eq by simp have "R (normalize x) (Rings.normalize X)" by transfer_prover thus ?thesis unfolding simps unfolding True * by simp next case False with eq have yz: "y \ zero" by simp have "R (gcd_eucl_i y (modulo x y)) (Euclidean_Algorithm.gcd Y (X mod Y))" by (rule 1(1)[OF False], transfer_prover+) thus ?thesis unfolding simps using False yz by simp qed qed qed end locale euclidean_ring_ops = euclidean_semiring_ops ops R for ops :: "'i arith_ops_record" and R :: "'i \ 'a :: {idom,euclidean_ring_gcd} \ bool" + assumes divide[transfer_rule]: "(R ===> R ===> R) divide (div)" begin lemma euclid_ext_aux_i[transfer_rule]: "(R ===> R ===> R ===> R ===> R ===> R ===> rel_prod (rel_prod R R) R) euclid_ext_aux_i euclid_ext_aux" proof (intro rel_funI, goal_cases) case (1 z Z a A b B c C x X y Y) thus ?case proof (induct Z A B C X Y arbitrary: z a b c x y rule: euclid_ext_aux.induct) case (1 Z A B C X Y z a b c x y) note [transfer_rule] = 1(2-) note simps = euclid_ext_aux_i.simps[of z a b c x y] euclid_ext_aux.simps[of Z A B C X Y] have eq: "(y = zero) = (Y = 0)" by transfer_prover show ?case proof (cases "Y = 0") case True hence *: "(y = zero) = True" "(Y = 0) = True" using eq by auto show ?thesis unfolding simps unfolding * if_True by transfer_prover next case False hence *: "(y = zero) = False" "(Y = 0) = False" using eq by auto have XY: "R (modulo x y) (X mod Y)" by transfer_prover have YA: "R (minus z (times (divide x y) a)) (Z - X div Y * A)" by transfer_prover have YC: "R (minus b (times (divide x y) c)) (B - X div Y * C)" by transfer_prover note [transfer_rule] = 1(1)[OF False refl 1(3) YA 1(5) YC 1(7) XY] show ?thesis unfolding simps * if_False Let_def by transfer_prover qed qed qed lemma euclid_ext_i [transfer_rule]: "(R ===> R ===> rel_prod (rel_prod R R) R) euclid_ext_i euclid_ext" by transfer_prover end locale field_ops = idom_divide_ops ops R + euclidean_semiring_ops ops R for ops :: "'i arith_ops_record" and - R :: "'i \ 'a :: {field, normalization_euclidean_semiring, euclidean_ring, factorial_ring_gcd} \ bool" + + R :: "'i \ 'a :: {field_gcd} \ bool" + assumes inverse[transfer_rule]: "(R ===> R) inverse Fields.inverse" lemma nth_default_rel[transfer_rule]: "(S ===> list_all2 S ===> (=) ===> S) nth_default nth_default" proof (intro rel_funI, clarify, goal_cases) case (1 x y xs ys _ n) from 1(2) show ?case proof (induct arbitrary: n) case Nil thus ?case using 1(1) by simp next case (Cons x y xs ys n) thus ?case by (cases n, auto) qed qed lemma strip_while_rel[transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) strip_while strip_while" unfolding strip_while_def[abs_def] by transfer_prover lemma list_all2_last[simp]: "list_all2 A (xs @ [x]) (ys @ [y]) \ list_all2 A xs ys \ A x y" proof (cases "length xs = length ys") case True show ?thesis by (simp add: list_all2_append[OF True]) next case False note len = list_all2_lengthD[of A] from len[of xs ys] len[of "xs @ [x]" "ys @ [y]"] False show ?thesis by auto qed end diff --git a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy --- a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy @@ -1,3144 +1,3145 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) section \The Berlekamp Algorithm\ theory Berlekamp_Type_Based imports Jordan_Normal_Form.Matrix_Kernel Jordan_Normal_Form.Gauss_Jordan_Elimination Jordan_Normal_Form.Missing_VectorSpace Polynomial_Factorization.Square_Free_Factorization Polynomial_Factorization.Missing_Multiset Finite_Field Chinese_Remainder_Poly Poly_Mod_Finite_Field "HOL-Computational_Algebra.Field_as_Ring" begin hide_const (open) up_ring.coeff up_ring.monom Modules.module subspace Modules.module_hom subsection \Auxiliary lemmas\ context fixes g :: "'b \ 'a :: comm_monoid_mult" begin lemma prod_list_map_filter: "prod_list (map g (filter f xs)) * prod_list (map g (filter (\ x. \ f x) xs)) = prod_list (map g xs)" by (induct xs, auto simp: ac_simps) lemma prod_list_map_partition: assumes "List.partition f xs = (ys, zs)" shows "prod_list (map g xs) = prod_list (map g ys) * prod_list (map g zs)" using assms by (subst prod_list_map_filter[symmetric, of _ f], auto simp: o_def) end lemma coprime_id_is_unit: fixes a::"'b::semiring_gcd" shows "coprime a a \ is_unit a" using dvd_unit_imp_unit by auto lemma dim_vec_of_list[simp]: "dim_vec (vec_of_list x) = length x" by (transfer, auto) lemma length_list_of_vec[simp]: "length (list_of_vec A) = dim_vec A" by (transfer', auto) lemma list_of_vec_vec_of_list[simp]: "list_of_vec (vec_of_list a) = a" proof - { fix aa :: "'a list" have "map (\n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0..n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0.. list_of_vec) {v. dim_vec v = n}" proof (rule comp_inj_on) show "inj_on list_of_vec {v. dim_vec v = n}" by (auto simp add: inj_on_def, transfer, auto simp add: mk_vec_def) show "inj_on Poly (list_of_vec ` {v. dim_vec v = n})" proof (auto simp add: inj_on_def) fix x y::"'c vec" assume "n = dim_vec x" and dim_xy: "dim_vec y = dim_vec x" and Poly_eq: "Poly (list_of_vec x) = Poly (list_of_vec y)" note [simp del] = nth_list_of_vec show "list_of_vec x = list_of_vec y" proof (rule nth_equalityI, auto simp: dim_xy) have length_eq: "length (list_of_vec x ) = length (list_of_vec y)" using dim_xy by (transfer, auto) fix i assume "i < dim_vec x" thus "list_of_vec x ! i = list_of_vec y ! i" using Poly_eq unfolding poly_eq_iff coeff_Poly_eq using dim_xy unfolding nth_default_def by (auto, presburger) qed qed qed corollary inj_Poly_list_of_vec: "inj_on (Poly \ list_of_vec) (carrier_vec n)" using inj_Poly_list_of_vec' unfolding carrier_vec_def . lemma list_of_vec_rw_map: "list_of_vec m = map (\n. m $ n) [0.. []" shows "degree (Poly xs) < length xs" using xs by (induct xs, auto intro: Poly.simps(1)) lemma vec_of_list_list_of_vec[simp]: "vec_of_list (list_of_vec a) = a" by (transfer, auto simp add: mk_vec_def) lemma row_mat_of_rows_list: assumes b: "b < length A" and nc: "\i. i < length A \ length (A ! i) = nc" shows "(row (mat_of_rows_list nc A) b) = vec_of_list (A ! b)" proof (auto simp add: vec_eq_iff) show "dim_col (mat_of_rows_list nc A) = length (A ! b)" unfolding mat_of_rows_list_def using b nc by auto fix i assume i: "i < length (A ! b)" show "row (mat_of_rows_list nc A) b $ i = vec_of_list (A ! b) $ i" using i b nc unfolding mat_of_rows_list_def row_def by (transfer, auto simp add: mk_vec_def mk_mat_def) qed lemma degree_Poly_list_of_vec: assumes n: "x \ carrier_vec n" and n0: "n > 0" shows "degree (Poly (list_of_vec x)) < n" proof - have x_dim: "dim_vec x = n" using n by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] n0 n x_dim) have "degree (Poly (list_of_vec x)) < length (list_of_vec x)" by (rule degree_Poly'[OF l]) also have "... = n" using x_dim by auto finally show ?thesis . qed lemma list_of_vec_nth: assumes i: "i < dim_vec x" shows "list_of_vec x ! i = x $ i" using i by (transfer, auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth': assumes i: "i < dim_vec x" shows "coeff (Poly (list_of_vec x)) i = x $ i" using i by (auto simp add: list_of_vec_nth nth_default_def) lemma list_of_vec_row_nth: assumes x: "x < dim_col A" shows "list_of_vec (row A i) ! x = A $$ (i, x)" using x unfolding row_def by (transfer', auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth: assumes x: "x < dim_col A" shows "coeff (Poly (list_of_vec (row A i))) x = A $$ (i, x)" proof - have "coeff (Poly (list_of_vec (row A i))) x = nth_default 0 (list_of_vec (row A i)) x" unfolding coeff_Poly_eq by simp also have "... = A $$ (i, x)" using x list_of_vec_row_nth unfolding nth_default_def by (auto simp del: nth_list_of_vec) finally show ?thesis . qed lemma inj_on_list_of_vec: "inj_on list_of_vec (carrier_vec n)" unfolding inj_on_def unfolding list_of_vec_rw_map by auto lemma vec_of_list_carrier[simp]: "vec_of_list x \ carrier_vec (length x)" unfolding carrier_vec_def by simp lemma card_carrier_vec: "card (carrier_vec n:: 'b::finite vec set) = CARD('b) ^ n" proof - let ?A = "UNIV::'b set" let ?B = "{xs. set xs \ ?A \ length xs = n}" let ?C = "(carrier_vec n:: 'b::finite vec set)" have "card ?C = card ?B" proof - have "bij_betw (list_of_vec) ?C ?B" proof (unfold bij_betw_def, auto) show "inj_on list_of_vec (carrier_vec n)" by (rule inj_on_list_of_vec) fix x::"'b list" assume n: "n = length x" thus "x \ list_of_vec ` carrier_vec (length x)" unfolding image_def by auto (rule bexI[of _ "vec_of_list x"], auto) qed thus ?thesis using bij_betw_same_card by blast qed also have "... = card ?A ^ n" by (rule card_lists_length_eq, simp) finally show ?thesis . qed lemma finite_carrier_vec[simp]: "finite (carrier_vec n:: 'b::finite vec set)" by (rule card_ge_0_finite, unfold card_carrier_vec, auto) lemma row_echelon_form_dim0_row: assumes "A \ carrier_mat 0 n" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_dim0_col: assumes "A \ carrier_mat n 0" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_one_dim0[simp]: "row_echelon_form (1\<^sub>m 0)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma Poly_list_of_vec_0[simp]: "Poly (list_of_vec (0\<^sub>v 0)) = [:0:]" by (simp add: poly_eq_iff nth_default_def) lemma monic_normalize: assumes "(p :: 'b :: {field,euclidean_ring_gcd} poly) \ 0" shows "monic (normalize p)" by (simp add: assms normalize_poly_old_def) lemma exists_factorization_prod_list: fixes P::"'b::field poly list" assumes "degree (prod_list P) > 0" and "\ u. u \ set P \ degree u > 0 \ monic u" and "square_free (prod_list P)" shows "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" using assms proof (induct P) case Nil thus ?case by auto next case (Cons x P) have sf_P: "square_free (prod_list P)" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons mult.commute square_free_factor) have deg_x: "degree x > 0" using Cons.prems by auto have distinct_P: "distinct P" by (meson Cons.prems(2) Cons.prems(3) distinct.simps(2) square_free_prod_list_distinct) have "\A. finite A \ x = \A \ A \ {q. irreducible q \ monic q}" proof (rule monic_square_free_irreducible_factorization) show "monic x" by (simp add: Cons.prems(2)) show "square_free x" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons square_free_factor) qed from this obtain A where fin_A: "finite A" and xA: "x = \A" and A: "A \ {q. irreducible\<^sub>d q \ monic q}" by auto obtain A' where s: "set A' = A" and length_A': "length A' = card A" using \finite A\ distinct_card finite_distinct_list by force have A_not_empty: "A \ {}" using xA deg_x by auto have x_prod_list_A': "x = prod_list A'" proof - have "x = \A" using xA by simp also have "... = prod id A" by simp also have "... = prod id (set A')" unfolding s by simp also have "... = prod_list (map id A')" by (rule prod.distinct_set_conv_list, simp add: card_distinct length_A' s) also have "... = prod_list A'" by auto finally show ?thesis . qed show ?case proof (cases "P = []") case True show ?thesis proof (rule exI[of _ "A'"], auto simp add: True) show "prod_list A' = x" using x_prod_list_A' by simp show "Suc 0 \ length A'" using A_not_empty using s length_A' by (simp add: Suc_leI card_gt_0_iff fin_A) show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed next case False have hyp: "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule Cons.hyps[OF _ _ sf_P]) have set_P: "set P \ {}" using False by auto have "prod_list P = prod_list (map id P)" by simp also have "... = prod id (set P)" using prod.distinct_set_conv_list[OF distinct_P, of id] by simp also have "... = \(set P)" by simp finally have "prod_list P = \(set P)" . hence "degree (prod_list P) = degree (\(set P))" by simp also have "... = degree (prod id (set P))" by simp also have "... = (\i\(set P). degree (id i))" proof (rule degree_prod_eq_sum_degree) show "\i\set P. id i \ 0" using Cons.prems(2) by force qed also have "... > 0" by (metis Cons.prems(2) List.finite_set set_P gr0I id_apply insert_iff list.set(2) sum_pos) finally show "degree (prod_list P) > 0" by simp show "\u. u \ set P \ degree u > 0 \ monic u" using Cons.prems by auto qed from this obtain Q where QP: "prod_list Q = prod_list P" and length_PQ: "length P \ length Q" and monic_irr_Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast show ?thesis proof (rule exI[of _ "A' @ Q"], auto simp add: monic_irr_Q) show "prod_list A' * prod_list Q = x * prod_list P" unfolding QP x_prod_list_A' by auto have "length A' \ 0" using A_not_empty using s length_A' by auto thus "Suc (length P) \ length A' + length Q" using QP length_PQ by linarith show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed qed qed lemma normalize_eq_imp_smult: fixes p :: "'b :: {euclidean_ring_gcd} poly" assumes n: "normalize p = normalize q" shows "\ c. c \ 0 \ q = smult c p" proof(cases "p = 0") case True with n show ?thesis by (auto intro:exI[of _ 1]) next case p0: False have degree_eq: "degree p = degree q" using n degree_normalize by metis hence q0: "q \ 0" using p0 n by auto have p_dvd_q: "p dvd q" using n by (simp add: associatedD1) from p_dvd_q obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with q0 have "k \ 0" by auto then have "degree k = 0" - using associatedD2 n p0 q - by (metis (no_types, lifting) mult_cancel_right1 normalize_eq_0_iff normalize_mult poly_dvd_1) + using degree_eq degree_mult_eq p0 q by fastforce then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed - -lemma prod_list_normalize: "normalize (prod_list P) = prod_list (map normalize P)" +lemma prod_list_normalize: + fixes P :: "'b :: {idom_divide,normalization_semidom_multiplicative} poly list" + shows "normalize (prod_list P) = prod_list (map normalize P)" proof (induct P) case Nil show ?case by auto next case (Cons p P) have "normalize (prod_list (p # P)) = normalize p * normalize (prod_list P)" using normalize_mult by auto also have "... = normalize p * prod_list (map normalize P)" using Cons.hyps by auto also have "... = prod_list (normalize p # (map normalize P))" by auto also have "... = prod_list (map normalize (p # P))" by auto finally show ?case . qed lemma prod_list_dvd_prod_list_subset: fixes A::"'b::comm_monoid_mult list" assumes dA: "distinct A" and dB: "distinct B" (*Maybe this condition could be avoided*) and s: "set A \ set B" shows "prod_list A dvd prod_list B" proof - have "prod_list A = prod_list (map id A)" by auto also have "... = prod id (set A)" by (rule prod.distinct_set_conv_list[symmetric, OF dA]) also have "... dvd prod id (set B)" by (rule prod_dvd_prod_subset[OF _ s], auto) also have "... = prod_list (map id B)" by (rule prod.distinct_set_conv_list[OF dB]) also have "... = prod_list B" by simp finally show ?thesis . qed end lemma gcd_monic_constant: "gcd f g \ {1, f}" if "monic f" and "degree g = 0" - for f g :: "'a :: {field,euclidean_ring_gcd} poly" + for f g :: "'a :: {field_gcd} poly" proof (cases "g = 0") case True moreover from \monic f\ have "normalize f = f" by (rule normalize_monic) ultimately show ?thesis by simp next case False with \degree g = 0\ have "is_unit g" by simp then have "Rings.coprime f g" by (rule is_unit_right_imp_coprime) then show ?thesis by simp qed lemma distinct_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "distinct (find_base_vectors A)" proof - note non_pivot_base = non_pivot_base[OF ref A] let ?pp = "set (pivot_positions A)" from A have dim: "dim_row A = nr" "dim_col A = nc" by auto { fix j j' assume j: "j < nc" "j \ snd ` ?pp" and j': "j' < nc" "j' \ snd ` ?pp" and neq: "j' \ j" from non_pivot_base(2)[OF j] non_pivot_base(4)[OF j' j neq] have "non_pivot_base A (pivot_positions A) j \ non_pivot_base A (pivot_positions A) j'" by auto } hence inj: "inj_on (non_pivot_base A (pivot_positions A)) (set [j\[0.. snd ` ?pp])" unfolding inj_on_def by auto thus ?thesis unfolding find_base_vectors_def Let_def unfolding distinct_map dim by auto qed lemma length_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "length (find_base_vectors A) = card (set (find_base_vectors A))" using distinct_card[OF distinct_find_base_vectors[OF ref A]] by auto subsection \Previous Results\ definition power_poly_f_mod :: "'a::field poly \ 'a poly \ nat \ 'a poly" where "power_poly_f_mod modulus = (\a n. a ^ n mod modulus)" lemma power_poly_f_mod_binary: "power_poly_f_mod m a n = (if n = 0 then 1 mod m else let (d, r) = Divides.divmod_nat n 2; rec = power_poly_f_mod m ((a * a) mod m) d in if r = 0 then rec else (rec * a) mod m)" - for m a :: "'a :: {factorial_ring_gcd, field} poly" + for m a :: "'a :: {field_gcd} poly" proof - note d = power_poly_f_mod_def show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding d by simp next case False obtain q r where div: "Divides.divmod_nat n 2 = (q,r)" by force hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding divmod_nat_def by auto have id: "a ^ (2 * q) = (a * a) ^ q" by (simp add: power_mult_distrib semiring_normalization_rules) show ?thesis proof (cases "r = 0") case True show ?thesis using power_mod [of "a * a" m q] by (auto simp add: divmod_nat_def Let_def True n d div id) next case False with r have r: "r = 1" by simp show ?thesis by (auto simp add: d r div Let_def mod_simps) (simp add: n r mod_simps ac_simps power_mult_distrib power_mult power2_eq_square) qed qed qed fun power_polys where "power_polys mul_p u curr_p (Suc i) = curr_p # power_polys mul_p u ((curr_p * mul_p) mod u) i" | "power_polys mul_p u curr_p 0 = []" context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma fermat_theorem_mod_ring [simp]: fixes a::"'a mod_ring" shows "a ^ CARD('a) = a" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis proof transfer fix a assume "a \ {0.. 0" then have a: "1 \ a" "a < int CARD('a)" by simp_all then have [simp]: "a mod int CARD('a) = a" by simp from a have "\ int CARD('a) dvd a" by (auto simp add: zdvd_not_zless) then have "\ CARD('a) dvd nat \a\" by simp with a have "\ CARD('a) dvd nat a" by simp with prime_card have "[nat a ^ (CARD('a) - 1) = 1] (mod CARD('a))" by (rule fermat_theorem) with a have "int (nat a ^ (CARD('a) - 1) mod CARD('a)) = 1" by (simp add: cong_def) with a have "a ^ (CARD('a) - 1) mod CARD('a) = 1" by (simp add: of_nat_mod) then have "a * (a ^ (CARD('a) - 1) mod int CARD('a)) = a" by simp then have "(a * (a ^ (CARD('a) - 1) mod int CARD('a))) mod int CARD('a) = a mod int CARD('a)" by (simp only:) then show "a ^ CARD('a) mod int CARD('a) = a" by (simp add: mod_simps semiring_normalization_rules(27)) qed qed lemma mod_eq_dvd_iff_poly: "((x::'a mod_ring poly) mod n = y mod n) = (n dvd x - y)" proof assume H: "x mod n = y mod n" hence "x mod n - y mod n = 0" by simp hence "(x mod n - y mod n) mod n = 0" by simp hence "(x - y) mod n = 0" by (simp add: mod_diff_eq) thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) next assume H: "n dvd x - y" then obtain k where k: "x-y = n*k" unfolding dvd_def by blast hence "x = n*k + y" using diff_eq_eq by blast hence "x mod n = (n*k + y) mod n" by simp thus "x mod n = y mod n" by (simp add: mod_add_left_eq) qed lemma cong_gcd_eq_poly: "gcd a m = gcd b m" if "[(a::'a mod_ring poly) = b] (mod m)" using that by (simp add: cong_def) (metis gcd_mod_left mod_by_0) lemma coprime_h_c_poly: fixes h::"'a mod_ring poly" assumes "c1 \ c2" shows "coprime (h - [:c1:]) (h - [:c2:])" proof (intro coprimeI) fix d assume "d dvd h - [:c1:]" and "d dvd h - [:c2:]" hence "h mod d = [:c1:] mod d" and "h mod d = [:c2:] mod d" using mod_eq_dvd_iff_poly by simp+ hence "[:c1:] mod d = [:c2:] mod d" by simp hence "d dvd [:c2 - c1:]" by (metis (no_types) mod_eq_dvd_iff_poly diff_pCons right_minus_eq) thus "is_unit d" by (metis (no_types) assms dvd_trans is_unit_monom_0 monom_0 right_minus_eq) qed lemma coprime_h_c_poly2: fixes h::"'a mod_ring poly" assumes "coprime (h - [:c1:]) (h - [:c2:])" and "\ is_unit (h - [:c1:])" shows "c1 \ c2" using assms coprime_id_is_unit by blast lemma degree_minus_eq_right: fixes p::"'b::ab_group_add poly" shows "degree q < degree p \ degree (p - q) = degree p" using degree_add_eq_left[of "-q" p] degree_minus by auto lemma coprime_prod: fixes A::"'a mod_ring set" and g::"'a mod_ring \ 'a mod_ring poly" assumes "\x\A. coprime (g a) (g x)" shows "coprime (g a) (prod (\x. g x) A)" proof - have f: "finite A" by simp show ?thesis using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = (g x) * (\c\A. g c)" by (simp add: insert.hyps(2)) with insert.prems show ?case by (auto simp: insert.hyps(3) prod_coprime_right) qed auto qed lemma coprime_prod2: fixes A::"'b::semiring_gcd set" assumes "\x\A. coprime (a) (x)" and f: "finite A" shows "coprime (a) (prod (\x. x) A)" using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. c) = (x) * (\c\A. c)" by (simp add: insert.hyps) with insert.prems show ?case by (simp add: insert.hyps prod_coprime_right) qed auto lemma divides_prod: fixes g::"'a mod_ring \ 'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2 \ coprime (g c1) (g c2)" assumes "\c\ A. g c dvd f" shows "(\c\A. g c) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = g x * (\c\ A. g c)" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "g x dvd f" using insert.prems by auto show "prod g A dvd f" using insert.hyps(3) insert.prems by auto from insert show "Rings.coprime (g x) (prod g A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto qed (* Proof of equation 9 in the book by Knuth x^p - x = (x-0)(x-1)...(x-(p-1)) (mod p) *) lemma poly_monom_identity_mod_p: "monom (1::'a mod_ring) (CARD('a)) - monom 1 1 = prod (\x. [:0,1:] - [:x:]) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof - let ?f="(\x::'a mod_ring. [:0,1:] - [:x:])" have "?rhs dvd ?lhs" proof (rule divides_prod) { fix a::"'a mod_ring" have "poly ?lhs a = 0" by (simp add: poly_monom) hence "([:0,1:] - [:a:]) dvd ?lhs" using poly_eq_0_iff_dvd by fastforce } thus "\x\UNIV::'a mod_ring set. [:0, 1:] - [:x:] dvd monom 1 CARD('a) - monom 1 1" by fast show "\c1 c2. c1 \ UNIV \ c2 \ UNIV \ c1 \ (c2 :: 'a mod_ring) \ coprime ([:0, 1:] - [:c1:]) ([:0, 1:] - [:c2:])" by (auto dest!: coprime_h_c_poly[of _ _ "[:0,1:]"]) qed from this obtain g where g: "?lhs = ?rhs * g" using dvdE by blast have degree_lhs_card: "degree ?lhs = CARD('a)" proof - have "degree (monom (1::'a mod_ring) 1) = 1" by (simp add: degree_monom_eq) moreover have d_c: "degree (monom (1::'a mod_ring) CARD('a)) = CARD('a)" by (simp add: degree_monom_eq) ultimately have "degree (monom (1::'a mod_ring) 1) < degree (monom (1::'a mod_ring) CARD('a))" using prime_card unfolding prime_nat_iff by auto hence "degree ?lhs = degree (monom (1::'a mod_ring) CARD('a))" by (rule degree_minus_eq_right) thus ?thesis unfolding d_c . qed have degree_rhs_card: "degree ?rhs = CARD('a)" proof - have "degree (prod ?f UNIV) = sum (degree \ ?f) UNIV \ coeff (prod ?f UNIV) (sum (degree \ ?f) UNIV) = 1" by (rule degree_prod_sum_monic, auto) moreover have "sum (degree \ ?f) UNIV = CARD('a)" by auto ultimately show ?thesis by presburger qed have monic_lhs: "monic ?lhs" using degree_lhs_card by auto have monic_rhs: "monic ?rhs" by (rule monic_prod, simp) have degree_eq: "degree ?rhs = degree ?lhs" unfolding degree_lhs_card degree_rhs_card .. have g_not_0: "g \ 0" using g monic_lhs by auto have degree_g0: "degree g = 0" proof - have "degree (?rhs * g) = degree ?rhs + degree g" by (rule degree_monic_mult[OF monic_rhs g_not_0]) thus ?thesis using degree_eq g by simp qed have monic_g: "monic g" using monic_factor g monic_lhs monic_rhs by auto have "g = 1" using monic_degree_0[OF monic_g] degree_g0 by simp thus ?thesis using g by auto qed (* Proof of equation 10 in the book by Knuth v(x)^p - v(x) = (v(x)-0)(v(x)-1)...(v(x)-(p-1)) (mod p) *) lemma poly_identity_mod_p: "v^(CARD('a)) - v = prod (\x. v - [:x:]) (UNIV::'a mod_ring set)" proof - have id: "monom 1 1 \\<^sub>p v = v" "[:0, 1:] \\<^sub>p v = v" unfolding pcompose_def apply (auto) by (simp add: fold_coeffs_def) have id2: "monom 1 (CARD('a)) \\<^sub>p v = v ^ (CARD('a))" by (metis id(1) pcompose_hom.hom_power x_pow_n) show ?thesis using arg_cong[OF poly_monom_identity_mod_p, of "\ f. f \\<^sub>p v"] unfolding pcompose_hom.hom_minus pcompose_hom.hom_prod id pcompose_const id2 . qed lemma coprime_gcd: fixes h::"'a mod_ring poly" assumes "Rings.coprime (h-[:c1:]) (h-[:c2:])" shows "Rings.coprime (gcd f(h-[:c1:])) (gcd f (h-[:c2:]))" using assms coprime_divisors by blast lemma divides_prod_gcd: fixes h::"'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2\ coprime (h-[:c1:]) (h-[:c2:])" shows "(\c\A. gcd f (h - [:c:])) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "gcd f (h - [:x:]) dvd f" by simp show "(\c\A. gcd f (h - [:c:])) dvd f" using insert.hyps(3) insert.prems by auto show "Rings.coprime (gcd f (h - [:x:])) (\c\A. gcd f (h - [:c:]))" by (rule prod_coprime_right) (metis Berlekamp_Type_Based.coprime_h_c_poly coprime_gcd coprime_iff_coprime insert.hyps(2)) qed finally show ?case . qed auto qed lemma monic_prod_gcd: -assumes f: "finite A" and f0: "(f :: 'b :: {field,factorial_ring_gcd} poly) \ 0" +assumes f: "finite A" and f0: "(f :: 'b :: {field_gcd} poly) \ 0" shows "monic (\c\A. gcd f (h - [:c:]))" using f proof (induct A) case (insert x A) have rw: "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps) show ?case proof (unfold rw, rule monic_mult) show "monic (gcd f (h - [:x:]))" using poly_gcd_monic[of f] f0 using insert.prems insert_iff by blast show "monic (\c\A. gcd f (h - [:c:]))" using insert.hyps(3) insert.prems by blast qed qed auto lemma coprime_not_unit_not_dvd: fixes a::"'b::semiring_gcd" assumes "a dvd b" and "coprime b c" and "\ is_unit a" shows "\ a dvd c" using assms coprime_divisors coprime_id_is_unit by fastforce lemma divides_prod2: fixes A::"'b::semiring_gcd set" assumes f: "finite A" and "\a\A. a dvd c" and "\a1 a2. a1 \ A \ a2 \ A \ a1 \ a2 \ coprime a1 a2" shows "\A dvd c" using assms proof (induct A) case (insert x A) have "\(insert x A) = x * \A" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... dvd c" proof (rule divides_mult) show "x dvd c" by (simp add: insert.prems) show "\A dvd c" using insert by auto from insert show "Rings.coprime x (\A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto lemma coprime_polynomial_factorization: - fixes a1 :: "'b :: {field,factorial_ring_gcd} poly" + fixes a1 :: "'b :: {field_gcd} poly" assumes irr: "as \ {q. irreducible q \ monic q}" and "finite as" and a1: "a1 \ as" and a2: "a2 \ as" and a1_not_a2: "a1 \ a2" shows "coprime a1 a2" proof (rule ccontr) assume not_coprime: "\ coprime a1 a2" let ?b= "gcd a1 a2" have b_dvd_a1: "?b dvd a1" and b_dvd_a2: "?b dvd a2" by simp+ have irr_a1: "irreducible a1" using a1 irr by blast have irr_a2: "irreducible a2" using a2 irr by blast have a2_not0: "a2 \ 0" using a2 irr by auto have degree_a1: "degree a1 \ 0" using irr_a1 by auto have degree_a2: "degree a2 \ 0" using irr_a2 by auto have not_a2_dvd_a1: "\ a2 dvd a1" proof (rule ccontr, simp) assume a2_dvd_a1: "a2 dvd a1" from this obtain k where k: "a1 = a2 * k" unfolding dvd_def by auto have k_not0: "k \ 0" using degree_a1 k by auto show False proof (cases "degree a2 = degree a1") case False have "degree a2 < degree a1" using False a2_dvd_a1 degree_a1 divides_degree by fastforce hence "\ irreducible a1" using degree_a2 a2_dvd_a1 degree_a2 by (metis degree_a1 irreducible\<^sub>dD(2) irreducible\<^sub>d_multD irreducible_connect_field k neq0_conv) thus False using irr_a1 by contradiction next case True have "degree a1 = degree a2 + degree k" unfolding k using degree_mult_eq[OF a2_not0 k_not0] by simp hence "degree k = 0" using True by simp hence "k = 1" using monic_factor a1 a2 irr k monic_degree_0 by auto hence "a1 = a2" using k by simp thus False using a1_not_a2 by contradiction qed qed have b_not0: "?b \ 0" by (simp add: a2_not0) have degree_b: "degree ?b > 0" using not_coprime[simplified] b_not0 is_unit_gcd is_unit_iff_degree by blast have "degree ?b < degree a2" by (meson b_dvd_a1 b_dvd_a2 irreducibleD' dvd_trans gcd_dvd_1 irr_a2 not_a2_dvd_a1 not_coprime) hence "\ irreducible\<^sub>d a2" using degree_a2 b_dvd_a2 degree_b by (metis degree_smult_eq irreducible\<^sub>d_dvd_smult less_not_refl3) thus False using irr_a2 by auto qed (* Proof of equation 14 in the book by Knuth *) theorem Berlekamp_gcd_step: fixes f::"'a mod_ring poly" and h::"'a mod_ring poly" assumes hq_mod_f: "[h^(CARD('a)) = h] (mod f)" and monic_f: "monic f" and sf_f: "square_free f" shows "f = prod (\c. gcd f (h - [:c:])) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof (cases "f=0") case True thus ?thesis using coeff_0 monic_f zero_neq_one by auto next case False note f_not_0 = False show ?thesis proof (rule poly_dvd_antisym) show "?rhs dvd f" using coprime_h_c_poly by (intro divides_prod_gcd, auto) have "monic ?rhs" by (rule monic_prod_gcd[OF _ f_not_0], simp) thus "coeff f (degree f) = coeff ?rhs (degree ?rhs)" using monic_f by auto next show "f dvd ?rhs" proof - let ?p = "CARD('a)" obtain P where finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" using monic_square_free_irreducible_factorization[OF monic_f sf_f] by auto have f_dvd_hqh: "f dvd (h^?p - h)" using hq_mod_f unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "f = \P" using f_desc_square_free by simp also have "... dvd ?rhs" proof (rule divides_prod2[OF finite_P]) show "\a1 a2. a1 \ P \ a2 \ P \ a1 \ a2 \ coprime a1 a2" using coprime_polynomial_factorization[OF P finite_P] by simp show "\a\P. a dvd (\c\UNIV. gcd f (h - [:c:]))" proof fix fi assume fi_P: "fi \ P" show "fi dvd ?rhs" proof (rule dvd_prod, auto) show "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible (fi)" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) have fi_dvd_hc: "\c\UNIV::'a mod_ring set. fi dvd (h-[:c:])" by (rule irreducible_dvd_prod[OF _ fi_dvd_prod_hc], simp add: irr_fi) thus "\c. fi dvd h - [:c:]" by simp qed qed qed finally show "f dvd ?rhs" . qed qed qed (******* Implementation of Berlekamp's algorithm (type-based version) *******) subsection \Definitions\ definition berlekamp_mat :: "'a mod_ring poly \ 'a mod_ring mat" where "berlekamp_mat u = (let n = degree u; mul_p = power_poly_f_mod u [:0,1:] (CARD('a)); xks = power_polys mul_p u 1 n in mat_of_rows_list n (map (\ cs. let coeffs_cs = (coeffs cs); k = n - length (coeffs cs) in (coeffs cs) @ replicate k 0) xks))" definition berlekamp_resulting_mat :: "('a mod_ring) poly \ 'a mod_ring mat" where "berlekamp_resulting_mat u = (let Q = berlekamp_mat u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then Q $$ (i,j) - 1 else Q $$ (i,j)) in (gauss_jordan_single (transpose_mat QI)))" definition berlekamp_basis :: "'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_basis u = (map (Poly o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" lemma berlekamp_basis_code[code]: "berlekamp_basis u = (map (poly_of_list o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" unfolding berlekamp_basis_def poly_of_list_def .. primrec berlekamp_factorization_main :: "nat \ 'a mod_ring poly list \ 'a mod_ring poly list \ nat \ 'a mod_ring poly list" where "berlekamp_factorization_main i divs (v # vs) n = (if v = 1 then berlekamp_factorization_main i divs vs n else if length divs = n then divs else let facts = [ w . u \ divs, s \ [0 ..< CARD('a)], w \ [gcd u (v - [:of_int s:])], w \ 1]; (lin,nonlin) = List.partition (\ q. degree q = i) facts in lin @ berlekamp_factorization_main i nonlin vs (n - length lin))" | "berlekamp_factorization_main i divs [] n = divs" definition berlekamp_monic_factorization :: "nat \ 'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_monic_factorization d f = (let vs = berlekamp_basis f; n = length vs; fs = berlekamp_factorization_main d [f] vs n in fs)" subsection \Properties\ lemma power_polys_works: fixes u::"'b::unique_euclidean_semiring" assumes i: "i < n" and c: "curr_p = curr_p mod u" (*Equivalent to degree curr_p < degree u*) shows "power_polys mult_p u curr_p n ! i = curr_p * mult_p ^ i mod u" using i c proof (induct n arbitrary: curr_p i) case 0 thus ?case by simp next case (Suc n) have p_rw: "power_polys mult_p u curr_p (Suc n) ! i = (curr_p # power_polys mult_p u (curr_p * mult_p mod u) n) ! i" by simp show ?case proof (cases "i=0") case True show ?thesis using Suc.prems unfolding p_rw True by auto next case False note i_not_0 = False show ?thesis proof (cases "i < n") case True note i_less_n = True have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (i - 1)" unfolding p_rw using nth_Cons_pos False by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (i-1) mod u" by (rule Suc.hyps) (auto simp add: i_less_n less_imp_diff_less) also have "... = curr_p * mult_p ^ i mod u" using False by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . next case False hence i_n: "i = n" using Suc.prems by auto have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (n - 1)" unfolding p_rw using nth_Cons_pos i_n i_not_0 by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (n-1) mod u" proof (rule Suc.hyps) show "n - 1 < n" using i_n i_not_0 by linarith show "curr_p * mult_p mod u = curr_p * mult_p mod u mod u" by simp qed also have "... = curr_p * mult_p ^ i mod u" using i_n [symmetric] i_not_0 by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . qed qed qed lemma length_power_polys[simp]: "length (power_polys mult_p u curr_p n) = n" by (induct n arbitrary: curr_p, auto) (* Equation 12 *) lemma Poly_berlekamp_mat: assumes k: "k < degree u" shows "Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k) mod u" proof - let ?map ="(map (\cs. coeffs cs @ replicate (degree u - length (coeffs cs)) 0) (power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)))" have "row (berlekamp_mat u) k = row (mat_of_rows_list (degree u) ?map) k" by (simp add: berlekamp_mat_def Let_def) also have "... = vec_of_list (?map ! k)" proof- { fix i assume i: "i < degree u" let ?c= "power_polys (power_poly_f_mod u [:0, 1:] CARD('a)) u 1 (degree u) ! i" let ?coeffs_c="(coeffs ?c)" have "?c = 1*([:0, 1:] ^ CARD('a) mod u)^i mod u" proof (unfold power_poly_f_mod_def, rule power_polys_works[OF i]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = [:0, 1:] ^ (CARD('a) * i) mod u" by (simp add: power_mod power_mult) finally have c_rw: "?c = [:0, 1:] ^ (CARD('a) * i) mod u" . have "length ?coeffs_c \ degree u" proof - show ?thesis proof (cases "?c = 0") case True thus ?thesis by auto next case False have "length ?coeffs_c = degree (?c) + 1" by (rule length_coeffs[OF False]) also have "... = degree ([:0, 1:] ^ (CARD('a) * i) mod u) + 1" using c_rw by simp also have "... \ degree u" by (metis One_nat_def add.right_neutral add_Suc_right c_rw calculation coeffs_def degree_0 degree_mod_less discrete gr_implies_not0 k list.size(3) one_neq_zero) finally show ?thesis . qed qed then have "length ?coeffs_c + (degree u - length ?coeffs_c) = degree u" by auto } with k show ?thesis by (intro row_mat_of_rows_list, auto) qed finally have row_rw: "row (berlekamp_mat u) k = vec_of_list (?map ! k)" . have "Poly (list_of_vec (row (berlekamp_mat u) k)) = Poly (list_of_vec (vec_of_list (?map ! k)))" unfolding row_rw .. also have "... = Poly (?map ! k)" by simp also have "... = [:0,1:]^(CARD('a) * k) mod u" proof - let ?cs = "(power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)) ! k" let ?c = "coeffs ?cs @ replicate (degree u - length (coeffs ?cs)) 0" have map_k_c: "?map ! k = ?c" by (rule nth_map, simp add: k) have "(Poly (?map ! k)) = Poly (coeffs ?cs)" unfolding map_k_c Poly_append_replicate_0 .. also have "... = ?cs" by simp also have "... = power_polys ([:0, 1:] ^ CARD('a) mod u) u 1 (degree u) ! k" by (simp add: power_poly_f_mod_def) also have "... = 1* ([:0,1:]^(CARD('a)) mod u) ^ k mod u" proof (rule power_polys_works[OF k]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = ([:0,1:]^(CARD('a)) mod u) ^ k mod u" by auto also have "... = [:0,1:]^(CARD('a) * k) mod u" by (simp add: power_mod power_mult) finally show ?thesis . qed finally show ?thesis . qed corollary Poly_berlekamp_cong_mat: assumes k: "k < degree u" shows "[Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k)] (mod u)" using Poly_berlekamp_mat[OF k] unfolding cong_def by auto lemma mat_of_rows_list_dim[simp]: "mat_of_rows_list n vs \ carrier_mat (length vs) n" "dim_row (mat_of_rows_list n vs) = length vs" "dim_col (mat_of_rows_list n vs) = n" unfolding mat_of_rows_list_def by auto lemma berlekamp_mat_closed[simp]: "berlekamp_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_mat u) = degree u" "dim_col (berlekamp_mat u) = degree u" unfolding carrier_mat_def berlekamp_mat_def Let_def by auto lemma vec_of_list_coeffs_nth: assumes i: "i \ {..degree h}" and h_not0: "h \ 0" shows "vec_of_list (coeffs h) $ i = coeff h i" proof - have "vec_of_list (map (coeff h) [0..i. f i mod z) A" using f by (induct, auto simp add: poly_mod_add_left) lemma prime_not_dvd_fact: assumes kn: "k < n" and prime_n: "prime n" shows "\ n dvd fact k" using kn proof (induct k) case 0 thus ?case using prime_n unfolding prime_nat_iff by auto next case (Suc k) show ?case proof (rule ccontr, unfold not_not) assume "n dvd fact (Suc k)" also have "... = Suc k * \{1..k}" unfolding fact_Suc unfolding fact_prod by simp finally have "n dvd Suc k * \{1..k}" . hence "n dvd Suc k \ n dvd \{1..k}" using prime_dvd_mult_eq_nat[OF prime_n] by blast moreover have "\ n dvd Suc k" by (simp add: Suc.prems(1) nat_dvd_not_less) moreover hence "\ n dvd \{1..k}" using Suc.hyps Suc.prems using Suc_lessD fact_prod[of k] by (metis of_nat_id) ultimately show False by simp qed qed lemma dvd_choose_prime: assumes kn: "k < n" and k: "k \ 0" and n: "n \ 0" and prime_n: "prime n" shows "n dvd (n choose k)" proof - have "n dvd (fact n)" by (simp add: fact_num_eq_if n) moreover have "\ n dvd (fact k * fact (n-k))" proof (rule ccontr, simp) assume "n dvd fact k * fact (n - k)" hence "n dvd fact k \ n dvd fact (n - k)" using prime_dvd_mult_eq_nat[OF prime_n] by simp moreover have "\ n dvd (fact k)" by (rule prime_not_dvd_fact[OF kn prime_n]) moreover have "\ n dvd fact (n - k)" using prime_not_dvd_fact[OF _ prime_n] kn k by simp ultimately show False by simp qed moreover have "(fact n::nat) = fact k * fact (n-k) * (n choose k)" using binomial_fact_lemma kn by auto ultimately show ?thesis using prime_n by (auto simp add: prime_dvd_mult_iff) qed lemma add_power_poly_mod_ring: fixes x :: "'a mod_ring poly" shows "(x + y) ^ CARD('a) = x ^ CARD('a) + y ^ CARD('a)" proof - let ?A="{0..CARD('a)}" let ?f="\k. of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k)" have A_rw: "?A = insert CARD('a) (insert 0 (?A - {0} - {CARD('a)}))" by fastforce have sum0: "sum ?f (?A - {0} - {CARD('a)}) = 0" proof (rule sum.neutral, rule) fix xa assume xa: "xa \ {0..CARD('a)} - {0} - {CARD('a)}" have card_dvd_choose: "CARD('a) dvd (CARD('a) choose xa)" proof (rule dvd_choose_prime) show "xa < CARD('a)" using xa by simp show "xa \ 0" using xa by simp show "CARD('a) \ 0" by simp show "prime CARD('a)" by (rule prime_card) qed hence rw0: "of_int (CARD('a) choose xa) = (0 :: 'a mod_ring)" by transfer simp have "of_nat (CARD('a) choose xa) = [:of_int (CARD('a) choose xa) :: 'a mod_ring:]" by (simp add: of_nat_poly) also have "... = [:0:]" using rw0 by simp finally show "of_nat (CARD('a) choose xa) * x ^ xa * y ^ (CARD('a) - xa) = 0" by auto qed have "(x + y)^CARD('a) = (\k = 0..CARD('a). of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k))" unfolding binomial_ring by (rule sum.cong, auto) also have "... = sum ?f (insert CARD('a) (insert 0 (?A - {0} - {CARD('a)})))" using A_rw by simp also have "... = ?f 0 + ?f CARD('a) + sum ?f (?A - {0} - {CARD('a)})" by auto also have "... = x^CARD('a) + y^CARD('a)" unfolding sum0 by auto finally show ?thesis . qed lemma power_poly_sum_mod_ring: fixes f :: "'b \ 'a mod_ring poly" assumes f: "finite A" shows "(sum f A) ^ CARD('a) = sum (\i. (f i) ^ CARD('a)) A" using f by (induct, auto simp add: add_power_poly_mod_ring) lemma poly_power_card_as_sum_of_monoms: fixes h :: "'a mod_ring poly" shows "h ^ CARD('a) = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof - have "h ^ CARD('a) = (\i\degree h. monom (coeff h i) i) ^ CARD('a)" by (simp add: poly_as_sum_of_monoms) also have "... = (\i\degree h. (monom (coeff h i) i) ^ CARD('a))" by (simp add: power_poly_sum_mod_ring) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" show "monom (coeff h x) x ^ CARD('a) = monom (coeff h x) (CARD('a) * x)" by (unfold poly_eq_iff, auto simp add: monom_power) qed finally show ?thesis . qed lemma degree_Poly_berlekamp_le: assumes i: "i < degree u" shows "degree (Poly (list_of_vec (row (berlekamp_mat u) i))) < degree u" by (metis Poly_berlekamp_mat degree_0 degree_mod_less gr_implies_not0 i linorder_neqE_nat) (* Equation 12: alternative statement. *) lemma monom_card_pow_mod_sum_berlekamp: assumes i: "i < degree u" shows "monom 1 (CARD('a) * i) mod u = (\j 0" using i by simp hence set_rw: "{..degree u - 1} = {.. degree u - 1" by auto have "monom 1 (CARD('a) * i) mod u = [:0, 1:] ^ (CARD('a) * i) mod u" using x_as_monom x_pow_n by metis also have "... = ?p" unfolding Poly_berlekamp_mat[OF i] by simp also have "... = (\i\degree u - 1. monom (coeff ?p i) i)" using degree_le2 poly_as_sum_of_monoms' by fastforce also have "... = (\ij {.. v = (\i = 0.. v = (\i = 0.. v = col A j \ v" using j row_transpose by auto also have "... = (\i = 0..iiiiiiiiiii degree u" by (metis Suc_leI assms coeffs_0_eq_Nil degree_0 length_coeffs_degree list.size(3) not_le_imp_less order.asym) thus ?thesis by simp qed lemma vec_of_list_coeffs_nth': assumes i: "i \ {..degree h}" and h_not0: "h \ 0" assumes "degree h < degree u" shows "vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) $ i = coeff h i" using assms by (transfer', auto simp add: mk_vec_def coeffs_nth length_coeffs_degree nth_append) lemma vec_of_list_coeffs_replicate_nth_0: assumes i: "i \ {.. {..{..degree h}") case True thus ?thesis using assms vec_of_list_coeffs_nth' h_not0 by simp next case False have c0: "coeff h i = 0" using False le_degree by auto thus ?thesis using assms False h_not0 by (transfer', auto simp add: mk_vec_def length_coeffs_degree nth_append c0) qed qed (* Equation 13 *) lemma equation_13: fixes u h defines H: "H \ vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0)" assumes deg_le: "degree h < degree u" (*Mandatory from equation 8*) shows "[h^CARD('a) = h] (mod u) \ (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H" (is "?lhs = ?rhs") proof - have f: "finite {..degree u}" by auto have [simp]: "dim_vec H = degree u" unfolding H using dim_vec_of_list_h deg_le by simp let ?B = "(berlekamp_mat u)" let ?f = "\i. (transpose_mat ?B *\<^sub>v H) $ i" show ?thesis proof assume rhs: ?rhs have dimv_h_dimr_B: "dim_vec H = dim_row ?B" by (metis berlekamp_mat_closed(2) berlekamp_mat_closed(3) dim_mult_mat_vec index_transpose_mat(2) rhs) have degree_h_less_dim_H: "degree h < dim_vec H" by (auto simp add: deg_le) have set_rw: "{..degree u - 1} = {.. degree u - 1" using deg_le by simp hence "h = (\j\degree u - 1. monom (coeff h j) j)" using poly_as_sum_of_monoms' by fastforce also have "... = (\jj {..j H) j)" by (rule sum.cong, auto) also have "... = (\ji = 0.. {.. H) x = monom (\i = 0..ji = 0..i = 0..ji = 0..ji = 0..ji = 0.. {0..jji = 0..i = 0..i = 0..x. x mod u"], rule sum.cong, rule) fix x assume x: "x \ {0.. {..i\degree h. monom (coeff h i) (CARD('a) * i)) mod u" proof (rule arg_cong[of _ _ "\x. x mod u"]) let ?f="\i. monom (coeff h i) (CARD('a) * i)" have ss0: "(\i = degree h + 1 ..< dim_vec H. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0..< dim_vec H} = {0..degree h} \ {degree h + 1 ..< dim_vec H}" using degree_h_less_dim_H by auto have "(\i = 0..i = 0..degree h. ?f i) + (\i = degree h + 1 ..< dim_vec H. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i = 0..degree h. ?f i)" unfolding ss0 by auto finally show "(\i = 0..i\degree h. ?f i)" by (simp add: atLeast0AtMost) qed also have "... = h^CARD('a) mod u" using poly_power_card_as_sum_of_monoms by auto finally show ?lhs unfolding cong_def using deg_le by (simp add: mod_poly_less) next assume lhs: ?lhs have deg_le': "degree h \ degree u - 1" using deg_le by auto have set_rw: "{..ii \ degree u - 1. monom (coeff h i) i)" by simp also have "... = (\i\degree h. monom (coeff h i) i)" unfolding poly_as_sum_of_monoms using poly_as_sum_of_monoms' deg_le' by auto also have "... = (\i\degree h. monom (coeff h i) i) mod u" by (simp add: deg_le mod_poly_less poly_as_sum_of_monoms) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i)) mod u" using lhs unfolding cong_def poly_as_sum_of_monoms poly_power_card_as_sum_of_monoms by auto also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i)) mod u" by (rule arg_cong[of _ _ "\x. x mod u"], rule sum.cong, simp_all add: mult_monom) also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i) mod u)" by (simp add: poly_mod_sum) also have "... = (\i\degree h. monom (coeff h i) 0 * (monom 1 (CARD('a)*i) mod u))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" have h_rw: "monom (coeff h x) 0 mod u = monom (coeff h x) 0" by (metis deg_le degree_pCons_eq_if gr_implies_not_zero linorder_neqE_nat mod_poly_less monom_0) have "monom (coeff h x) 0 * monom 1 (CARD('a) * x) = smult (coeff h x) (monom 1 (CARD('a) * x))" by (simp add: monom_0) also have "... mod u = Polynomial.smult (coeff h x) (monom 1 (CARD('a) * x) mod u)" using mod_smult_left by auto also have "... = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" by (simp add: monom_0) finally show "monom (coeff h x) 0 * monom 1 (CARD('a) * x) mod u = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" . qed also have "... = (\i\degree h. monom (coeff h i) 0 * (\j {..degree h}" have "(monom 1 (CARD('a) * x) mod u) = (\jjiji=degree h+1 ..< degree u. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0.. {degree h+1..i=0..i=0..degree h. ?f i) + (\i=degree h+1 ..< degree u. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i=0..degree h. ?f i)" using ss0 by simp finally show ?thesis by (simp add: atLeast0AtMost atLeast0LessThan) qed also have "... = (\ijijjijiijii. i < degree u \ coeff h i = (\ji. i < degree u \ H $ i = (\jjjjv H = H" proof (rule eq_vecI) fix i show "dim_vec (transpose_mat ?B *\<^sub>v H) = dim_vec (H)" by auto assume i: "i < dim_vec (H)" have "(transpose_mat ?B *\<^sub>v H) $ i = row (transpose_mat ?B) i \ H" using i by simp also have "... = (\j = 0..jv H) $ i = H $ i" . qed qed qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma exists_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\s. fi dvd (h - [:s:])" proof - let ?p = "CARD('a)" have f_dvd_hqh: "f dvd (h^?p - h)" using h unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis using irreducible_dvd_prod[OF _ fi_dvd_prod_hc] irr_fi by auto qed corollary exists_unique_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\!s. fi dvd (h - [:s:])" proof - obtain c where fi_dvd: "fi dvd (h - [:c:])" using assms exists_s_factor_dvd_h_s by blast have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" by (simp add: irr_fi irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis proof (rule ex1I[of _ c], auto simp add: fi_dvd) fix c2 assume fi_dvd_hc2: "fi dvd h - [:c2:]" have *: "fi dvd (h - [:c:]) * (h - [:c2:])" using fi_dvd by auto hence "fi dvd (h - [:c:]) \ fi dvd (h - [:c2:])" using irr_fi by auto thus "c2 = c" using coprime_h_c_poly coprime_not_unit_not_dvd fi_dvd fi_dvd_hc2 fi_not_unit by blast qed qed lemma exists_two_distint: "\a b::'a mod_ring. a \ b" by (rule exI[of _ 0], rule exI[of _ 1], auto) lemma coprime_cong_mult_factorization_poly: fixes f::"'b::{field} poly" - and a b p :: "'c :: {factorial_ring_gcd,field} poly" + and a b p :: "'c :: {field_gcd} poly" assumes finite_P: "finite P" and P: "P \ {q. irreducible q}" and p: "\p\P. [a=b] (mod p)" and coprime_P: "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" shows "[a = b] (mod (\a\P. a))" using finite_P P p coprime_P proof (induct P) case empty thus ?case by simp next case (insert p P) have ab_mod_pP: "[a=b] (mod (p * \P))" proof (rule coprime_cong_mult_poly) show "[a = b] (mod p)" using insert.prems by auto show "[a = b] (mod \P)" using insert.prems insert.hyps by auto from insert show "Rings.coprime p (\P)" by (auto intro: prod_coprime_right) qed thus ?case by (simp add: insert.hyps(1) insert.hyps(2)) qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma W_eq_berlekamp_mat: fixes u::"'a mod_ring poly" shows "{v. [v^CARD('a) = v] (mod u) \ degree v < degree u} = {h. let H = vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) in (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H \ degree h < degree u}" using equation_13 by (auto simp add: Let_def) lemma transpose_minus_1: assumes "dim_row Q = dim_col Q" shows "transpose_mat (Q - (1\<^sub>m (dim_row Q))) = (transpose_mat Q - (1\<^sub>m (dim_row Q)))" using assms unfolding mat_eq_iff by auto lemma system_iff: fixes v::"'b::comm_ring_1 vec" assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" proof - have t1:"transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v) \ (transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v)" by (subst minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have t2:"(transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v) \ transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v)" by (subst (asm) minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have "transpose_mat Q *\<^sub>v v - v = v - v \ transpose_mat Q *\<^sub>v v = v" proof - assume a1: "transpose_mat Q *\<^sub>v v - v = v - v" have f2: "transpose_mat Q *\<^sub>v v \ carrier_vec (dim_vec v)" by (metis dim_mult_mat_vec index_transpose_mat(2) sq_Q v carrier_vec_dim_vec) then have f3: "0\<^sub>v (dim_vec v) + transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v" by (meson left_zero_vec) have f4: "0\<^sub>v (dim_vec v) = transpose_mat Q *\<^sub>v v - v" using a1 by auto have f5: "- v \ carrier_vec (dim_vec v)" by simp then have f6: "- v + transpose_mat Q *\<^sub>v v = v - v" using f2 a1 using comm_add_vec minus_add_uminus_vec by fastforce have "v - v = - v + v" by auto then have "transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v - v + v" using f6 f4 f3 f2 by (metis (no_types, lifting) a1 assoc_add_vec comm_add_vec f5 carrier_vec_dim_vec) then show ?thesis using a1 by auto qed hence "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q *\<^sub>v v) - v = v - v)" by auto also have "... = ((transpose_mat Q *\<^sub>v v) - v = 0\<^sub>v (dim_vec v))" by auto also have "... = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using t1 t2 by auto finally show ?thesis. qed lemma system_if_mat_kernel: assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q))))" proof - have "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using assms system_iff by blast also have "... = (v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q)))))" unfolding mat_kernel_def unfolding transpose_minus_1[OF sq_Q] unfolding v by auto finally show ?thesis . qed lemma degree_u_mod_irreducible\<^sub>d_factor_0: fixes v and u::"'a mod_ring poly" defines W: "W \ {v. [v ^ CARD('a) = v] (mod u)}" assumes v: "v \ W" and finite_U: "finite U" and u_U: "u = \U" and U_irr_monic: "U \ {q. irreducible q \ monic q}" and fi_U: "fi \ U" shows "degree (v mod fi) = 0" proof - have deg_fi: "degree fi > 0" using U_irr_monic using fi_U irreducible\<^sub>dD[of fi] by auto have "fi dvd u" using u_U U_irr_monic finite_U dvd_prod_eqI fi_U by blast moreover have "u dvd (v^CARD('a) - v)" using v unfolding W cong_def by (simp add: mod_eq_dvd_iff_poly) ultimately have "fi dvd (v^CARD('a) - v)" by (rule dvd_trans) then have fi_dvd_prod_vc: "fi dvd prod (\c. v - [:c:]) (UNIV::'a mod_ring set)" by (simp add: poly_identity_mod_p) have irr_fi: "irreducible fi" using fi_U U_irr_monic by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (auto simp: poly_dvd_1) have fi_dvd_vc: "\c. fi dvd v - [:c:]" using irreducible_dvd_prod[OF _ fi_dvd_prod_vc] irr_fi by auto from this obtain a where "fi dvd v - [:a:]" by blast hence "v mod fi = [:a:] mod fi" using mod_eq_dvd_iff_poly by blast also have "... = [:a:]" by (simp add: deg_fi mod_poly_less) finally show ?thesis by simp qed (* Also polynomials over a field as a vector space in HOL-Algebra.*) definition "poly_abelian_monoid = \carrier = UNIV::'a mod_ring poly set, monoid.mult = ((*)), one = 1, zero = 0, add = (+), module.smult = smult\" interpretation vector_space_poly: vectorspace class_ring poly_abelian_monoid rewrites [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 0" and [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 1" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (+)" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (*)" and [simp]: "carrier poly_abelian_monoid = UNIV" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = smult" apply unfold_locales apply (auto simp: poly_abelian_monoid_def class_field_def smult_add_left smult_add_right Units_def) by (metis add.commute add.right_inverse) lemma subspace_Berlekamp: assumes f: "degree f \ 0" shows "subspace (class_ring :: 'a mod_ring ring) {v. [v^(CARD('a)) = v] (mod f) \ (degree v < degree f)} poly_abelian_monoid" proof - { fix v :: "'a mod_ring poly" and w :: "'a mod_ring poly" assume a1: "v ^ card (UNIV::'a set) mod f = v mod f" assume "w ^ card (UNIV::'a set) mod f = w mod f" then have "(v ^ card (UNIV::'a set) + w ^ card (UNIV::'a set)) mod f = (v + w) mod f" using a1 by (meson mod_add_cong) then have "(v + w) ^ card (UNIV::'a set) mod f = (v + w) mod f" by (simp add: add_power_poly_mod_ring) } note r=this thus ?thesis using f by (unfold_locales, auto simp: zero_power mod_smult_left smult_power cong_def degree_add_less) qed lemma berlekamp_resulting_mat_closed[simp]: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" proof - let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" let ?G="(gauss_jordan_single ?A)" have "?G \carrier_mat (degree u) (degree u)" by (rule gauss_jordan_single(2)[of ?A], auto) thus "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" unfolding berlekamp_resulting_mat_def Let_def by auto qed (*find_base vectors returns a basis:*) lemma berlekamp_resulting_mat_basis: "kernel.basis (degree u) (berlekamp_resulting_mat u) (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule find_base_vectors(3)) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" have "row_echelon_form (gauss_jordan_single ?A)" by (rule gauss_jordan_single(3)[of ?A], auto) thus "row_echelon_form (berlekamp_resulting_mat u)" unfolding berlekamp_resulting_mat_def Let_def by auto qed lemma set_berlekamp_basis_eq: "(set (berlekamp_basis u)) = (Poly \ list_of_vec)` (set (find_base_vectors (berlekamp_resulting_mat u)))" by (auto simp add: image_def o_def berlekamp_basis_def) lemma berlekamp_resulting_mat_constant: assumes deg_u: "degree u = 0" shows "berlekamp_resulting_mat u = 1\<^sub>m 0" by (unfold mat_eq_iff, auto simp add: deg_u) context fixes u::"'a::prime_card mod_ring poly" begin lemma set_berlekamp_basis_constant: assumes deg_u: "degree u = 0" shows "set (berlekamp_basis u) = {}" proof - have one_carrier: "1\<^sub>m 0 \ carrier_mat 0 0" by auto have m: "mat_kernel (1\<^sub>m 0) = {(0\<^sub>v 0) :: 'a mod_ring vec}" unfolding mat_kernel_def by auto have r: "row_echelon_form (1\<^sub>m 0 :: 'a mod_ring mat)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto have "set (find_base_vectors (1\<^sub>m 0)) \ {0\<^sub>v 0 :: 'a mod_ring vec}" using find_base_vectors(1)[OF r one_carrier] unfolding m . hence "set (find_base_vectors (1\<^sub>m 0) :: 'a mod_ring vec list) = {}" using find_base_vectors(2)[OF r one_carrier] using subset_singletonD by fastforce thus ?thesis unfolding set_berlekamp_basis_eq unfolding berlekamp_resulting_mat_constant[OF deg_u] by auto qed (*Maybe [simp]*) lemma row_echelon_form_berlekamp_resulting_mat: "row_echelon_form (berlekamp_resulting_mat u)" by (rule gauss_jordan_single(3), auto simp add: berlekamp_resulting_mat_def Let_def) lemma mat_kernel_berlekamp_resulting_mat_degree_0: assumes d: "degree u = 0" shows "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (auto simp add: mat_kernel_def mult_mat_vec_def d) lemma in_mat_kernel_berlekamp_resulting_mat: assumes x: "transpose_mat (berlekamp_mat u) *\<^sub>v x = x" and x_dim: "x \ carrier_vec (degree u)" shows "x \ mat_kernel (berlekamp_resulting_mat u)" proof - let ?QI="(mat(dim_row (berlekamp_mat u)) (dim_row (berlekamp_mat u)) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" have *: "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (degree u)) = transpose_mat ?QI" by auto have "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (dim_row (berlekamp_mat u))) *\<^sub>v x = 0\<^sub>v (dim_vec x)" using system_iff[of "berlekamp_mat u" x] x_dim x by auto hence "transpose_mat ?QI *\<^sub>v x = 0\<^sub>v (degree u)" using x_dim * by auto hence "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" unfolding berlekamp_resulting_mat_def Let_def using gauss_jordan_single(1)[of "transpose_mat ?QI" "degree u" "degree u" _ x] x_dim by auto thus ?thesis by (auto simp add: mat_kernel_def x_dim) qed private abbreviation "V \ kernel.VK (degree u) (berlekamp_resulting_mat u)" private abbreviation "W \ vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u) \ (degree v < degree u)}" interpretation V: vectorspace class_ring V proof - interpret k: kernel "(degree u)" "(degree u)" "(berlekamp_resulting_mat u)" by (unfold_locales; auto) show "vectorspace class_ring V" by intro_locales qed lemma linear_Poly_list_of_vec: shows "(Poly \ list_of_vec) \ module_hom class_ring V (vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u)})" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed qed lemma linear_Poly_list_of_vec': assumes "degree u > 0" shows "(Poly \ list_of_vec) \ module_hom R V W" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "degree (Poly (list_of_vec x)) < degree u" by (rule degree_Poly_list_of_vec, insert assms x, auto simp: mat_kernel_def) qed lemma berlekamp_basis_eq_8: assumes v: "v \ set (berlekamp_basis u)" shows "[v ^ CARD('a) = v] (mod u)" proof - { fix x assume x: "x \ set (find_base_vectors (berlekamp_resulting_mat u))" have "set (find_base_vectors (berlekamp_resulting_mat u)) \ mat_kernel (berlekamp_resulting_mat u)" proof (rule find_base_vectors(1)) show "row_echelon_form (berlekamp_resulting_mat u)" by (rule row_echelon_form_berlekamp_resulting_mat) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp qed hence "x \ mat_kernel (berlekamp_resulting_mat u)" using x by auto hence "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto } thus "[v ^ CARD('a) = v] (mod u)" using v unfolding set_berlekamp_basis_eq by auto qed lemma surj_Poly_list_of_vec: assumes deg_u: "degree u > 0" shows "(Poly \ list_of_vec)` (carrier V) = carrier W" proof (auto simp add: image_def) fix xa assume xa: "xa \ mat_kernel (berlekamp_resulting_mat u)" thus "[Poly (list_of_vec xa) ^ CARD('a) = Poly (list_of_vec xa)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto show "degree (Poly (list_of_vec xa)) < degree u" proof (rule degree_Poly_list_of_vec[OF _ deg_u]) show "xa \ carrier_vec (degree u)" using xa unfolding mat_kernel_def by simp qed next fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" show "\xa \ mat_kernel (berlekamp_resulting_mat u). x = Poly (list_of_vec xa)" proof (rule bexI[of _ "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)"]) let ?X = "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)" show "x = Poly (list_of_vec (vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)))" by auto have X: "?X \ carrier_vec (degree u)" unfolding carrier_vec_def by (auto, metis Suc_leI coeffs_0_eq_Nil deg_x degree_0 le_add_diff_inverse length_coeffs_degree linordered_semidom_class.add_diff_inverse list.size(3) order.asym) have t: "transpose_mat (berlekamp_mat u) *\<^sub>v ?X = ?X" using equation_13[OF deg_x] x by auto show "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0) \ mat_kernel (berlekamp_resulting_mat u)" by (rule in_mat_kernel_berlekamp_resulting_mat[OF t X]) qed qed lemma card_set_berlekamp_basis: "card (set (berlekamp_basis u)) = length (berlekamp_basis u)" proof - have b: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by auto have "(set (berlekamp_basis u)) = (Poly \ list_of_vec) ` set (find_base_vectors (berlekamp_resulting_mat u))" unfolding set_berlekamp_basis_eq .. also have " card ... = card (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule card_image, rule subset_inj_on[OF inj_Poly_list_of_vec]) show "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier_vec (degree u)" using find_base_vectors(1)[OF row_echelon_form_berlekamp_resulting_mat b] unfolding carrier_vec_def mat_kernel_def by auto qed also have "... = length (find_base_vectors (berlekamp_resulting_mat u))" by (rule length_find_base_vectors[symmetric, OF row_echelon_form_berlekamp_resulting_mat b]) finally show ?thesis unfolding berlekamp_basis_def by auto qed context assumes deg_u0[simp]: "degree u > 0" begin interpretation Berlekamp_subspace: vectorspace class_ring W by (rule vector_space_poly.subspace_is_vs[OF subspace_Berlekamp], simp) lemma linear_map_Poly_list_of_vec': "linear_map class_ring V W (Poly \ list_of_vec)" proof (auto simp add: linear_map_def) show "vectorspace class_ring V" by intro_locales show "vectorspace class_ring W" by (rule Berlekamp_subspace.vectorspace_axioms) show "mod_hom class_ring V W (Poly \ list_of_vec)" proof (rule mod_hom.intro, unfold mod_hom_axioms_def) show "module class_ring V" by intro_locales show "module class_ring W" using Berlekamp_subspace.vectorspace_axioms by intro_locales show "Poly \ list_of_vec \ module_hom class_ring V W" by (rule linear_Poly_list_of_vec'[OF deg_u0]) qed qed lemma berlekamp_basis_basis: "Berlekamp_subspace.basis (set (berlekamp_basis u))" proof (unfold set_berlekamp_basis_eq, rule linear_map.linear_inj_image_is_basis) show "linear_map class_ring V W (Poly \ list_of_vec)" by (rule linear_map_Poly_list_of_vec') show "inj_on (Poly \ list_of_vec) (carrier V)" proof (rule subset_inj_on[OF inj_Poly_list_of_vec]) show "carrier V \ carrier_vec (degree u)" by (auto simp add: mat_kernel_def) qed show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show b: "V.basis (set (find_base_vectors (berlekamp_resulting_mat u)))" by (rule berlekamp_resulting_mat_basis) show "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using b unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed qed lemma finsum_sum: fixes f::"'a mod_ring poly" assumes f: "finite B" and a_Pi: "a \ B \ carrier R" and V: "B \ carrier W" shows "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" using f a_Pi V proof (induct B) case empty thus ?case unfolding Berlekamp_subspace.module.M.finsum_empty by auto next case (insert x V) have hyp: "(\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) V" proof (rule insert.hyps) show "a \ V \ carrier R" using insert.prems unfolding class_field_def by auto show "V \ carrier W" using insert.prems by simp qed have "(\\<^bsub>W\<^esub>v\insert x V. a v \\<^bsub>W\<^esub> v) = (a x \\<^bsub>W\<^esub> x) \\<^bsub>W\<^esub> (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" proof (rule abelian_monoid.finsum_insert) show "abelian_monoid W" by (unfold_locales) show "finite V" by fact show "x \ V" by fact show "(\v. a v \\<^bsub>W\<^esub> v) \ V \ carrier W" proof (unfold Pi_def, rule, rule allI, rule impI) fix v assume v: "v\V" show "a v \\<^bsub>W\<^esub> v \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a v \ carrier class_ring" using insert.prems v unfolding Pi_def by (simp add: class_field_def) show "v \ carrier W" using v insert.prems by auto qed qed show "a x \\<^bsub>W\<^esub> x \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a x \ carrier class_ring" using insert.prems unfolding Pi_def by (simp add: class_field_def) show "x \ carrier W" using insert.prems by auto qed qed also have "... = (a x \\<^bsub>W\<^esub> x) + (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" by auto also have "... = (a x \\<^bsub>W\<^esub> x) + sum (\v. smult (a v) v) V" unfolding hyp by simp also have "... = (smult (a x) x) + sum (\v. smult (a v) v) V" by simp also have "... = sum (\v. smult (a v) v) (insert x V)" by (simp add: insert.hyps(1) insert.hyps(2)) finally show ?case . qed lemma exists_vector_in_Berlekamp_subspace_dvd: fixes p_i::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v. v \ {h. [h^(CARD('a)) = h] (mod u) \ degree h < degree u} \ v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1)" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by auto have irr_pj: "irreducible p_j" using pj P by auto obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m f_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m f_desc_square_free f_not_0 by auto obtain i where mi: "m i = p_i" and i: "i < n" using P_m pi by blast obtain j where mj: "m j = p_j" and j: "j < n" using P_m pj by blast have ij: "i \ j" using mi mj pi_pj by auto obtain s_i and s_j::"'a mod_ring" where si_sj: "s_i \ s_j" using exists_two_distint by blast let ?u="\x. if x = i then [:s_i:] else if x = j then [:s_j:] else [:0:]" have degree_si: "degree [:s_i:] = 0" by auto have degree_sj: "degree [:s_j:] = 0" by auto have "\!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\a\{i. i < n}. [v = ?u a] (mod m a))" proof (rule chinese_remainder_unique_poly) show "\a\{i. i < n}. \b\{i. i < n}. a \ b \ Rings.coprime (m a) (m b)" proof (rule+) fix a b assume "a \ {i. i < n}" and "b \ {i. i < n}" and "a \ b" thus "Rings.coprime (m a) (m b)" using coprime_polynomial_factorization[OF P finite_P, simplified] P_m by (metis image_eqI inj_onD inj_on_m) qed show "\i\{i. i < n}. m i \ 0" by (rule not_zero) show "0 < degree (prod m {i. i < n})" unfolding degree_prod using deg_u0 by blast qed from this obtain v where v: "\a\{i. i < n}. [v = ?u a] (mod m a)" and degree_v: "degree v < (\i\{i. i < n}. degree (m i))" by blast show ?thesis proof (rule exI[of _ v], auto) show vp_v_mod: "[v ^ CARD('a) = v] (mod u)" proof (unfold f_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P]) show "P \ {q. irreducible q}" using P by blast show "\p\P. [v ^ CARD('a) = v] (mod p)" proof (rule ballI) fix p assume p: "p \ P" hence irr_p: "irreducible\<^sub>d p" using P by auto obtain k where mk: "m k = p" and k: "k < n" using P_m p by blast have "[v = ?u k] (mod p)" using v mk k by auto moreover have "?u k mod p = ?u k" apply (rule mod_poly_less) using irreducible\<^sub>dD(1)[OF irr_p] by auto ultimately obtain s where v_mod_p: "v mod p = [:s:]" unfolding cong_def by force hence deg_v_p: "degree (v mod p) = 0" by auto have "v mod p = [:s:]" by (rule v_mod_p) also have "... = [:s:]^CARD('a)" unfolding poly_const_pow by auto also have "... = (v mod p) ^ CARD('a)" using v_mod_p by auto also have "... = (v mod p) ^ CARD('a) mod p" using calculation by auto also have "... = v^CARD('a) mod p" using power_mod by blast finally show "[v ^ CARD('a) = v] (mod p)" unfolding cong_def .. qed show "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" using P coprime_polynomial_factorization finite_P by auto qed have "[v = ?u i] (mod m i)" using v i by auto hence v_pi_si_mod: "v mod p_i = [:s_i:] mod p_i" unfolding cong_def mi by auto also have "... = [:s_i:]" apply (rule mod_poly_less) using irr_pi by auto finally have v_pi_si: "v mod p_i = [:s_i:]" . have "[v = ?u j] (mod m j)" using v j by auto hence v_pj_sj_mod: "v mod p_j = [:s_j:] mod p_j" unfolding cong_def mj using ij by auto also have "... = [:s_j:]" apply (rule mod_poly_less) using irr_pj by auto finally have v_pj_sj: "v mod p_j = [:s_j:]" . show "v mod p_i = v mod p_j \ False" using si_sj v_pi_si v_pj_sj by auto show "degree (v mod p_i) = 0" unfolding v_pi_si by simp show "degree (v mod p_j) = 0" unfolding v_pj_sj by simp show "\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" using v_pi_si_mod mod_eq_dvd_iff_poly by blast have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" using v_pj_sj_mod mod_eq_dvd_iff_poly by blast have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using vp_v_mod cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" proof (rule ccontr, simp) assume gcd_w: "gcd w (v - [:s_i:]) = w" show False apply (rule \v mod p_i = v mod p_j \ False\) by (metis irreducibleE \degree (v mod p_i) = 0\ gcd_greatest_iff gcd_w irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_si) qed show "gcd w (v - [:s_i:]) \ 1" by (metis irreducibleE gcd_greatest_iff irr_pi pi_dvd_v_si pi_dvd_w) qed show "degree v < degree u" proof - have "(\i | i < n. degree (m i)) = degree (prod m {i. i < n})" by (rule degree_prod_eq_sum_degree[symmetric, OF not_zero]) thus ?thesis using degree_v unfolding degree_prod by auto qed qed qed lemma exists_vector_in_Berlekamp_basis_dvd_aux: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j" proof (rule ccontr, auto) have V_in_carrier: "B \ carrier W" using basis_V unfolding Berlekamp_subspace.basis_def by auto assume all_eq: "\v\B. v mod p_i = v mod p_j" obtain x where x: "x \ {h. [h ^ CARD('a) = h] (mod u) \ degree h < degree u}" and x_pi_pj: "x mod p_i \ x mod p_j" and "degree (x mod p_i) = 0" and "degree (x mod p_j) = 0" "(\s. gcd w (x - [:s:]) \ w \ gcd w (x - [:s:]) \ 1)" using exists_vector_in_Berlekamp_subspace_dvd[OF _ _ _ pi pj _ _ _ _ w_dvd_f monic_w pi_dvd_w] assms by meson have x_in: "x \ carrier W" using x by auto hence "(\!a. a \ B \\<^sub>E carrier class_ring \ Berlekamp_subspace.lincomb a B = x)" using Berlekamp_subspace.basis_criterion[OF finite_V V_in_carrier] using basis_V by (simp add: class_field_def) from this obtain a where a_Pi: "a \ B \\<^sub>E carrier class_ring" and lincomb_x: "Berlekamp_subspace.lincomb a B = x" by blast have fs_ss: "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" proof (rule finsum_sum) show "finite B" by fact show "a \ B \ carrier class_ring" using a_Pi by auto show "B \ carrier W" by (rule V_in_carrier) qed have "x mod p_i = Berlekamp_subspace.lincomb a B mod p_i" using lincomb_x by simp also have " ... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_i" unfolding Berlekamp_subspace.lincomb_def .. also have "... = (sum (\v. smult (a v) v) B) mod p_i" unfolding fs_ss .. also have "... = sum (\v. smult (a v) v mod p_i) B" using finite_V poly_mod_sum by blast also have "... = sum (\v. smult (a v) (v mod p_i)) B" by (meson mod_smult_left) also have "... = sum (\v. smult (a v) (v mod p_j)) B" using all_eq by auto also have "... = sum (\v. smult (a v) v mod p_j) B" by (metis mod_smult_left) also have "... = (sum (\v. smult (a v) v) B) mod p_j" by (metis (mono_tags, lifting) finite_V poly_mod_sum sum.cong) also have "... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_j" unfolding fs_ss .. also have "... = Berlekamp_subspace.lincomb a B mod p_j" unfolding Berlekamp_subspace.lincomb_def .. also have "... = x mod p_j" using lincomb_x by simp finally have "x mod p_i = x mod p_j" . thus False using x_pi_pj by contradiction qed lemma exists_vector_in_Berlekamp_basis_dvd: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ \ coprime w (v - [:s:]))" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by fast have irr_pj: "irreducible p_j" using pj P by fast obtain v where vV: "v \ B" and v_pi_pj: "v mod p_i \ v mod p_j" using assms exists_vector_in_Berlekamp_basis_dvd_aux by blast have v: "v \ {v. [v ^ CARD('a) = v] (mod u)}" using basis_V vV unfolding Berlekamp_subspace.basis_def by auto have deg_v_pi: "degree (v mod p_i) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pi]) from this obtain s_i where v_pi_si: "v mod p_i = [:s_i:]" using degree_eq_zeroE by blast have deg_v_pj: "degree (v mod p_j) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pj]) from this obtain s_j where v_pj_sj: "v mod p_j = [:s_j:]" using degree_eq_zeroE by blast have si_sj: "s_i \ s_j" using v_pi_si v_pj_sj v_pi_pj by auto have "(\s. gcd w (v - [:s:]) \ w \ \ Rings.coprime w (v - [:s:]))" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pi_si) have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pj_sj) have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using v cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" by (metis irreducibleE deg_v_pi gcd_greatest_iff irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_pj v_pi_si) show "\ Rings.coprime w (v - [:s_i:])" using irr_pi pi_dvd_v_si pi_dvd_w by (simp add: irreducible\<^sub>dD(1) not_coprimeI) qed thus ?thesis using v_pi_pj vV deg_v_pi deg_v_pj by auto qed lemma exists_bijective_linear_map_W_vec: assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" shows "\f. linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f \ bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" proof - let ?B="carrier_vec (card P)::'a mod_ring vec set" have u_not_0: "u \ 0" using deg_u0 degree_0 by force obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence n: "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m u_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m u_desc_square_free u_not_0 by auto have deg_sum_eq: "(\i\{i. i < n}. degree (m i)) = degree u" by (metis degree_prod degree_prod_eq_sum_degree not_zero) have coprime_mi_mj:"\i\{i. i < n}. \j\{i. i < n}. i \ j \ coprime (m i) (m j)" proof (rule+) fix i j assume i: "i \ {i. i < n}" and j: "j \ {i. i < n}" and ij: "i \ j" show "coprime (m i) (m j)" proof (rule coprime_polynomial_factorization[OF P finite_P]) show "m i \ P" using i P_m by auto show "m j \ P" using j P_m by auto show "m i \ m j" using inj_on_m i ij j unfolding inj_on_def by blast qed qed let ?f = "\v. vec n (\i. coeff (v mod (m i)) 0)" interpret vec_VS: vectorspace class_ring "(module_vec TYPE('a mod_ring) n)" by (rule VS_Connect.vec_vs) interpret linear_map class_ring W "(module_vec TYPE('a mod_ring) n)" ?f by (intro_locales, unfold mod_hom_axioms_def LinearCombinations.module_hom_def, auto simp add: vec_eq_iff module_vec_def mod_smult_left poly_mod_add_left) have "linear_map class_ring W (module_vec TYPE('a mod_ring) n) ?f" by (intro_locales) moreover have inj_f: "inj_on ?f (carrier W)" proof (rule Ke0_imp_inj, auto simp add: mod_hom.ker_def) show "[0 ^ CARD('a) = 0] (mod u)" by (simp add: cong_def zero_power) show "vec n (\i. 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" by (auto simp add: module_vec_def) fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" and v: "vec n (\i. coeff (x mod m i) 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" have cong_0: "\i\{i. i < n}. [x = (\i. 0) i] (mod m i)" proof (rule, unfold cong_def) fix i assume i: "i \ {i. i < n}" have deg_x_mod_mi: "degree (x mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "x \ {v. [v ^ CARD('a) = v] (mod u)}" using x by auto show "m i \ P" using P_m i by auto qed thus "x mod m i = 0 mod m i" using v unfolding module_vec_def by (auto, metis i leading_coeff_neq_0 mem_Collect_eq index_vec index_zero_vec(1)) qed moreover have deg_x2: "degree x < (\i\{i. i < n}. degree (m i))" using deg_sum_eq deg_x by simp moreover have "\i\{i. i < n}. [0 = (\i. 0) i] (mod m i)" by (auto simp add: cong_def) moreover have "degree 0 < (\i\{i. i < n}. degree (m i))" using degree_prod deg_sum_eq deg_u0 by force moreover have "\!x. degree x < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [x = (\i. 0) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) ultimately show "x = 0" by blast qed moreover have "?f ` (carrier W) = ?B" proof (auto simp add: image_def) fix xa show "n = card P" by (auto simp add: n) next fix x::"'a mod_ring vec" assume x: "x \ carrier_vec (card P)" have " \!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) from this obtain v where deg_v: "degree v < (\i\{i. i < n}. degree (m i))" and v_x_cong: "(\i \ {i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" by auto show "\xa. [xa ^ CARD('a) = xa] (mod u) \ degree xa < degree u \ x = vec n (\i. coeff (xa mod m i) 0)" proof (rule exI[of _ v], auto) show v: "[v ^ CARD('a) = v] (mod u)" proof (unfold u_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P], auto) fix y assume y: "y \ P" thus "irreducible y" using P by blast obtain i where i: "i \ {i. i < n}" and mi: "y = m i" using P_m y by blast have "irreducible (m i)" using i P_m P by auto moreover have "[v = [:x $ i:]] (mod m i)" using v_x_cong i by auto ultimately have v_mi_eq_xi: "v mod m i = [:x $ i:]" by (auto simp: cong_def intro!: mod_poly_less) have xi_pow_xi: "[:x $ i:]^CARD('a) = [:x $ i:]" by (simp add: poly_const_pow) hence "(v mod m i)^CARD('a) = v mod m i" using v_mi_eq_xi by auto hence "(v mod m i)^CARD('a) = (v^CARD('a) mod m i)" by (metis mod_mod_trivial power_mod) thus "[v ^ CARD('a) = v] (mod y)" unfolding mi cong_def v_mi_eq_xi xi_pow_xi by simp next fix p1 p2 assume "p1 \ P" and "p2 \ P" and "p1 \ p2" then show "Rings.coprime p1 p2" using coprime_polynomial_factorization[OF P finite_P] by auto qed show "degree v < degree u" using deg_v deg_sum_eq degree_prod by presburger show "x = vec n (\i. coeff (v mod m i) 0)" proof (unfold vec_eq_iff, rule conjI) show "dim_vec x = dim_vec (vec n (\i. coeff (v mod m i) 0))" using x n by simp show "\ii. coeff (v mod m i) 0)). x $ i = vec n (\i. coeff (v mod m i) 0) $ i" proof (auto) fix i assume i: "i < n" have deg_mi: "irreducible (m i)" using i P_m P by auto have deg_v_mi: "degree (v mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "v \ {v. [v ^ CARD('a) = v] (mod u)}" using v by fast show "m i \ P" using P_m i by auto qed have "v mod m i = [:x $ i:] mod m i" using v_x_cong i unfolding cong_def by auto also have "... = [:x $ i:]" using deg_mi by (auto intro!: mod_poly_less) finally show "x $ i = coeff (v mod m i) 0" by simp qed qed qed qed ultimately show ?thesis unfolding bij_betw_def n by auto qed lemma fin_dim_kernel_berlekamp: "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using berlekamp_resulting_mat_basis[of u] unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed lemma Berlekamp_subspace_fin_dim: "Berlekamp_subspace.fin_dim" proof (rule linear_map.surj_fin_dim[OF linear_map_Poly_list_of_vec']) show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show "V.fin_dim" by (rule fin_dim_kernel_berlekamp) qed context fixes P assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" begin interpretation RV: vec_space "TYPE('a mod_ring)" "card P" . lemma Berlekamp_subspace_eq_dim_vec: "Berlekamp_subspace.dim = RV.dim" proof - obtain f where lm_f: "linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f" and bij_f: "bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" using exists_bijective_linear_map_W_vec[OF finite_P u_desc_square_free P] by blast show ?thesis proof (rule linear_map.dim_eq[OF lm_f Berlekamp_subspace_fin_dim]) show "inj_on f (carrier W)" by (rule bij_betw_imp_inj_on[OF bij_f]) show " f ` carrier W = carrier RV.V" using bij_f unfolding bij_betw_def by auto qed qed lemma Berlekamp_subspace_dim: "Berlekamp_subspace.dim = card P" using Berlekamp_subspace_eq_dim_vec RV.dim_is_n by simp corollary card_berlekamp_basis_number_factors: "card (set (berlekamp_basis u)) = card P" unfolding Berlekamp_subspace_dim[symmetric] by (rule Berlekamp_subspace.dim_basis[symmetric], auto simp add: berlekamp_basis_basis) lemma length_berlekamp_basis_numbers_factors: "length (berlekamp_basis u) = card P" using card_set_berlekamp_basis card_berlekamp_basis_number_factors by auto end end end end context assumes "SORT_CONSTRAINT('a :: prime_card)" begin context fixes f :: "'a mod_ring poly" and n assumes sf: "square_free f" and n: "n = length (berlekamp_basis f)" and monic_f: "monic f" begin lemma berlekamp_basis_length_factorization: assumes f: "f = prod_list us" and d: "\ u. u \ set us \ degree u > 0" shows "length us \ n" proof (cases "degree f = 0") case True have "us = []" proof (rule ccontr) assume "us \ []" from this obtain u where u: "u \ set us" by fastforce hence deg_u: "degree u > 0" using d by auto have "degree f = degree (prod_list us)" unfolding f .. also have "... = sum_list (map degree us)" proof (rule degree_prod_list_eq) fix p assume p: "p \ set us" show "p \ 0" using d[OF p] degree_0 by auto qed also have " ... \ degree u" by (simp add: member_le_sum_list u) finally have "degree f > 0" using deg_u by auto thus False using True by auto qed thus ?thesis by simp next case False hence f_not_0: "f \ 0" using degree_0 by fastforce obtain P where fin_P: "finite P" and f_P: "f = \P" and P: "P \ {p. irreducible p \ monic p}" using monic_square_free_irreducible_factorization[OF monic_f sf] by auto have n_card_P: "n = card P" using P False f_P fin_P length_berlekamp_basis_numbers_factors n by blast have distinct_us: "distinct us" using d f sf square_free_prod_list_distinct by blast let ?us'="(map normalize us)" have distinct_us': "distinct ?us'" proof (auto simp add: distinct_map) show "distinct us" by (rule distinct_us) show "inj_on normalize (set us)" proof (auto simp add: inj_on_def, rule ccontr) fix x y assume x: "x \ set us" and y: "y \ set us" and n: "normalize x = normalize y" and x_not_y: "x \ y" from normalize_eq_imp_smult[OF n] obtain c where c0: "c \ 0" and y_smult: "y = smult c x" by blast have sf_xy: "square_free (x*y)" proof (rule square_free_factor[OF _ sf]) have "x*y = prod_list [x,y]" by simp also have "... dvd prod_list us" by (rule prod_list_dvd_prod_list_subset, auto simp add: x y x_not_y distinct_us) also have "... = f" unfolding f .. finally show "x * y dvd f" . qed have "x * y = smult c (x*x)" using y_smult mult_smult_right by auto hence sf_smult: "square_free (smult c (x*x))" using sf_xy by auto have "x*x dvd (smult c (x*x))" by (simp add: dvd_smult) hence "\ square_free (smult c (x*x))" by (metis d square_free_def x) thus "False" using sf_smult by contradiction qed qed have length_us_us': "length us = length ?us'" by simp have f_us': "f = prod_list ?us'" proof - have "f = normalize f" using monic_f f_not_0 by (simp add: normalize_monic) also have "... = prod_list ?us'" by (unfold f, rule prod_list_normalize[of us]) finally show ?thesis . qed have "\Q. prod_list Q = prod_list ?us' \ length ?us' \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule exists_factorization_prod_list) show "degree (prod_list ?us') > 0" using False f_us' by auto show "square_free (prod_list ?us')" using f_us' sf by auto fix u assume u: "u \ set ?us'" have u_not0: "u \ 0" using d u degree_0 by fastforce have "degree u > 0" using d u by auto moreover have "monic u" using u monic_normalize[OF u_not0] by auto ultimately show "degree u > 0 \ monic u" by simp qed from this obtain Q where Q_us': "prod_list Q = prod_list ?us'" and length_us'_Q: "length ?us' \ length Q" and Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast have distinct_Q: "distinct Q" proof (rule square_free_prod_list_distinct) show "square_free (prod_list Q)" using Q_us' f_us' sf by auto show "\u. u \ set Q \ degree u > 0" using Q irreducible_degree_field by auto qed have set_Q_P: "set Q = P" proof (rule monic_factorization_uniqueness) show "\(set Q) = \P" using Q_us' by (metis distinct_Q f_P f_us' list.map_ident prod.distinct_set_conv_list) qed (insert P Q fin_P, auto) hence "length Q = card P" using distinct_Q distinct_card by fastforce have "length us = length ?us'" by (rule length_us_us') also have "... \ length Q" using length_us'_Q by auto also have "... = card (set Q)" using distinct_card[OF distinct_Q] by simp also have "... = card P" using set_Q_P by simp finally show ?thesis using n_card_P by simp qed lemma berlekamp_basis_irreducible: assumes f: "f = prod_list us" and n_us: "length us = n" and us: "\ u. u \ set us \ degree u > 0" and u: "u \ set us" shows "irreducible u" proof (fold irreducible_connect_field, intro irreducible\<^sub>dI[OF us[OF u]]) fix q r :: "'a mod_ring poly" assume dq: "degree q > 0" and qu: "degree q < degree u" and dr: "degree r > 0" and uqr: "u = q * r" with us[OF u] have q: "q \ 0" and r: "r \ 0" by auto from split_list[OF u] obtain xs ys where id: "us = xs @ u # ys" by auto let ?us = "xs @ q # r # ys" have f: "f = prod_list ?us" unfolding f id uqr by simp { fix x assume "x \ set ?us" with us[unfolded id] dr dq have "degree x > 0" by auto } from berlekamp_basis_length_factorization[OF f this] have "length ?us \ n" by simp also have "\ = length us" unfolding n_us by simp also have "\ < length ?us" unfolding id by simp finally show False by simp qed end lemma not_irreducible_factor_yields_prime_factors: - assumes uf: "u dvd (f :: 'b :: {field,euclidean_ring_gcd} poly)" and fin: "finite P" and fP: "f = \P" and P: "P \ {q. irreducible q \ monic q}" + assumes uf: "u dvd (f :: 'b :: {field_gcd} poly)" and fin: "finite P" + and fP: "f = \P" and P: "P \ {q. irreducible q \ monic q}" and u: "degree u > 0" "\ irreducible u" shows "\ pi pj. pi \ P \ pj \ P \ pi \ pj \ pi dvd u \ pj dvd u" proof - from finite_distinct_list[OF fin] obtain ps where Pps: "P = set ps" and dist: "distinct ps" by auto have fP: "f = prod_list ps" unfolding fP Pps using dist by (simp add: prod.distinct_set_conv_list) note P = P[unfolded Pps] have "set ps \ P" unfolding Pps by auto from uf[unfolded fP] P dist this show ?thesis proof (induct ps) case Nil with u show ?case using divides_degree[of u 1] by auto next case (Cons p ps) from Cons(3) have ps: "set ps \ {q. irreducible q \ monic q}" by auto from Cons(2) have dvd: "u dvd p * prod_list ps" by simp obtain k where gcd: "u = gcd p u * k" by (meson dvd_def gcd_dvd2) from Cons(3) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of u] *(2) have "gcd p u = 1 \ gcd p u = p" by auto thus ?case proof assume "gcd p u = 1" then have "Rings.coprime p u" by (rule gcd_eq_1_imp_coprime) with dvd have "u dvd prod_list ps" using coprime_dvd_mult_right_iff coprime_imp_coprime by blast from Cons(1)[OF this ps] Cons(4-5) show ?thesis by auto next assume "gcd p u = p" with gcd have upk: "u = p * k" by auto hence p: "p dvd u" by auto from dvd[unfolded upk] *(3) have kps: "k dvd prod_list ps" by auto from dvd u * have dk: "degree k > 0" by (metis gr0I irreducible_mult_unit_right is_unit_iff_degree mult_zero_right upk) from ps kps have "\ q \ set ps. q dvd k" proof (induct ps) case Nil with dk show ?case using divides_degree[of k 1] by auto next case (Cons p ps) from Cons(3) have dvd: "k dvd p * prod_list ps" by simp obtain l where gcd: "k = gcd p k * l" by (meson dvd_def gcd_dvd2) from Cons(2) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of k] *(2) have "gcd p k = 1 \ gcd p k = p" by auto thus ?case proof assume "gcd p k = 1" with dvd have "k dvd prod_list ps" by (metis dvd_triv_left gcd_greatest_mult mult.left_neutral) from Cons(1)[OF _ this] Cons(2) show ?thesis by auto next assume "gcd p k = p" with gcd have upk: "k = p * l" by auto hence p: "p dvd k" by auto thus ?thesis by auto qed qed then obtain q where q: "q \ set ps" and dvd: "q dvd k" by auto from dvd upk have qu: "q dvd u" by auto from Cons(4) q have "p \ q" by auto thus ?thesis using q p qu Cons(5) by auto qed qed qed lemma berlekamp_factorization_main: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and vs: "vs = vs1 @ vs2" and vsf: "vs = berlekamp_basis f" and n_bb: "n = length (berlekamp_basis f)" and n: "n = length us1 + n2" and us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" and us1: "\ u. u \ set us1 \ monic u \ irreducible u" and divs: "\ d. d \ set divs \ monic d \ degree d > 0" and vs1: "\ u v i. v \ set vs1 \ u \ set us1 \ set divs \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1,u}" and f: "f = prod_list (us1 @ divs)" and deg_f: "degree f > 0" and d: "\ g. g dvd f \ degree g = d \ irreducible g" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - have mon_f: "monic f" unfolding f by (rule monic_prod_list, insert divs us1, auto) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \ P" "P \ {q. irreducible q \ monic q}" by auto hence f0: "f \ 0" by auto show ?thesis using vs n us divs f us1 vs1 proof (induct vs2 arbitrary: divs n2 us1 vs1) case (Cons v vs2) show ?case proof (cases "v = 1") case False from Cons(2) vsf have v: "v \ set (berlekamp_basis f)" by auto from berlekamp_basis_eq_8[OF this] have vf: "[v ^ CARD('a) = v] (mod f)" . let ?gcd = "\ u i. gcd u (v - [:of_int i:])" let ?gcdn = "\ u i. gcd u (v - [:of_nat i:])" let ?map = "\ u. (map (\ i. ?gcd u i) [0 ..< CARD('a)])" define udivs where "udivs \ \ u. filter (\ w. w \ 1) (?map u)" { obtain xs where xs: "[0.. u. [w. i \ [0 ..< CARD('a)], w \ [?gcd u i], w \ 1])" unfolding udivs_def xs by (intro ext, auto simp: o_def, induct xs, auto) } note udivs_def' = this define facts where "facts \ [ w . u \ divs, w \ udivs u]" { fix u assume u: "u \ set divs" then obtain bef aft where divs: "divs = bef @ u # aft" by (meson split_list) from Cons(5)[OF u] have mon_u: "monic u" by simp have uf: "u dvd f" unfolding Cons(6) divs by auto from vf uf have vu: "[v ^ CARD('a) = v] (mod u)" by (rule cong_dvd_modulus_poly) from square_free_factor[OF uf sf_f] have sf_u: "square_free u" . let ?g = "?gcd u" from mon_u have u0: "u \ 0" by auto have "u = (\c\UNIV. gcd u (v - [:c:]))" using Berlekamp_gcd_step[OF vu mon_u sf_u] . also have "\ = (\i \ {0..< int CARD('a)}. ?g i)" by (rule sym, rule prod.reindex_cong[OF to_int_mod_ring_hom.inj_f range_to_int_mod_ring[symmetric]], simp add: of_int_of_int_mod_ring) finally have u_prod: "u = (\i \ {0..< int CARD('a)}. ?g i)" . let ?S = "{0.. ?S" hence "?g i \ 1" by auto moreover have mgi: "monic (?g i)" by (rule poly_gcd_monic, insert u0, auto) ultimately have "degree (?g i) > 0" using monic_degree_0 by blast note this mgi } note gS = this have int_set: "int ` set [0.. ?S" and j: "j \ ?S" and gij: "?g i = ?g j" show "i = j" proof (rule ccontr) define S where "S = {0.. S" "j \ S" "finite S" using i j unfolding S_def by auto assume ij: "i \ j" have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = ?g i * ?g j * (\i \ S. ?g i)" unfolding id using S ij by auto also have "\ = ?g i * ?g i * (\i \ S. ?g i)" unfolding gij by simp finally have dvd: "?g i * ?g i dvd u" unfolding dvd_def by auto with sf_u[unfolded square_free_def, THEN conjunct2, rule_format, OF gS(1)[OF i]] show False by simp qed qed have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = (\i \ ?S. ?g i)" by (rule sym, rule prod.setdiff_irrelevant, auto) also have "\ = \ (set (udivs u))" unfolding udivs_def set_filter set_map by (rule sym, rule prod.reindex_cong[of ?g, OF inj _ refl], auto simp: int_set[symmetric]) finally have u_udivs: "u = \(set (udivs u))" . { fix w assume mem: "w \ set (udivs u)" then obtain i where w: "w = ?g i" and i: "i \ ?S" unfolding udivs_def set_filter set_map int_set by auto have wu: "w dvd u" by (simp add: w) let ?v = "\ j. v - [:of_nat j:]" define j where "j = nat i" from i have j: "of_int i = (of_nat j :: 'a mod_ring)" "j < CARD('a)" unfolding j_def by auto from gS[OF i, folded w] have *: "degree w > 0" "monic w" "w \ 0" by auto from w have "w dvd ?v j" using j by simp hence gcdj: "?gcdn w j = w" by (metis gcd.commute gcd_left_idem j(1) w) { fix j' assume j': "j' < CARD('a)" have "?gcdn w j' \ {1,w}" proof (rule ccontr) assume not: "?gcdn w j' \ {1,w}" with gcdj have neq: "int j' \ int j" by auto (* next step will yield contradiction to square_free u *) let ?h = "?gcdn w j'" from *(3) not have deg: "degree ?h > 0" using monic_degree_0 poly_gcd_monic by auto have hw: "?h dvd w" by auto have "?h dvd ?gcdn u j'" using wu using dvd_trans by auto also have "?gcdn u j' = ?g j'" by simp finally have hj': "?h dvd ?g j'" by auto from divides_degree[OF this] deg u0 have degj': "degree (?g j') > 0" by auto hence j'1: "?g j' \ 1" by auto with j' have mem': "?g j' \ set (udivs u)" unfolding udivs_def by auto from degj' j' have j'S: "int j' \ ?S" by auto from i j have jS: "int j \ ?S" by auto from inj_on_contraD[OF inj neq j'S jS] have neq: "w \ ?g j'" using w j by auto have cop: "\ coprime w (?g j')" using hj' hw deg by (metis coprime_not_unit_not_dvd poly_dvd_1 Nat.neq0_conv) obtain w' where w': "?g j' = w'" by auto from u_udivs sf_u have "square_free (\(set (udivs u)))" by simp from square_free_prodD[OF this finite_set mem mem'] cop neq show False by simp qed } from gS[OF i, folded w] i this have "degree w > 0" "monic w" "\ j. j < CARD('a) \ ?gcdn w j \ {1,w}" by auto } note udivs = this let ?is = "filter (\ i. ?g i \ 1) (map int [0 ..< CARD('a)])" have id: "udivs u = map ?g ?is" unfolding udivs_def filter_map o_def .. have dist: "distinct (udivs u)" unfolding id distinct_map proof (rule conjI[OF distinct_filter], unfold distinct_map) have "?S = set ?is" unfolding int_set[symmetric] by auto thus "inj_on ?g (set ?is)" using inj by auto qed (auto simp: inj_on_def) from u_udivs prod.distinct_set_conv_list[OF dist, of id] have "prod_list (udivs u) = u" by auto note udivs this dist } note udivs = this have facts: "facts = concat (map udivs divs)" unfolding facts_def by auto obtain lin nonlin where part: "List.partition (\ q. degree q = d) facts = (lin,nonlin)" by force from Cons(6) have "f = prod_list us1 * prod_list divs" by auto also have "prod_list divs = prod_list facts" unfolding facts using udivs(4) by (induct divs, auto) finally have f: "f = prod_list us1 * prod_list facts" . note facts' = facts { fix u assume u: "u \ set facts" from u[unfolded facts] obtain u' where u': "u' \ set divs" and u: "u \ set (udivs u')" by auto from u' udivs(1-2)[OF u' u] prod_list_dvd[OF u, unfolded udivs(4)[OF u']] have "degree u > 0" "monic u" "\ u' \ set divs. u dvd u'" by auto } note facts = this have not1: "(v = 1) = False" using False by auto have "us = us1 @ (if length divs = n2 then divs else let (lin, nonlin) = List.partition (\q. degree q = d) facts in lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding Cons(4) facts_def udivs_def' berlekamp_factorization_main.simps Let_def not1 if_False by (rule arg_cong[where f = "\ x. us1 @ x"], rule if_cong, simp_all) (* takes time *) hence res: "us = us1 @ (if length divs = n2 then divs else lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding part by auto show ?thesis proof (cases "length divs = n2") case False with res have us: "us = (us1 @ lin) @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin)" by auto from Cons(2) have vs: "vs = (vs1 @ [v]) @ vs2" by auto have f: "f = prod_list ((us1 @ lin) @ nonlin)" unfolding f using prod_list_partition[OF part] by simp { fix u assume "u \ set ((us1 @ lin) @ nonlin)" with part have "u \ set facts \ set us1" by auto with facts Cons(7) have "degree u > 0" by (auto simp: irreducible_degree_field) } note deg = this from berlekamp_basis_length_factorization[OF sf_f n_bb mon_f f deg, unfolded Cons(3)] have "n2 \ length lin" by auto hence n: "n = length (us1 @ lin) + (n2 - length lin)" unfolding Cons(3) by auto show ?thesis proof (rule Cons(1)[OF vs n us _ f]) fix u assume "u \ set nonlin" with part have "u \ set facts" by auto from facts[OF this] show "monic u \ degree u > 0" by auto next fix u assume u: "u \ set (us1 @ lin)" { assume *: "\ (monic u \ irreducible\<^sub>d u)" with Cons(7) u have "u \ set lin" by auto with part have uf: "u \ set facts" and deg: "degree u = d" by auto from facts[OF uf] obtain u' where "u' \ set divs" and uu': "u dvd u'" by auto from this(1) have "u' dvd f" unfolding Cons(6) using prod_list_dvd[of u'] by auto with uu' have "u dvd f" by (rule dvd_trans) from facts[OF uf] d[OF this deg] * have False by auto } thus "monic u \ irreducible u" by auto next fix w u i assume w: "w \ set (vs1 @ [v])" and u: "u \ set (us1 @ lin) \ set nonlin" and i: "i < CARD('a)" from u part have u: "u \ set us1 \ set facts" by auto show "gcd u (w - [:of_nat i:]) \ {1, u}" proof (cases "u \ set us1") case True from Cons(7)[OF this] have "monic u" "irreducible u" by auto thus ?thesis by (rule monic_irreducible_gcd) next case False with u have u: "u \ set facts" by auto show ?thesis proof (cases "w = v") case True from u[unfolded facts'] obtain u' where u: "u \ set (udivs u')" and u': "u' \ set divs" by auto from udivs(3)[OF u' u i] show ?thesis unfolding True . next case False with w have w: "w \ set vs1" by auto from u obtain u' where u': "u' \ set divs" and dvd: "u dvd u'" using facts(3)[of u] dvd_refl[of u] by blast from w have "w \ set vs1 \ w = v" by auto from facts(1-2)[OF u] have u: "monic u" by auto from Cons(8)[OF w _ i] u' have "gcd u' (w - [:of_nat i:]) \ {1, u'}" by auto with dvd u show ?thesis by (rule monic_gcd_dvd) qed qed qed next case True with res have us: "us = us1 @ divs" by auto from Cons(3) True have n: "n = length us" unfolding us by auto show ?thesis unfolding us[symmetric] proof (intro conjI ballI) show f: "f = prod_list us" unfolding us using Cons(6) by simp { fix u assume "u \ set us" hence "degree u > 0" using Cons(5) Cons(7)[unfolded irreducible\<^sub>d_def] unfolding us by (auto simp: irreducible_degree_field) } note deg = this fix u assume u: "u \ set us" thus "monic u" unfolding us using Cons(5) Cons(7) by auto show "irreducible u" by (rule berlekamp_basis_irreducible[OF sf_f n_bb mon_f f n[symmetric] deg u]) qed qed next case True (* v = 1 *) with Cons(4) have us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" by simp from Cons(2) True have vs: "vs = (vs1 @ [1]) @ vs2" by auto show ?thesis proof (rule Cons(1)[OF vs Cons(3) us Cons(5-7)], goal_cases) case (3 v u i) show ?case proof (cases "v = 1") case False with 3 Cons(8)[of v u i] show ?thesis by auto next case True hence deg: "degree (v - [: of_nat i :]) = 0" by (metis (no_types, hide_lams) degree_pCons_0 diff_pCons diff_zero pCons_one) from 3(2) Cons(5,7)[of u] have "monic u" by auto from gcd_monic_constant[OF this deg] show ?thesis . qed qed qed next case Nil with vsf have vs1: "vs1 = berlekamp_basis f" by auto from Nil(3) have us: "us = us1 @ divs" by auto from Nil(4,6) have md: "\ u. u \ set us \ monic u \ degree u > 0" unfolding us by (auto simp: irreducible_degree_field) from Nil(7)[unfolded vs1] us have no_further_splitting_possible: "\ u v i. v \ set (berlekamp_basis f) \ u \ set us \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1, u}" by auto from Nil(5) us have prod: "f = prod_list us" by simp show ?case proof (intro conjI ballI) fix u assume u: "u \ set us" from md[OF this] have mon_u: "monic u" and deg_u: "degree u > 0" by auto from prod u have uf: "u dvd f" by (simp add: prod_list_dvd) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \P" "P \ {q. irreducible q \ monic q}" by auto show "irreducible u" proof (rule ccontr) assume irr_u: "\ irreducible u" from not_irreducible_factor_yields_prime_factors[OF uf P deg_u this] obtain pi pj where pij: "pi \ P" "pj \ P" "pi \ pj" "pi dvd u" "pj dvd u" by blast from exists_vector_in_Berlekamp_basis_dvd[OF deg_f berlekamp_basis_basis[OF deg_f, folded vs1] finite_set P pij(1-3) mon_f sf_f irr_u uf mon_u pij(4-5), unfolded vs1] obtain v s where v: "v \ set (berlekamp_basis f)" and gcd: "gcd u (v - [:s:]) \ {1,u}" using is_unit_gcd by auto from surj_of_nat_mod_ring[of s] obtain i where i: "i < CARD('a)" and s: "s = of_nat i" by auto from no_further_splitting_possible[OF v u i] gcd[unfolded s] show False by auto qed qed (insert prod md, auto) qed qed lemma berlekamp_monic_factorization: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and us: "berlekamp_monic_factorization d f = us" and d: "\ g. g dvd f \ degree g = d \ irreducible g" and deg: "degree f > 0" and mon: "monic f" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - from us[unfolded berlekamp_monic_factorization_def Let_def] deg have us: "us = [] @ berlekamp_factorization_main d [f] (berlekamp_basis f) (length (berlekamp_basis f))" by (auto) have id: "berlekamp_basis f = [] @ berlekamp_basis f" "length (berlekamp_basis f) = length [] + length (berlekamp_basis f)" "f = prod_list ([] @ [f])" by auto show "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" by (rule berlekamp_factorization_main[OF sf_f id(1) refl refl id(2) us _ _ _ id(3)], insert mon deg d, auto) qed end end diff --git a/thys/Berlekamp_Zassenhaus/Chinese_Remainder_Poly.thy b/thys/Berlekamp_Zassenhaus/Chinese_Remainder_Poly.thy --- a/thys/Berlekamp_Zassenhaus/Chinese_Remainder_Poly.thy +++ b/thys/Berlekamp_Zassenhaus/Chinese_Remainder_Poly.thy @@ -1,258 +1,258 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) subsection \Chinese Remainder Theorem for Polynomials\ text \We prove the Chinese Remainder Theorem, and strengthen it by showing uniqueness\ theory Chinese_Remainder_Poly imports "HOL-Number_Theory.Residues" Polynomial_Factorization.Polynomial_Divisibility Polynomial_Interpolation.Missing_Polynomial begin lemma cong_add_poly: - "[(a::'b::{factorial_ring_gcd,field} poly) = b] (mod m) \ [c = d] (mod m) \ [a + c = b + d] (mod m)" + "[(a::'b::{field_gcd} poly) = b] (mod m) \ [c = d] (mod m) \ [a + c = b + d] (mod m)" by (fact cong_add) lemma cong_mult_poly: - "[(a::'b::{field, factorial_ring_gcd} poly) = b] (mod m) \ [c = d] (mod m) \ [a * c = b * d] (mod m)" + "[(a::'b::{field_gcd} poly) = b] (mod m) \ [c = d] (mod m) \ [a * c = b * d] (mod m)" by (fact cong_mult) -lemma cong_mult_self_poly: "[(a::'b::{factorial_ring_gcd,field} poly) * m = 0] (mod m)" +lemma cong_mult_self_poly: "[(a::'b::{field_gcd} poly) * m = 0] (mod m)" by (fact cong_mult_self_right) -lemma cong_scalar2_poly: "[(a::'b::{field, factorial_ring_gcd} poly)= b] (mod m) \ [k * a = k * b] (mod m)" +lemma cong_scalar2_poly: "[(a::'b::{field_gcd} poly)= b] (mod m) \ [k * a = k * b] (mod m)" by (fact cong_scalar_left) lemma cong_sum_poly: - "(\x. x \ A \ [((f x)::'b::{factorial_ring_gcd,field} poly) = g x] (mod m)) \ + "(\x. x \ A \ [((f x)::'b::{field_gcd} poly) = g x] (mod m)) \ [(\x\A. f x) = (\x\A. g x)] (mod m)" by (rule cong_sum) -lemma cong_iff_lin_poly: "([(a::'b::{factorial_ring_gcd,field} poly) = b] (mod m)) = (\k. b = a + m * k)" +lemma cong_iff_lin_poly: "([(a::'b::{field_gcd} poly) = b] (mod m)) = (\k. b = a + m * k)" using cong_diff_iff_cong_0 [of b a m] by (auto simp add: cong_0_iff dvd_def algebra_simps dest: cong_sym) -lemma cong_solve_poly: "(a::'b::{normalization_euclidean_semiring, factorial_ring_gcd,field} poly) \ 0 \ \x. [a * x = gcd a n] (mod n)" +lemma cong_solve_poly: "(a::'b::{field_gcd} poly) \ 0 \ \x. [a * x = gcd a n] (mod n)" proof (cases "n = 0") case True note n0=True show ?thesis proof (cases "monic a") case True have n: "normalize a = a" by (rule normalize_monic[OF True]) show ?thesis by (rule exI[of _ 1], auto simp add: n0 n cong_def) next case False show ?thesis by (auto simp add: True cong_def normalize_poly_old_def map_div_is_smult_inverse) (metis mult.right_neutral mult_smult_right) qed next case False note n_not_0 = False show ?thesis using bezout_coefficients_fst_snd [of a n, symmetric] by (auto simp add: cong_iff_lin_poly mult.commute [of a] mult.commute [of n]) qed lemma cong_solve_coprime_poly: -assumes coprime_an:"coprime (a::'b::{normalization_euclidean_semiring, factorial_ring_gcd,field} poly) n" +assumes coprime_an:"coprime (a::'b::{field_gcd} poly) n" shows "\x. [a * x = 1] (mod n)" proof (cases "a = 0") case True show ?thesis unfolding cong_def using True coprime_an by auto next case False show ?thesis using coprime_an cong_solve_poly[OF False, of n] unfolding cong_def by presburger qed lemma cong_dvd_modulus_poly: - "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" for x y :: "'b::{factorial_ring_gcd,field} poly" + "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" for x y :: "'b::{field_gcd} poly" by (auto simp add: cong_iff_lin_poly elim!: dvdE) lemma chinese_remainder_aux_poly: fixes A :: "'a set" - and m :: "'a \ 'b::{normalization_euclidean_semiring,factorial_ring_gcd,field} poly" + and m :: "'a \ 'b::{field_gcd} poly" assumes fin: "finite A" and cop: "\i \ A. (\j \ A. i \ j \ coprime (m i) (m j))" shows "\b. (\i \ A. [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j)))" proof (rule finite_set_choice, rule fin, rule ballI) fix i assume "i : A" with cop have "coprime (\j \ A - {i}. m j) (m i)" by (auto intro: prod_coprime_left) then have "\x. [(\j \ A - {i}. m j) * x = 1] (mod m i)" by (elim cong_solve_coprime_poly) then obtain x where "[(\j \ A - {i}. m j) * x = 1] (mod m i)" by auto moreover have "[(\j \ A - {i}. m j) * x = 0] (mod (\j \ A - {i}. m j))" by (subst mult.commute, rule cong_mult_self_poly) ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod prod m (A - {i}))" by blast qed (*The Chinese Remainder Theorem for polynomials: *) lemma chinese_remainder_poly: fixes A :: "'a set" - and m :: "'a \ 'b::{normalization_euclidean_semiring,factorial_ring_gcd,field} poly" + and m :: "'a \ 'b::{field_gcd} poly" and u :: "'a \ 'b poly" assumes fin: "finite A" and cop: "\i\A. (\j\A. i \ j \ coprime (m i) (m j))" shows "\x. (\i\A. [x = u i] (mod m i))" proof - from chinese_remainder_aux_poly [OF fin cop] obtain b where bprop: "\i\A. [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j))" by blast let ?x = "\i\A. (u i) * (b i)" show "?thesis" proof (rule exI, clarify) fix i assume a: "i : A" show "[?x = u i] (mod m i)" proof - from fin a have "?x = (\j \ {i}. u j * b j) + (\j \ A - {i}. u j * b j)" by (subst sum.union_disjoint [symmetric], auto intro: sum.cong) then have "[?x = u i * b i + (\j \ A - {i}. u j * b j)] (mod m i)" unfolding cong_def by auto also have "[u i * b i + (\j \ A - {i}. u j * b j) = u i * 1 + (\j \ A - {i}. u j * 0)] (mod m i)" apply (rule cong_add_poly) apply (rule cong_scalar2_poly) using bprop a apply blast apply (rule cong_sum) apply (rule cong_scalar2_poly) using bprop apply auto apply (rule cong_dvd_modulus_poly) apply (drule (1) bspec) apply (erule conjE) apply assumption apply rule using fin a apply auto done thus ?thesis by (metis (no_types, lifting) a add.right_neutral fin mult_cancel_left1 mult_cancel_right1 sum.not_neutral_contains_not_neutral sum.remove) qed qed qed (*********************** Now we try to prove the uniqueness **********************) lemma cong_trans_poly: - "[(a::'b::{factorial_ring_gcd,field} poly) = b] (mod m) \ [b = c] (mod m) \ [a = c] (mod m)" + "[(a::'b::{field_gcd} poly) = b] (mod m) \ [b = c] (mod m) \ [a = c] (mod m)" by (fact cong_trans) -lemma cong_mod_poly: "(n::'b::{factorial_ring_gcd,field} poly) ~= 0 \ [a mod n = a] (mod n)" +lemma cong_mod_poly: "(n::'b::{field_gcd} poly) ~= 0 \ [a mod n = a] (mod n)" by auto -lemma cong_sym_poly: "[(a::'b::{factorial_ring_gcd,field} poly) = b] (mod m) \ [b = a] (mod m)" +lemma cong_sym_poly: "[(a::'b::{field_gcd} poly) = b] (mod m) \ [b = a] (mod m)" by (fact cong_sym) -lemma cong_1_poly: "[(a::'b::{factorial_ring_gcd,field} poly) = b] (mod 1)" +lemma cong_1_poly: "[(a::'b::{field_gcd} poly) = b] (mod 1)" by (fact cong_1) lemma coprime_cong_mult_poly: - assumes "[(a::'b::{factorial_ring_gcd,field} poly) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n" + assumes "[(a::'b::{field_gcd} poly) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n" shows "[a = b] (mod m * n)" using divides_mult assms by (metis (no_types, hide_lams) cong_dvd_modulus_poly cong_iff_lin_poly dvd_mult2 dvd_refl minus_add_cancel mult.right_neutral) lemma coprime_cong_prod_poly: "(\i\A. (\j\A. i \ j \ coprime (m i) (m j))) \ - (\i\A. [(x::'b::{factorial_ring_gcd,field} poly) = y] (mod m i)) \ + (\i\A. [(x::'b::{field_gcd} poly) = y] (mod m i)) \ [x = y] (mod (\i\A. m i))" apply (induct A rule: infinite_finite_induct) apply auto apply (metis coprime_cong_mult_poly prod_coprime_right) done lemma cong_less_modulus_unique_poly: - "[(x::'b::{factorial_ring_gcd,field} poly) = y] (mod m) \ degree x < degree m \ degree y < degree m \ x = y" + "[(x::'b::{field_gcd} poly) = y] (mod m) \ degree x < degree m \ degree y < degree m \ x = y" by (simp add: cong_def mod_poly_less) lemma chinese_remainder_unique_poly: fixes A :: "'a set" - and m :: "'a \ 'b::{normalization_euclidean_semiring,factorial_ring_gcd,field} poly" + and m :: "'a \ 'b::{field_gcd} poly" and u :: "'a \ 'b poly" assumes nz: "\i\A. (m i) \ 0" and cop: "\i\A. (\j\A. i \ j \ coprime (m i) (m j))" (*The following assumption should not be necessary, but I need it since in Isabelle degree 0 is 0 instead of -\*) and not_constant: "0 < degree (prod m A)" shows "\!x. degree x < (\i\A. degree (m i)) \ (\i\A. [x = u i] (mod m i))" proof - from not_constant have fin: "finite A" by (metis degree_1 gr_implies_not0 prod.infinite) from chinese_remainder_poly [OF fin cop] obtain y where one: "(\i\A. [y = u i] (mod m i))" by blast let ?x = "y mod (\i\A. m i)" have degree_prod_sum: "degree (prod m A) = (\i\A. degree (m i))" by (rule degree_prod_eq_sum_degree[OF nz]) from fin nz have prodnz: "(\i\A. (m i)) \ 0" by auto (*This would hold without the premise not_constant if degree 0 = -\*) have less: "degree ?x < (\i\A. degree (m i))" unfolding degree_prod_sum[symmetric] using degree_mod_less[OF prodnz, of y] using not_constant by auto have cong: "\i\A. [?x = u i] (mod m i)" apply auto apply (rule cong_trans_poly) prefer 2 using one apply auto apply (rule cong_dvd_modulus_poly) apply (rule cong_mod_poly) using prodnz apply auto apply rule apply (rule fin) apply assumption done have unique: "\z. degree z < (\i\A. degree (m i)) \ (\i\A. [z = u i] (mod m i)) \ z = ?x" proof (clarify) fix z::"'b poly" assume zless: "degree z < (\i\A. degree (m i))" assume zcong: "(\i\A. [z = u i] (mod m i))" have deg1: "degree z < degree (prod m A)" using degree_prod_sum zless by simp have deg2: "degree ?x < degree (prod m A)" by (metis deg1 degree_0 degree_mod_less gr0I gr_implies_not0) have "\i\A. [?x = z] (mod m i)" apply clarify apply (rule cong_trans_poly) using cong apply (erule bspec) apply (rule cong_sym_poly) using zcong by auto with fin cop have "[?x = z] (mod (\i\A. m i))" by (intro coprime_cong_prod_poly) auto with zless show "z = ?x" apply (intro cong_less_modulus_unique_poly) apply (erule cong_sym_poly) apply (auto simp add: deg1 deg2) done qed from less cong unique show ?thesis by blast qed end diff --git a/thys/Berlekamp_Zassenhaus/Code_Abort_Gcd.thy b/thys/Berlekamp_Zassenhaus/Code_Abort_Gcd.thy --- a/thys/Berlekamp_Zassenhaus/Code_Abort_Gcd.thy +++ b/thys/Berlekamp_Zassenhaus/Code_Abort_Gcd.thy @@ -1,28 +1,29 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) theory Code_Abort_Gcd imports "HOL-Computational_Algebra.Polynomial_Factorial" begin text \Dummy code-setup for @{const Gcd} and @{const Lcm} in the presence of Container.\ definition dummy_Gcd where "dummy_Gcd x = Gcd x" definition dummy_Lcm where "dummy_Lcm x = Lcm x" declare [[code abort: dummy_Gcd]] lemma dummy_Gcd_Lcm: "Gcd x = dummy_Gcd x" "Lcm x = dummy_Lcm x" unfolding dummy_Gcd_def dummy_Lcm_def by auto -lemmas dummy_Gcd_Lcm_poly [code] = dummy_Gcd_Lcm [where ?'a = "'a :: factorial_ring_gcd poly"] +lemmas dummy_Gcd_Lcm_poly [code] = dummy_Gcd_Lcm + [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] lemmas dummy_Gcd_Lcm_int [code] = dummy_Gcd_Lcm [where ?'a = int] lemmas dummy_Gcd_Lcm_nat [code] = dummy_Gcd_Lcm [where ?'a = nat] declare [[code abort: Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm]] end diff --git a/thys/Berlekamp_Zassenhaus/Factorize_Int_Poly.thy b/thys/Berlekamp_Zassenhaus/Factorize_Int_Poly.thy --- a/thys/Berlekamp_Zassenhaus/Factorize_Int_Poly.thy +++ b/thys/Berlekamp_Zassenhaus/Factorize_Int_Poly.thy @@ -1,727 +1,727 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) subsection \Factoring Arbitrary Integer Polynomials\ text \We combine the factorization algorithm for square-free integer polynomials with a square-free factorization algorithm to a factorization algorithm for integer polynomials which does not make any assumptions.\ theory Factorize_Int_Poly imports Berlekamp_Zassenhaus Square_Free_Factorization_Int begin hide_const coeff monom lifting_forget poly.lifting typedef int_poly_factorization_algorithm = "{alg. \ (f :: int poly) fs. square_free f \ degree f > 0 \ alg f = fs \ (f = prod_list fs \ (\ fi \ set fs. irreducible\<^sub>d fi))}" by (rule exI[of _ berlekamp_zassenhaus_factorization], insert berlekamp_zassenhaus_factorization_irreducible\<^sub>d, auto) setup_lifting type_definition_int_poly_factorization_algorithm lift_definition int_poly_factorization_algorithm :: "int_poly_factorization_algorithm \ (int poly \ int poly list)" is "\ x. x" . lemma int_poly_factorization_algorithm_irreducible\<^sub>d: assumes "int_poly_factorization_algorithm alg f = fs" and "square_free f" and "degree f > 0" shows "f = prod_list fs \ (\ fi \ set fs. irreducible\<^sub>d fi)" using assms by (transfer, auto) corollary int_poly_factorization_algorithm_irreducible: assumes res: "int_poly_factorization_algorithm alg f = fs" and sf: "square_free f" and deg: "degree f > 0" and pr: "primitive f" shows "f = prod_list fs \ (\ fi \ set fs. irreducible fi \ degree fi > 0 \ primitive fi)" proof (intro conjI ballI) note * = int_poly_factorization_algorithm_irreducible\<^sub>d[OF res sf deg] from * show f: "f = prod_list fs" by auto fix fi assume fi: "fi \ set fs" with primitive_prod_list[OF pr[unfolded f]] show "primitive fi" by auto from irreducible_primitive_connect[OF this] * pr[unfolded f] fi show "irreducible fi" by auto from * fi show "degree fi > 0" by (auto) qed lemma irreducible_imp_square_free: assumes irr: "irreducible (p::'a::idom poly)" shows "square_free p" proof(intro square_freeI) from irr show p0: "p \ 0" by auto fix a assume "a * a dvd p" then obtain b where paab: "p = a * (a * b)" by (elim dvdE, auto) assume "degree a > 0" then have a1: "\ a dvd 1" by (auto simp: poly_dvd_1) then have ab1: "\ a * b dvd 1" using dvd_mult_left by auto from paab irr a1 ab1 show False by force qed (* TODO: Move *) lemma not_mem_set_dropWhileD: "x \ set (dropWhile P xs) \ x \ set xs \ P x" by (metis dropWhile_append3 in_set_conv_decomp) lemma primitive_reflect_poly: fixes f :: "'a :: comm_semiring_1 poly" shows "primitive (reflect_poly f) = primitive f" proof- have "(\ a \ set (coeffs f). x dvd a) \ (\a \ set (dropWhile ((=) 0) (coeffs f)). x dvd a)" for x by (auto dest: not_mem_set_dropWhileD set_dropWhileD) then show ?thesis by (auto simp: primitive_def coeffs_reflect_poly) qed (* TODO: move *) lemma gcd_list_sub: assumes "set xs \ set ys" shows "gcd_list ys dvd gcd_list xs" by (metis Gcd_fin.subset assms semiring_gcd_class.gcd_dvd1) lemma content_reflect_poly: "content (reflect_poly f) = content f" (is "?l = ?r") proof- have l: "?l = gcd_list (dropWhile ((=) 0) (coeffs f))" (is "_ = gcd_list ?xs") by (simp add: content_def reflect_poly_def) have "set ?xs \ set (coeffs f)" by (auto dest: set_dropWhileD) from gcd_list_sub[OF this] have "?r dvd gcd_list ?xs" by (simp add: content_def) with l have rl: "?r dvd ?l" by auto have "set (coeffs f) \ set (0 # ?xs)" by (auto dest: not_mem_set_dropWhileD) from gcd_list_sub[OF this] have "gcd_list ?xs dvd ?r" by (simp add: content_def) with l have lr: "?l dvd ?r" by auto from rl lr show "?l = ?r" by (simp add: associated_eqI) qed lemma coeff_primitive_part: "content f * coeff (primitive_part f) i = coeff f i" using arg_cong[OF content_times_primitive_part[of f], of "\f. coeff f _", unfolded coeff_smult]. (* TODO: move *) lemma smult_cancel[simp]: fixes c :: "'a :: idom" shows "smult c f = smult c g \ c = 0 \ f = g" proof- have l: "smult c f = [:c:] * f" by simp have r: "smult c g = [:c:] * g" by simp show ?thesis unfolding l r mult_cancel_left by simp qed lemma primitive_part_reflect_poly: fixes f :: "'a :: {semiring_gcd,idom} poly" shows "primitive_part (reflect_poly f) = reflect_poly (primitive_part f)" (is "?l = ?r") using content_times_primitive_part[of "reflect_poly f"] proof- note content_reflect_poly[of f, symmetric] also have "smult (content (reflect_poly f)) ?l = reflect_poly f" by simp also have "... = reflect_poly (smult (content f) (primitive_part f))" by simp finally show ?thesis unfolding reflect_poly_smult smult_cancel by auto qed (* TODO: move *) lemma reflect_poly_eq_zero[simp]: "reflect_poly f = 0 \ f = 0" proof assume "reflect_poly f = 0" then have "coeff (reflect_poly f) 0 = 0" by simp then have "lead_coeff f = 0" by simp then show "f = 0" by simp qed simp lemma irreducible\<^sub>d_reflect_poly_main: fixes f :: "'a :: {idom, semiring_gcd} poly" assumes nz: "coeff f 0 \ 0" and irr: "irreducible\<^sub>d (reflect_poly f)" shows "irreducible\<^sub>d f" proof let ?r = reflect_poly from irr degree_reflect_poly_eq[OF nz] show "degree f > 0" by auto fix g h assume deg: "degree g < degree f" "degree h < degree f" and fgh: "f = g * h" from arg_cong[OF fgh, of "\ f. coeff f 0"] nz have nz': "coeff g 0 \ 0" by (auto simp: coeff_mult_0) note rfgh = arg_cong[OF fgh, of reflect_poly, unfolded reflect_poly_mult[of g h]] from deg degree_reflect_poly_le[of g] degree_reflect_poly_le[of h] degree_reflect_poly_eq[OF nz] have "degree (?r h) < degree (?r f)" "degree (?r g) < degree (?r f)" by auto with irr rfgh show False by auto qed lemma irreducible\<^sub>d_reflect_poly: fixes f :: "'a :: {idom, semiring_gcd} poly" assumes nz: "coeff f 0 \ 0" shows "irreducible\<^sub>d (reflect_poly f) = irreducible\<^sub>d f" proof assume "irreducible\<^sub>d (reflect_poly f)" from irreducible\<^sub>d_reflect_poly_main[OF nz this] show "irreducible\<^sub>d f" . next from nz have nzr: "coeff (reflect_poly f) 0 \ 0" by auto assume "irreducible\<^sub>d f" with nz have "irreducible\<^sub>d (reflect_poly (reflect_poly f))" by simp from irreducible\<^sub>d_reflect_poly_main[OF nzr this] show "irreducible\<^sub>d (reflect_poly f)" . qed lemma irreducible_reflect_poly: fixes f :: "'a :: {idom,semiring_gcd} poly" assumes nz: "coeff f 0 \ 0" shows "irreducible (reflect_poly f) = irreducible f" (is "?l = ?r") proof (cases "degree f = 0") case True then obtain f0 where "f = [:f0:]" by (auto dest: degree0_coeffs) then show ?thesis by simp next case deg: False show ?thesis proof (cases "primitive f") case False with deg irreducible_imp_primitive[of f] irreducible_imp_primitive[of "reflect_poly f"] nz show ?thesis unfolding primitive_reflect_poly by auto next case cf: True let ?r = "reflect_poly" from nz have nz': "coeff (?r f) 0 \ 0" by auto let ?ir = irreducible\<^sub>d from irreducible\<^sub>d_reflect_poly[OF nz] irreducible\<^sub>d_reflect_poly[OF nz'] nz have "?ir f \ ?ir (reflect_poly f)" by auto also have "... \ irreducible (reflect_poly f)" by (rule irreducible_primitive_connect, unfold primitive_reflect_poly, fact cf) finally show ?thesis by (unfold irreducible_primitive_connect[OF cf], auto) qed qed (* TODO: Move *) lemma reflect_poly_dvd: "(f :: 'a :: idom poly) dvd g \ reflect_poly f dvd reflect_poly g" unfolding dvd_def by (auto simp: reflect_poly_mult) lemma square_free_reflect_poly: fixes f :: "'a :: idom poly" assumes sf: "square_free f" and nz: "coeff f 0 \ 0" shows "square_free (reflect_poly f)" unfolding square_free_def proof (intro allI conjI impI notI) let ?r = reflect_poly from sf[unfolded square_free_def] have f0: "f \ 0" and sf: "\ q. 0 < degree q \ q * q dvd f \ False" by auto from f0 nz show "?r f = 0 \ False" by auto fix q assume 0: "0 < degree q" and dvd: "q * q dvd ?r f" from dvd have "q dvd ?r f" by auto then obtain x where id: "?r f = q * x" by fastforce { assume "coeff q 0 = 0" hence "coeff (?r f) 0 = 0" using id by (auto simp: coeff_mult) with nz have False by auto } hence nzq: "coeff q 0 \ 0" by auto from dvd have "?r (q * q) dvd ?r (?r f)" by (rule reflect_poly_dvd) also have "?r (?r f) = f" using nz by auto also have "?r (q * q) = ?r q * ?r q" by (rule reflect_poly_mult) finally have "?r q * ?r q dvd f" . from sf[OF _ this] 0 nzq show False by simp qed -lemma gcd_reflect_poly: fixes f :: "'a :: factorial_ring_gcd poly" +lemma gcd_reflect_poly: fixes f :: "'a :: {factorial_ring_gcd, semiring_gcd_mult_normalize} poly" assumes nz: "coeff f 0 \ 0" "coeff g 0 \ 0" shows "gcd (reflect_poly f) (reflect_poly g) = normalize (reflect_poly (gcd f g))" proof (rule sym, rule gcdI) have "gcd f g dvd f" by auto from reflect_poly_dvd[OF this] show "normalize (reflect_poly (gcd f g)) dvd reflect_poly f" by simp have "gcd f g dvd g" by auto from reflect_poly_dvd[OF this] show "normalize (reflect_poly (gcd f g)) dvd reflect_poly g" by simp show "normalize (normalize (reflect_poly (gcd f g))) = normalize (reflect_poly (gcd f g))" by auto fix h assume hf: "h dvd reflect_poly f" and hg: "h dvd reflect_poly g" from hf obtain k where "reflect_poly f = h * k" unfolding dvd_def by auto from arg_cong[OF this, of "\ f. coeff f 0", unfolded coeff_mult_0] nz(1) have h: "coeff h 0 \ 0" by auto from reflect_poly_dvd[OF hf] reflect_poly_dvd[OF hg] have "reflect_poly h dvd f" "reflect_poly h dvd g" using nz by auto hence "reflect_poly h dvd gcd f g" by auto from reflect_poly_dvd[OF this] h have "h dvd reflect_poly (gcd f g)" by auto thus "h dvd normalize (reflect_poly (gcd f g))" by auto qed lemma linear_primitive_irreducible: fixes f :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes deg: "degree f = 1" and cf: "primitive f" shows "irreducible f" proof (intro irreducibleI) fix a b assume fab: "f = a * b" with deg have a0: "a \ 0" and b0: "b \ 0" by auto from deg[unfolded fab] degree_mult_eq[OF this] have "degree a = 0 \ degree b = 0" by auto then show "a dvd 1 \ b dvd 1" proof assume "degree a = 0" then obtain a0 where a: "a = [:a0:]" by (auto dest:degree0_coeffs) with fab have "c \ set (coeffs f) \ a0 dvd c" for c by (cases "a0 = 0", auto simp: coeffs_smult) with cf show ?thesis by (auto dest: primitiveD simp: a) next assume "degree b = 0" then obtain b0 where b: "b = [:b0:]" by (auto dest:degree0_coeffs) with fab have "c \ set (coeffs f) \ b0 dvd c" for c by (cases "b0 = 0", auto simp: coeffs_smult) with cf show ?thesis by (auto dest: primitiveD simp: b) qed qed (insert deg, auto simp: poly_dvd_1) lemma square_free_factorization_last_coeff_nz: assumes sff: "square_free_factorization f (a, fs)" and mem: "(fi,i) \ set fs" and nz: "coeff f 0 \ 0" shows "coeff fi 0 \ 0" proof assume fi: "coeff fi 0 = 0" note sff_list = square_free_factorization_prod_list[OF sff] note sff = square_free_factorizationD[OF sff] from sff_list have "coeff f 0 = a * coeff (\(a, i)\fs. a ^ Suc i) 0" by simp with split_list[OF mem] fi have "coeff f 0 = 0" by (auto simp: coeff_mult) with nz show False by simp qed context fixes alg :: int_poly_factorization_algorithm begin (* main factorization algorithm for square-free, content-free, non-constant polynomial that do not have 0 as root, with special cases and reciprocal polynomials *) definition main_int_poly_factorization :: "int poly \ int poly list" where "main_int_poly_factorization f = (let df = degree f in if df = 1 then [f] else if abs (coeff f 0) < abs (coeff f df) \ \take reciprocal polynomial, if \f(0) < lc(f)\\ then map reflect_poly (int_poly_factorization_algorithm alg (reflect_poly f)) else int_poly_factorization_algorithm alg f)" (* preprocessing via square-free factorization *) definition internal_int_poly_factorization :: "int poly \ int \ (int poly \ nat) list" where "internal_int_poly_factorization f = ( case square_free_factorization_int f of (a,gis) \ (a, [ (h,i) . (g,i) \ gis, h \ main_int_poly_factorization g ]) )" lemma internal_int_poly_factorization_code[code]: "internal_int_poly_factorization f = ( case square_free_factorization_int f of (a,gis) \ (a, concat (map (\ (g,i). (map (\ f. (f,i)) (main_int_poly_factorization g))) gis)))" unfolding internal_int_poly_factorization_def by auto (* factorization for polynomials that do not have 0 as root, with special treatment of polynomials of degree at most 1 *) definition factorize_int_last_nz_poly :: "int poly \ int \ (int poly \ nat) list" where "factorize_int_last_nz_poly f = (let df = degree f in if df = 0 then (coeff f 0, []) else if df = 1 then (content f,[(primitive_part f,0)]) else internal_int_poly_factorization f)" (* factorization for arbitrary polynomials *) definition factorize_int_poly_generic :: "int poly \ int \ (int poly \ nat) list" where "factorize_int_poly_generic f = (case x_split f of (n,g) \ \extract \x^n\\ \ if g = 0 then (0,[]) else case factorize_int_last_nz_poly g of (a,fs) \ if n = 0 then (a,fs) else (a, (monom 1 1, n - 1) # fs))" lemma factorize_int_poly_0[simp]: "factorize_int_poly_generic 0 = (0,[])" unfolding factorize_int_poly_generic_def x_split_def by simp lemma main_int_poly_factorization: assumes res: "main_int_poly_factorization f = fs" and sf: "square_free f" and df: "degree f > 0" and nz: "coeff f 0 \ 0" shows "f = prod_list fs \ (\ fi \ set fs. irreducible\<^sub>d fi)" proof (cases "degree f = 1") case True with res[unfolded main_int_poly_factorization_def Let_def] have "fs = [f]" by auto with True show ?thesis by auto next case False hence *: "(if degree f = 1 then t :: int poly list else e) = e" for t e by auto note res = res[unfolded main_int_poly_factorization_def Let_def *] show ?thesis proof (cases "abs (coeff f 0) < abs (coeff f (degree f))") case False with res have "int_poly_factorization_algorithm alg f = fs" by auto from int_poly_factorization_algorithm_irreducible\<^sub>d[OF this sf df] show ?thesis . next case True let ?f = "reflect_poly f" from square_free_reflect_poly[OF sf nz] have sf: "square_free ?f" . from nz df have df: "degree ?f > 0" by simp from True res obtain gs where fs: "fs = map reflect_poly gs" and gs: "int_poly_factorization_algorithm alg (reflect_poly f) = gs" by auto from int_poly_factorization_algorithm_irreducible\<^sub>d[OF gs sf df] have id: "reflect_poly ?f = reflect_poly (prod_list gs)" "?f = prod_list gs" and irr: "\ gi. gi \ set gs \ irreducible\<^sub>d gi" by auto from id(1) have f_fs: "f = prod_list fs" unfolding fs using nz by (simp add: reflect_poly_prod_list) { fix fi assume "fi \ set fs" from this[unfolded fs] obtain gi where gi: "gi \ set gs" and fi: "fi = reflect_poly gi" by auto { assume "coeff gi 0 = 0" with id(2) split_list[OF gi] have "coeff ?f 0 = 0" by (auto simp: coeff_mult) with nz have False by auto } hence nzg: "coeff gi 0 \ 0" by auto from irreducible\<^sub>d_reflect_poly[OF nzg] irr[OF gi] have "irreducible\<^sub>d fi" unfolding fi by simp } with f_fs show ?thesis by auto qed qed lemma internal_int_poly_factorization_mem: assumes f: "coeff f 0 \ 0" and res: "internal_int_poly_factorization f = (c,fs)" and mem: "(fi,i) \ set fs" shows "irreducible fi" "irreducible\<^sub>d fi" and "primitive fi" and "degree fi \ 0" proof - obtain a psi where a_psi: "square_free_factorization_int f = (a, psi)" by force from square_free_factorization_int[OF this] have sff: "square_free_factorization f (a, psi)" and cnt: "\ fi i. (fi, i) \ set psi \ primitive fi" by blast+ from square_free_factorization_last_coeff_nz[OF sff _ f] have nz_fi: "\ fi i. (fi, i) \ set psi \ coeff fi 0 \ 0" by auto note res = res[unfolded internal_int_poly_factorization_def a_psi Let_def split] obtain fact where fact: "fact = (\ (q,i :: nat). (map (\ f. (f,i)) (main_int_poly_factorization q)))" by auto from res[unfolded split Let_def] have c: "c = a" and fs: "fs = concat (map fact psi)" unfolding fact by auto note sff' = square_free_factorizationD[OF sff] from mem[unfolded fs, simplified] obtain d j where psi: "(d,j) \ set psi" and fi: "(fi, i) \ set (fact (d,j))" by auto obtain hs where d: "main_int_poly_factorization d = hs" by force from fi[unfolded d split fact] have fi: "fi \ set hs" by auto from main_int_poly_factorization[OF d _ _ nz_fi[OF psi]] sff'(2)[OF psi] cnt[OF psi] have main: "d = prod_list hs" "\ fi. fi \ set hs \ irreducible\<^sub>d fi" by auto from main split_list[OF fi] have "content fi dvd content d" by auto with cnt[OF psi] show cnt: "primitive fi" by simp from main(2)[OF fi] show irr: "irreducible\<^sub>d fi" . show "irreducible fi" using irreducible_primitive_connect[OF cnt] irr by blast from irr show "degree fi \ 0" by auto qed lemma internal_int_poly_factorization: assumes f: "coeff f 0 \ 0" and res: "internal_int_poly_factorization f = (c,fs)" shows "square_free_factorization f (c,fs)" proof - obtain a psi where a_psi: "square_free_factorization_int f = (a, psi)" by force from square_free_factorization_int[OF this] have sff: "square_free_factorization f (a, psi)" and pr: "\ fi i. (fi, i) \ set psi \ primitive fi" by blast+ obtain fact where fact: "fact = (\ (q,i :: nat). (map (\ f. (f,i)) (main_int_poly_factorization q)))" by auto from res[unfolded split Let_def] have c: "c = a" and fs: "fs = concat (map fact psi)" unfolding fact internal_int_poly_factorization_def a_psi by auto note sff' = square_free_factorizationD[OF sff] show ?thesis unfolding square_free_factorization_def split proof (intro conjI impI allI) show "f = 0 \ c = 0" "f = 0 \ fs = []" using sff'(4) unfolding c fs by auto { fix a i assume "(a,i) \ set fs" from irreducible_imp_square_free internal_int_poly_factorization_mem[OF f res this] show "square_free a" "degree a > 0" by auto } from square_free_factorization_last_coeff_nz[OF sff _ f] have nz: "\ fi i. (fi, i) \ set psi \ coeff fi 0 \ 0" by auto have eq: "f = smult c (\(a, i)\fs. a ^ Suc i)" unfolding prod.distinct_set_conv_list[OF sff'(5)] sff'(1) c proof (rule arg_cong[where f = "smult a"], unfold fs, insert sff'(2) nz, induct psi) case (Cons pi psi) obtain p i where pi: "pi = (p,i)" by force obtain gs where gs: "main_int_poly_factorization p = gs" by auto from Cons(2)[of p i] have p: "square_free p" "degree p > 0" unfolding pi by auto from Cons(3)[of p i] have nz: "coeff p 0 \ 0" unfolding pi by auto from main_int_poly_factorization[OF gs p nz] have pgs: "p = prod_list gs" by auto have fact: "fact (p,i) = map (\ g. (g,i)) gs" unfolding fact split gs by auto have cong: "\ x y X Y. x = X \ y = Y \ x * y = X * Y" by auto show ?case unfolding pi list.simps prod_list.Cons split fact concat.simps prod_list.append map_append proof (rule cong) show "p ^ Suc i = (\(a, i)\map (\g. (g, i)) gs. a ^ Suc i)" unfolding pgs by (induct gs, auto simp: ac_simps power_mult_distrib) show "(\(a, i)\psi. a ^ Suc i) = (\(a, i)\concat (map fact psi). a ^ Suc i)" by (rule Cons(1), insert Cons(2-3), auto) qed qed simp { fix i j l fi assume *: "j < length psi" "l < length (fact (psi ! j))" "fact (psi ! j) ! l = (fi, i)" from * have psi: "psi ! j \ set psi" by auto obtain d k where dk: "psi ! j = (d,k)" by force with * have psij: "psi ! j = (d,i)" unfolding fact split by auto from sff'(2)[OF psi[unfolded psij]] have d: "square_free d" "degree d > 0" by auto from nz[OF psi[unfolded psij]] have d0: "coeff d 0 \ 0" . from * psij fact have bz: "main_int_poly_factorization d = map fst (fact (psi ! j))" by (auto simp: o_def) from main_int_poly_factorization[OF bz d d0] pr[OF psi[unfolded dk]] have dhs: "d = prod_list (map fst (fact (psi ! j)))" by auto from * have mem: "fi \ set (map fst (fact (psi ! j)))" by (metis fst_conv image_eqI nth_mem set_map) from mem dhs psij d have "\ d. fi \ set (map fst (fact (psi ! j))) \ d = prod_list (map fst (fact (psi ! j))) \ psi ! j = (d, i) \ square_free d" by blast } note deconstruct = this { fix k K fi i Fi I assume k: "k < length fs" "K < length fs" and f: "fs ! k = (fi, i)" "fs ! K = (Fi, I)" and diff: "k \ K" from nth_concat_diff[OF k[unfolded fs] diff, folded fs, unfolded length_map] obtain j l J L where diff: "(j, l) \ (J, L)" and j: "j < length psi" "J < length psi" and l: "l < length (map fact psi ! j)" "L < length (map fact psi ! J)" and fs: "fs ! k = map fact psi ! j ! l" "fs ! K = map fact psi ! J ! L" by blast+ hence psij: "psi ! j \ set psi" by auto from j have id: "map fact psi ! j = fact (psi ! j)" "map fact psi ! J = fact (psi ! J)" by auto note l = l[unfolded id] note fs = fs[unfolded id] from j have psi: "psi ! j \ set psi" "psi ! J \ set psi" by auto from deconstruct[OF j(1) l(1) fs(1)[unfolded f, symmetric]] obtain d where mem: "fi \ set (map fst (fact (psi ! j)))" and d: "d = prod_list (map fst (fact (psi ! j)))" "psi ! j = (d, i)" "square_free d" by blast from deconstruct[OF j(2) l(2) fs(2)[unfolded f, symmetric]] obtain D where Mem: "Fi \ set (map fst (fact (psi ! J)))" and D: "D = prod_list (map fst (fact (psi ! J)))" "psi ! J = (D, I)" "square_free D" by blast from pr[OF psij[unfolded d(2)]] have cnt: "primitive d" . have "coprime fi Fi" proof (cases "J = j") case False from sff'(5) False j have "(d,i) \ (D,I)" unfolding distinct_conv_nth d(2)[symmetric] D(2)[symmetric] by auto from sff'(3)[OF psi[unfolded d(2) D(2)] this] have cop: "coprime d D" by auto from prod_list_dvd[OF mem, folded d(1)] have fid: "fi dvd d" by auto from prod_list_dvd[OF Mem, folded D(1)] have FiD: "Fi dvd D" by auto from coprime_divisors[OF fid FiD] cop show ?thesis by simp next case True note id = this from id diff have diff: "l \ L" by auto obtain bz where bz: "bz = map fst (fact (psi ! j))" by auto from fs[unfolded f] l have fi: "fi = bz ! l" "Fi = bz ! L" unfolding id bz by (metis fst_conv nth_map)+ from d[folded bz] have sf: "square_free (prod_list bz)" by auto from d[folded bz] cnt have cnt: "content (prod_list bz) = 1" by auto from l have l: "l < length bz" "L < length bz" unfolding bz id by auto from l fi have "fi \ set bz" by auto from content_dvd_1[OF cnt prod_list_dvd[OF this]] have cnt: "content fi = 1" . obtain g where g: "g = gcd fi Fi" by auto have g': "g dvd fi" "g dvd Fi" unfolding g by auto define bef where "bef = take l bz" define aft where "aft = drop (Suc l) bz" from id_take_nth_drop[OF l(1)] l have bz: "bz = bef @ fi # aft" and bef: "length bef = l" unfolding bef_def aft_def fi by auto with l diff have mem: "Fi \ set (bef @ aft)" unfolding fi(2) by (auto simp: nth_append) from split_list[OF this] obtain Bef Aft where ba: "bef @ aft = Bef @ Fi # Aft" by auto have "prod_list bz = fi * prod_list (bef @ aft)" unfolding bz by simp also have "prod_list (bef @ aft) = Fi * prod_list (Bef @ Aft)" unfolding ba by auto finally have "fi * Fi dvd prod_list bz" by auto with g' have "g * g dvd prod_list bz" by (meson dvd_trans mult_dvd_mono) with sf[unfolded square_free_def] have deg: "degree g = 0" by auto from content_dvd_1[OF cnt g'(1)] have cnt: "content g = 1" . from degree0_coeffs[OF deg] obtain c where gc: "g = [: c :]" by auto from cnt[unfolded gc content_def, simplified] have "abs c = 1" by (cases "c = 0", auto) with g gc have "gcd fi Fi \ {1,-1}" by fastforce thus "coprime fi Fi" by (auto intro!: gcd_eq_1_imp_coprime) (metis dvd_minus_iff dvd_refl is_unit_gcd_iff one_neq_neg_one) qed } note cop = this show dist: "distinct fs" unfolding distinct_conv_nth proof (intro impI allI) fix k K assume k: "k < length fs" "K < length fs" and diff: "k \ K" obtain fi i Fi I where f: "fs ! k = (fi,i)" "fs ! K = (Fi,I)" by force+ from cop[OF k f diff] have cop: "coprime fi Fi" . from k(1) f(1) have "(fi,i) \ set fs" unfolding set_conv_nth by force from internal_int_poly_factorization_mem[OF assms(1) res this] have "degree fi > 0" by auto hence "\ is_unit fi" by (simp add: poly_dvd_1) with cop coprime_id_is_unit[of fi] have "fi \ Fi" by auto thus "fs ! k \ fs ! K" unfolding f by auto qed show "f = smult c (\(a, i)\set fs. a ^ Suc i)" unfolding eq prod.distinct_set_conv_list[OF dist] by simp fix fi i Fi I assume mem: "(fi, i) \ set fs" "(Fi,I) \ set fs" and diff: "(fi, i) \ (Fi, I)" then obtain k K where k: "k < length fs" "K < length fs" and f: "fs ! k = (fi, i)" "fs ! K = (Fi, I)" unfolding set_conv_nth by auto with diff have diff: "k \ K" by auto from cop[OF k f diff] show "Rings.coprime fi Fi" by auto qed qed lemma factorize_int_last_nz_poly: assumes res: "factorize_int_last_nz_poly f = (c,fs)" and nz: "coeff f 0 \ 0" shows "square_free_factorization f (c,fs)" "(fi,i) \ set fs \ irreducible fi" "(fi,i) \ set fs \ degree fi \ 0" proof (atomize(full)) from nz have lz: "lead_coeff f \ 0" by auto note res = res[unfolded factorize_int_last_nz_poly_def Let_def] consider (0) "degree f = 0" | (1) "degree f = 1" | (2) "degree f > 1" by linarith then show "square_free_factorization f (c,fs) \ ((fi,i) \ set fs \ irreducible fi) \ ((fi,i) \ set fs \ degree fi \ 0)" proof cases case 0 from degree0_coeffs[OF 0] obtain a where f: "f = [:a:]" by auto from res show ?thesis unfolding square_free_factorization_def f by auto next case 1 then have irr: "irreducible (primitive_part f)" by (auto intro!: linear_primitive_irreducible content_primitive_part) from irreducible_imp_square_free[OF irr] have sf: "square_free (primitive_part f)" . from 1 have f0: "f \ 0" by auto from res irr sf f0 show ?thesis unfolding square_free_factorization_def by (auto simp: 1) next case 2 with res have "internal_int_poly_factorization f = (c,fs)" by auto from internal_int_poly_factorization[OF nz this] internal_int_poly_factorization_mem[OF nz this] show ?thesis by auto qed qed lemma factorize_int_poly: assumes res: "factorize_int_poly_generic f = (c,fs)" shows "square_free_factorization f (c,fs)" "(fi,i) \ set fs \ irreducible fi" "(fi,i) \ set fs \ degree fi \ 0" proof (atomize(full)) obtain n g where xs: "x_split f = (n,g)" by force obtain d hs where fact: "factorize_int_last_nz_poly g = (d,hs)" by force from res[unfolded factorize_int_poly_generic_def xs split fact] have res: "(if g = 0 then (0, []) else if n = 0 then (d, hs) else (d, (monom 1 1, n - 1) # hs)) = (c, fs)" . note xs = x_split[OF xs] show "square_free_factorization f (c,fs) \ ((fi,i) \ set fs \ irreducible fi) \ ((fi,i) \ set fs \ degree fi \ 0)" proof (cases "g = 0") case True hence "f = 0" "c = 0" "fs = []" using res xs by auto thus ?thesis unfolding square_free_factorization_def by auto next case False with xs have "\ monom 1 1 dvd g" by auto hence "coeff g 0 \ 0" by (simp add: monom_1_dvd_iff') note fact = factorize_int_last_nz_poly[OF fact this] let ?x = "monom 1 1 :: int poly" have x: "content ?x = 1 \ lead_coeff ?x = 1 \ degree ?x = 1" by (auto simp add: degree_monom_eq monom_Suc content_def) from res False have res: "(if n = 0 then (d, hs) else (d, (?x, n - 1) # hs)) = (c, fs)" by auto show ?thesis proof (cases n) case 0 with res xs have id: "fs = hs" "c = d" "f = g" by auto from fact show ?thesis unfolding id by auto next case (Suc m) with res have id: "c = d" "fs = (?x,m) # hs" by auto from Suc xs have fg: "f = monom 1 (Suc m) * g" and dvd: "\ ?x dvd g" by auto from x linear_primitive_irreducible[of ?x] have irr: "irreducible ?x" by auto from irreducible_imp_square_free[OF this] have sfx: "square_free ?x" . from irr fact have one: "(fi, i) \ set fs \ irreducible fi \ degree fi \ 0" unfolding id by (auto simp: degree_monom_eq) have fg: "f = ?x ^ n * g" unfolding fg Suc by (metis x_pow_n) from x have degx: "degree ?x = 1" by simp note sf = square_free_factorizationD[OF fact(1)] { fix a i assume ai: "(a,i) \ set hs" with sf(4) have g0: "g \ 0" by auto from split_list[OF ai] obtain ys zs where hs: "hs = ys @ (a,i) # zs" by auto have "a dvd g" unfolding square_free_factorization_prod_list[OF fact(1)] hs by (rule dvd_smult, simp add: ac_simps) moreover have "\ ?x dvd g" using xs[unfolded Suc] by auto ultimately have dvd: "\ ?x dvd a" using dvd_trans by blast from sf(2)[OF ai] have "a \ 0" by auto have "1 = gcd ?x a" proof (rule gcdI) fix d assume d: "d dvd ?x" "d dvd a" from content_dvd_contentI[OF d(1)] x have cnt: "is_unit (content d)" by auto show "is_unit d" proof (cases "degree d = 1") case False with divides_degree[OF d(1), unfolded degx] have "degree d = 0" by auto from degree0_coeffs[OF this] obtain c where dc: "d = [:c:]" by auto from cnt[unfolded dc] have "is_unit c" by (auto simp: content_def, cases "c = 0", auto) hence "d * d = 1" unfolding dc by (auto, cases "c = -1"; cases "c = 1", auto) thus "is_unit d" by (metis dvd_triv_right) next case True from d(1) obtain e where xde: "?x = d * e" unfolding dvd_def by auto from arg_cong[OF this, of degree] degx have "degree d + degree e = 1" by (metis True add.right_neutral degree_0 degree_mult_eq one_neq_zero) with True have "degree e = 0" by auto from degree0_coeffs[OF this] xde obtain e where xde: "?x = [:e:] * d" by auto from arg_cong[OF this, of content, unfolded content_mult] x have "content [:e:] * content d = 1" by auto also have "content [:e :] = abs e" by (auto simp: content_def, cases "e = 0", auto) finally have "\e\ * content d = 1" . from pos_zmult_eq_1_iff_lemma[OF this] have "e * e = 1" by (cases "e = 1"; cases "e = -1", auto) with arg_cong[OF xde, of "smult e"] have "d = ?x * [:e:]" by auto hence "?x dvd d" unfolding dvd_def by blast with d(2) have "?x dvd a" by (metis dvd_trans) with dvd show ?thesis by auto qed qed auto hence "coprime ?x a" by (simp add: gcd_eq_1_imp_coprime) note this dvd } note hs_dvd_x = this from hs_dvd_x[of ?x m] have nmem: "(?x,m) \ set hs" by auto hence eq: "?x ^ n * g = smult d (\(a, i)\set fs. a ^ Suc i)" unfolding sf(1) unfolding id Suc by simp have eq0: "?x ^ n * g = 0 \ g = 0" by simp have "square_free_factorization f (d,fs)" unfolding fg id(1) square_free_factorization_def split eq0 unfolding eq proof (intro conjI allI impI, rule refl) fix a i assume ai: "(a,i) \ set fs" thus "square_free a" "degree a > 0" using sf(2) sfx degx unfolding id by auto fix b j assume bj: "(b,j) \ set fs" and diff: "(a,i) \ (b,j)" consider (hs_hs) "(a,i) \ set hs" "(b,j) \ set hs" | (hs_x) "(a,i) \ set hs" "b = ?x" | (x_hs) "(b,j) \ set hs" "a = ?x" using ai bj diff unfolding id by auto thus "Rings.coprime a b" proof cases case hs_hs from sf(3)[OF this diff] show ?thesis . next case hs_x from hs_dvd_x(1)[OF hs_x(1)] show ?thesis unfolding hs_x(2) by (simp add: ac_simps) next case x_hs from hs_dvd_x(1)[OF x_hs(1)] show ?thesis unfolding x_hs(2) by simp qed next show "g = 0 \ d = 0" using sf(4) by auto show "g = 0 \ fs = []" using sf(4) xs Suc by auto show "distinct fs" using sf(5) nmem unfolding id by auto qed thus ?thesis using one unfolding id by auto qed qed qed end lift_definition berlekamp_zassenhaus_factorization_algorithm :: int_poly_factorization_algorithm is berlekamp_zassenhaus_factorization using berlekamp_zassenhaus_factorization_irreducible\<^sub>d by blast abbreviation factorize_int_poly where "factorize_int_poly \ factorize_int_poly_generic berlekamp_zassenhaus_factorization_algorithm" end diff --git a/thys/Berlekamp_Zassenhaus/Factorize_Rat_Poly.thy b/thys/Berlekamp_Zassenhaus/Factorize_Rat_Poly.thy --- a/thys/Berlekamp_Zassenhaus/Factorize_Rat_Poly.thy +++ b/thys/Berlekamp_Zassenhaus/Factorize_Rat_Poly.thy @@ -1,117 +1,118 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) subsection \Factoring Rational Polynomials\ text \We combine the factorization algorithm for integer polynomials with Gauss Lemma to a factorization algorithm for rational polynomials.\ theory Factorize_Rat_Poly imports Factorize_Int_Poly begin (*TODO: Move*) -interpretation content_hom: monoid_mult_hom "content::'a::factorial_semiring_gcd poly \ _" +interpretation content_hom: monoid_mult_hom + "content::'a::{factorial_semiring, semiring_gcd, normalization_semidom_multiplicative} poly \ _" by (unfold_locales, auto simp: content_mult) lemma prod_dvd_1_imp_all_dvd_1: assumes "finite X" and "prod f X dvd 1" and "x \ X" shows "f x dvd 1" proof (insert assms, induct rule:finite_induct) case IH: (insert x' X) show ?case proof (cases "x = x'") case True with IH show ?thesis using dvd_trans[of "f x'" "f x' * _" 1] by (metis dvd_triv_left prod.insert) next case False then show ?thesis using IH by (auto intro!: IH(3) dvd_trans[of "prod f X" "_ * prod f X" 1]) qed qed simp context fixes alg :: int_poly_factorization_algorithm begin definition factorize_rat_poly_generic :: "rat poly \ rat \ (rat poly \ nat) list" where "factorize_rat_poly_generic f = (case rat_to_normalized_int_poly f of (c,g) \ case factorize_int_poly_generic alg g of (d,fs) \ (c * rat_of_int d, map (\ (fi,i). (map_poly rat_of_int fi, i)) fs))" lemma factorize_rat_poly_0[simp]: "factorize_rat_poly_generic 0 = (0,[])" unfolding factorize_rat_poly_generic_def rat_to_normalized_int_poly_def by simp lemma factorize_rat_poly: assumes res: "factorize_rat_poly_generic f = (c,fs)" shows "square_free_factorization f (c,fs)" and "(fi,i) \ set fs \ irreducible fi" proof(atomize(full), cases "f=0", goal_cases) case 1 with res show ?case by (auto simp: square_free_factorization_def) next case 2 show ?case proof (unfold square_free_factorization_def split, intro conjI impI allI) let ?r = rat_of_int let ?rp = "map_poly ?r" obtain d g where ri: "rat_to_normalized_int_poly f = (d,g)" by force obtain e gs where fi: "factorize_int_poly_generic alg g = (e,gs)" by force from res[unfolded factorize_rat_poly_generic_def ri fi split] have c: "c = d * ?r e" and fs: "fs = map (\ (fi,i). (?rp fi, i)) gs" by auto from factorize_int_poly[OF fi] have irr: "(fi, i) \ set gs \ irreducible fi \ content fi = 1" for fi i using irreducible_imp_primitive[of fi] by auto note sff = factorize_int_poly(1)[OF fi] note sff' = square_free_factorizationD[OF sff] { fix n f have "?rp (f ^ n) = (?rp f) ^ n" by (induct n, auto simp: hom_distribs) } note exp = this show dist: "distinct fs" using sff'(5) unfolding fs distinct_map inj_on_def by auto interpret mh: map_poly_inj_idom_hom rat_of_int.. have "f = smult d (?rp g)" using rat_to_normalized_int_poly[OF ri] by auto also have "\ = smult d (?rp (smult e (\(a, i)\set gs. a ^ Suc i)))" using sff'(1) by simp also have "\ = smult c (?rp (\(a, i)\set gs. a ^ Suc i))" unfolding c by (simp add: hom_distribs) also have "?rp (\(a, i)\set gs. a ^ Suc i) = (\(a, i)\set fs. a ^ Suc i)" unfolding prod.distinct_set_conv_list[OF sff'(5)] prod.distinct_set_conv_list[OF dist] unfolding fs by (insert exp, auto intro!: arg_cong[of _ _ "\x. prod_list (map x gs)"] simp: hom_distribs of_int_poly_hom.hom_prod_list) finally show f: "f = smult c (\(a, i)\set fs. a ^ Suc i)" by auto { fix a i assume ai: "(a,i) \ set fs" from ai obtain A where a: "a = ?rp A" and A: "(A,i) \ set gs" unfolding fs by auto fix b j assume "(b,j) \ set fs" and diff: "(a,i) \ (b,j)" from this(1) obtain B where b: "b = ?rp B" and B: "(B,j) \ set gs" unfolding fs by auto from diff[unfolded a b] have "(A,i) \ (B,j)" by auto from sff'(3)[OF A B this] show "Rings.coprime a b" by (auto simp add: coprime_iff_gcd_eq_1 gcd_rat_to_gcd_int a b) } { fix fi i assume "(fi,i) \ set fs" then obtain gi where fi: "fi = ?rp gi" and gi: "(gi,i) \ set gs" unfolding fs by auto from irr[OF gi] have cf_gi: "primitive gi" by auto then have "primitive (?rp gi)" by (auto simp: content_field_poly) note [simp] = irreducible_primitive_connect[OF cf_gi] irreducible_primitive_connect[OF this] show "irreducible fi" using irr[OF gi] fi irreducible\<^sub>d_int_rat[of gi,simplified] by auto then show "degree fi > 0" "square_free fi" unfolding fi by (auto intro: irreducible_imp_square_free) } { assume "f = 0" with ri have *: "d = 1" "g = 0" unfolding rat_to_normalized_int_poly_def by auto with sff'(4)[OF *(2)] show "c = 0" "fs = []" unfolding c fs by auto } qed qed end abbreviation factorize_rat_poly where "factorize_rat_poly \ factorize_rat_poly_generic berlekamp_zassenhaus_factorization_algorithm" end diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field.thy b/thys/Berlekamp_Zassenhaus/Finite_Field.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field.thy @@ -1,337 +1,350 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) section \Finite Rings and Fields\ text \We start by establishing some preliminary results about finite rings and finite fields\ subsection \Finite Rings\ theory Finite_Field imports "HOL-Computational_Algebra.Primes" "HOL-Number_Theory.Residues" Containers.Set_Impl Subresultants.Binary_Exponentiation Polynomial_Interpolation.Ring_Hom_Poly begin typedef ('a::finite) mod_ring = "{0..x\{0..x\{0..x\{0.. Rep_mod_ring (Abs_mod_ring xb)" using Rep_mod_ring atLeastLessThan_iff by blast assume xb1: "0 \ xb" and xb2: "xb < int CARD('a)" thus " Rep_mod_ring (Abs_mod_ring xb) < int CARD('a)" by (metis Abs_mod_ring_inverse Rep_mod_ring atLeastLessThan_iff le_less_trans linear) have xb: "xb \ {0..xa::'a mod_ring. (\x\{0.. xb = Rep_mod_ring xa" by (rule exI[of _ "Abs_mod_ring xb"], auto simp add: xb1 xb2, rule Abs_mod_ring_inverse[OF xb, symmetric]) qed ultimately show "bij_betw Rep_mod_ring {y. \x\{0.. 'a mod_ring \ bool" is "(=)" . instance by (intro_classes, transfer, auto) end instantiation mod_ring :: (finite) comm_ring begin lift_definition plus_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" is "\ x y. (x + y) mod int (CARD('a))" by simp lift_definition uminus_mod_ring :: "'a mod_ring \ 'a mod_ring" is "\ x. if x = 0 then 0 else int (CARD('a)) - x" by simp lift_definition minus_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" is "\ x y. (x - y) mod int (CARD('a))" by simp lift_definition times_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" is "\ x y. (x * y) mod int (CARD('a))" by simp lift_definition zero_mod_ring :: "'a mod_ring" is 0 by simp instance by standard (transfer; auto simp add: mod_simps algebra_simps intro: mod_diff_cong)+ end lift_definition to_int_mod_ring :: "'a::finite mod_ring \ int" is "\ x. x" . lift_definition of_int_mod_ring :: "int \ 'a::finite mod_ring" is "\ x. x mod int (CARD('a))" by simp interpretation to_int_mod_ring_hom: inj_zero_hom to_int_mod_ring by (unfold_locales; transfer, auto) lemma int_nat_card[simp]: "int (nat CARD('a::finite)) = CARD('a)" by auto interpretation of_int_mod_ring_hom: zero_hom of_int_mod_ring by (unfold_locales, transfer, auto) lemma of_int_mod_ring_to_int_mod_ring[simp]: "of_int_mod_ring (to_int_mod_ring x) = x" by (transfer, auto) lemma to_int_mod_ring_of_int_mod_ring[simp]: "0 \ x \ x < int CARD('a :: finite) \ to_int_mod_ring (of_int_mod_ring x :: 'a mod_ring) = x" by (transfer, auto) lemma range_to_int_mod_ring: "range (to_int_mod_ring :: ('a :: finite mod_ring \ int)) = {0 ..< CARD('a)}" apply (intro equalityI subsetI) apply (elim rangeE, transfer, force) by (auto intro!: range_eqI to_int_mod_ring_of_int_mod_ring[symmetric]) subsection \Nontrivial Finite Rings\ class nontriv = assumes nontriv: "CARD('a) > 1" subclass(in nontriv) finite by(intro_classes,insert nontriv,auto intro:card_ge_0_finite) instantiation mod_ring :: (nontriv) comm_ring_1 begin lift_definition one_mod_ring :: "'a mod_ring" is 1 using nontriv[where ?'a='a] by auto instance by (intro_classes; transfer, simp) end interpretation to_int_mod_ring_hom: inj_one_hom to_int_mod_ring by (unfold_locales, transfer, simp) lemma of_nat_of_int_mod_ring [code_unfold]: "of_nat = of_int_mod_ring o int" proof (rule ext, unfold o_def) show "of_nat n = of_int_mod_ring (int n)" for n proof (induct n) case (Suc n) show ?case by (simp only: of_nat_Suc Suc, transfer) (simp add: mod_simps) qed simp qed lemma of_nat_card_eq_0[simp]: "(of_nat (CARD('a::nontriv)) :: 'a mod_ring) = 0" by (unfold of_nat_of_int_mod_ring, transfer, auto) lemma of_int_of_int_mod_ring[code_unfold]: "of_int = of_int_mod_ring" proof (rule ext) fix x :: int obtain n1 n2 where x: "x = int n1 - int n2" by (rule int_diff_cases) show "of_int x = of_int_mod_ring x" unfolding x of_int_diff of_int_of_nat_eq of_nat_of_int_mod_ring o_def by (transfer, simp add: mod_diff_right_eq mod_diff_left_eq) qed unbundle lifting_syntax lemma pcr_mod_ring_to_int_mod_ring: "pcr_mod_ring = (\x y. x = to_int_mod_ring y)" unfolding mod_ring.pcr_cr_eq unfolding cr_mod_ring_def to_int_mod_ring.rep_eq .. lemma [transfer_rule]: "((=) ===> pcr_mod_ring) (\ x. int x mod int (CARD('a :: nontriv))) (of_nat :: nat \ 'a mod_ring)" by (intro rel_funI, unfold pcr_mod_ring_to_int_mod_ring of_nat_of_int_mod_ring, transfer, auto) lemma [transfer_rule]: "((=) ===> pcr_mod_ring) (\ x. x mod int (CARD('a :: nontriv))) (of_int :: int \ 'a mod_ring)" by (intro rel_funI, unfold pcr_mod_ring_to_int_mod_ring of_int_of_int_mod_ring, transfer, auto) lemma one_mod_card [simp]: "1 mod CARD('a::nontriv) = 1" using mod_less nontriv by blast lemma Suc_0_mod_card [simp]: "Suc 0 mod CARD('a::nontriv) = 1" using one_mod_card by simp lemma one_mod_card_int [simp]: "1 mod int CARD('a::nontriv) = 1" proof - from nontriv [where ?'a = 'a] have "int (1 mod CARD('a::nontriv)) = 1" by simp then show ?thesis using of_nat_mod [of 1 "CARD('a)", where ?'a = int] by simp qed lemma pow_mod_ring_transfer[transfer_rule]: "(pcr_mod_ring ===> (=) ===> pcr_mod_ring) (\a::int. \n. a^n mod CARD('a::nontriv)) ((^)::'a mod_ring \ nat \ 'a mod_ring)" unfolding pcr_mod_ring_to_int_mod_ring proof (intro rel_funI,simp) fix x::"'a mod_ring" and n show "to_int_mod_ring x ^ n mod int CARD('a) = to_int_mod_ring (x ^ n)" proof (induct n) case 0 thus ?case by auto next case (Suc n) have "to_int_mod_ring (x ^ Suc n) = to_int_mod_ring (x * x ^ n)" by auto also have "... = to_int_mod_ring x * to_int_mod_ring (x ^ n) mod CARD('a)" unfolding to_int_mod_ring_def using times_mod_ring.rep_eq by auto also have "... = to_int_mod_ring x * (to_int_mod_ring x ^ n mod CARD('a)) mod CARD('a)" using Suc.hyps by auto also have "... = to_int_mod_ring x ^ Suc n mod int CARD('a)" by (simp add: mod_simps) finally show ?case .. qed qed lemma dvd_mod_ring_transfer[transfer_rule]: "((pcr_mod_ring :: int \ 'a :: nontriv mod_ring \ bool) ===> (pcr_mod_ring :: int \ 'a mod_ring \ bool) ===> (=)) (\ i j. \k \ {0..k \ {0..k \ {0.. {0.. 'a mod_ring" where "inverse_mod_ring x = (if x = 0 then 0 else x ^ (nat (CARD('a) - 2)))" definition divide_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" where "divide_mod_ring x y = x * ((\c. if c = 0 then 0 else c ^ (nat (CARD('a) - 2))) y)" instance proof fix a b c::"'a mod_ring" show "inverse 0 = (0::'a mod_ring)" by (simp add: inverse_mod_ring_def) show "a div b = a * inverse b" unfolding inverse_mod_ring_def by (transfer', simp add: divide_mod_ring_def) show "a \ 0 \ inverse a * a = 1" proof (unfold inverse_mod_ring_def, transfer) let ?p="CARD('a)" fix x assume x: "x \ {0.. 0" have p0': "0\?p" by auto have "\ ?p dvd x" using x x0 zdvd_imp_le by fastforce then have "\ CARD('a) dvd nat \x\" by simp with x have "\ CARD('a) dvd nat x" by simp have rw: "x ^ nat (int (?p - 2)) * x = x ^ nat (?p - 1)" proof - have p2: "0 \ int (?p-2)" using x by simp have card_rw: "(CARD('a) - Suc 0) = nat (1 + int (CARD('a) - 2))" using nat_eq_iff x x0 by auto have "x ^ nat (?p - 2)*x = x ^ (Suc (nat (?p - 2)))" by simp also have "... = x ^ (nat (?p - 1))" using Suc_nat_eq_nat_zadd1[OF p2] card_rw by auto finally show ?thesis . qed have "[int (nat x ^ (CARD('a) - 1)) = int 1] (mod CARD('a))" using fermat_theorem [OF prime_card \\ CARD('a) dvd nat x\] by (simp only: cong_def cong_def of_nat_mod [symmetric]) then have *: "[x ^ (CARD('a) - 1) = 1] (mod CARD('a))" using x by auto have "x ^ (CARD('a) - 2) mod CARD('a) * x mod CARD('a) = (x ^ nat (CARD('a) - 2) * x) mod CARD('a)" by (simp add: mod_simps) also have "... = (x ^ nat (?p - 1) mod ?p)" unfolding rw by simp also have "... = (x ^ (nat ?p - 1) mod ?p)" using p0' by (simp add: nat_diff_distrib') also have "... = 1" using * by (simp add: cong_def) finally show "(if x = 0 then 0 else x ^ nat (int (CARD('a) - 2)) mod CARD('a)) * x mod CARD('a) = 1" using x0 by auto qed qed end instantiation mod_ring :: (prime_card) "{normalization_euclidean_semiring, euclidean_ring}" begin definition modulo_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" where "modulo_mod_ring x y = (if y = 0 then x else 0)" definition normalize_mod_ring :: "'a mod_ring \ 'a mod_ring" where "normalize_mod_ring x = (if x = 0 then 0 else 1)" definition unit_factor_mod_ring :: "'a mod_ring \ 'a mod_ring" where "unit_factor_mod_ring x = x" definition euclidean_size_mod_ring :: "'a mod_ring \ nat" where "euclidean_size_mod_ring x = (if x = 0 then 0 else 1)" instance proof (intro_classes) fix a :: "'a mod_ring" show "a \ 0 \ unit_factor a dvd 1" unfolding dvd_def unit_factor_mod_ring_def by (intro exI[of _ "inverse a"], auto) qed (auto simp: normalize_mod_ring_def unit_factor_mod_ring_def modulo_mod_ring_def euclidean_size_mod_ring_def field_simps) end instantiation mod_ring :: (prime_card) euclidean_ring_gcd begin definition gcd_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" where "gcd_mod_ring = Euclidean_Algorithm.gcd" definition lcm_mod_ring :: "'a mod_ring \ 'a mod_ring \ 'a mod_ring" where "lcm_mod_ring = Euclidean_Algorithm.lcm" definition Gcd_mod_ring :: "'a mod_ring set \ 'a mod_ring" where "Gcd_mod_ring = Euclidean_Algorithm.Gcd" definition Lcm_mod_ring :: "'a mod_ring set \ 'a mod_ring" where "Lcm_mod_ring = Euclidean_Algorithm.Lcm" instance by (intro_classes, auto simp: gcd_mod_ring_def lcm_mod_ring_def Gcd_mod_ring_def Lcm_mod_ring_def) end +instantiation mod_ring :: (prime_card) unique_euclidean_ring +begin + +definition [simp]: "division_segment_mod_ring (x :: 'a mod_ring) = (1 :: 'a mod_ring)" + +instance by intro_classes (auto simp: euclidean_size_mod_ring_def split: if_splits) + +end + +instance mod_ring :: (prime_card) field_gcd + by intro_classes auto + + lemma surj_of_nat_mod_ring: "\ i. i < CARD('a :: prime_card) \ (x :: 'a mod_ring) = of_nat i" by (rule exI[of _ "nat (to_int_mod_ring x)"], unfold of_nat_of_int_mod_ring o_def, subst nat_0_le, transfer, simp, simp, transfer, auto) lemma of_nat_0_mod_ring_dvd: assumes x: "of_nat x = (0 :: 'a ::prime_card mod_ring)" shows "CARD('a) dvd x" proof - let ?x = "of_nat x :: int" from x have "of_int_mod_ring ?x = (0 :: 'a mod_ring)" by (fold of_int_of_int_mod_ring, simp) hence "?x mod CARD('a) = 0" by (transfer, auto) hence "x mod CARD('a) = 0" by presburger thus ?thesis unfolding mod_eq_0_iff_dvd . qed end diff --git a/thys/Berlekamp_Zassenhaus/Hensel_Lifting.thy b/thys/Berlekamp_Zassenhaus/Hensel_Lifting.thy --- a/thys/Berlekamp_Zassenhaus/Hensel_Lifting.thy +++ b/thys/Berlekamp_Zassenhaus/Hensel_Lifting.thy @@ -1,1758 +1,1758 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) section \Hensel Lifting\ subsection \Properties about Factors\ text \We define and prove properties of Hensel-lifting. Here, we show the result that Hensel-lifting can lift a factorization mod $p$ to a factorization mod $p^n$. For the lifting we have proofs for both versions, the original linear Hensel-lifting or the quadratic approach from Zassenhaus. Via the linear version, we also show a uniqueness result, however only in the binary case, i.e., where $f = g \cdot h$. Uniqueness of the general case will later be shown in theory Berlekamp-Hensel by incorporating the factorization algorithm for finite fields algorithm.\ theory Hensel_Lifting imports "HOL-Computational_Algebra.Euclidean_Algorithm" Poly_Mod_Finite_Field_Record_Based Polynomial_Factorization.Square_Free_Factorization begin lemma uniqueness_poly_equality: - fixes f g :: "'a :: factorial_ring_gcd poly" + fixes f g :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes cop: "coprime f g" and deg: "B = 0 \ degree B < degree f" "B' = 0 \ degree B' < degree f" and f: "f \ 0" and eq: "A * f + B * g = A' * f + B' * g" shows "A = A'" "B = B'" proof - from eq have *: "(A - A') * f = (B' - B) * g" by (simp add: field_simps) hence "f dvd (B' - B) * g" unfolding dvd_def by (intro exI[of _ "A - A'"], auto simp: field_simps) with cop[simplified] have dvd: "f dvd (B' - B)" by (simp add: coprime_dvd_mult_right_iff ac_simps) from divides_degree[OF this] have "degree f \ degree (B' - B) \ B = B'" by auto with degree_diff_le_max[of B' B] deg show "B = B'" by auto with * f show "A = A'" by auto qed lemmas (in poly_mod_prime_type) uniqueness_poly_equality = uniqueness_poly_equality[where 'a="'a mod_ring", untransferred] lemmas (in poly_mod_prime) uniqueness_poly_equality = poly_mod_prime_type.uniqueness_poly_equality [unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemma pseudo_divmod_main_list_1_is_divmod_poly_one_main_list: "pseudo_divmod_main_list (1 :: 'a :: comm_ring_1) q f g n = divmod_poly_one_main_list q f g n" by (induct n arbitrary: q f g, auto simp: Let_def) lemma pdivmod_monic_pseudo_divmod: assumes g: "monic g" shows "pdivmod_monic f g = pseudo_divmod f g" proof - from g have id: "(coeffs g = []) = False" by auto from g have mon: "hd (rev (coeffs g)) = 1" by (metis coeffs_eq_Nil hd_rev id last_coeffs_eq_coeff_degree) show ?thesis unfolding pseudo_divmod_impl pseudo_divmod_list_def id if_False pdivmod_monic_def Let_def mon pseudo_divmod_main_list_1_is_divmod_poly_one_main_list by (auto split: prod.splits) qed lemma pdivmod_monic: assumes g: "monic g" and res: "pdivmod_monic f g = (q, r)" shows "f = g * q + r" "r = 0 \ degree r < degree g" proof - from g have g0: "g \ 0" by auto from pseudo_divmod[OF g0 res[unfolded pdivmod_monic_pseudo_divmod[OF g]], unfolded g] show "f = g * q + r" "r = 0 \ degree r < degree g" by auto qed definition dupe_monic :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly \ 'a poly \ 'a poly * 'a poly" where "dupe_monic D H S T U = (case pdivmod_monic (T * U) D of (q,r) \ (S * U + H * q, r))" lemma dupe_monic: assumes 1: "D*S + H*T = 1" and mon: "monic D" and dupe: "dupe_monic D H S T U = (A,B)" shows "A * D + B * H = U" "B = 0 \ degree B < degree D" proof - obtain Q R where div: "pdivmod_monic ((T * U)) D = (Q,R)" by force from dupe[unfolded dupe_monic_def div split] have A: "A = (S * U + H * Q)" and B: "B = R" by auto from pdivmod_monic[OF mon div] have TU: "T * U = D * Q + R" and deg: "R = 0 \ degree R < degree D" by auto hence R: "R = T * U - D * Q" by simp have "A * D + B * H = (D * S + H * T) * U" unfolding A B R by (simp add: field_simps) also have "\ = U" unfolding 1 by simp finally show eq: "A * D + B * H = U" . show "B = 0 \ degree B < degree D" using deg unfolding B . qed -lemma dupe_monic_unique: fixes D :: "'a :: factorial_ring_gcd poly" +lemma dupe_monic_unique: fixes D :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes 1: "D*S + H*T = 1" and mon: "monic D" and dupe: "dupe_monic D H S T U = (A,B)" and cop: "coprime D H" and other: "A' * D + B' * H = U" "B' = 0 \ degree B' < degree D" shows "A' = A" "B' = B" proof - from dupe_monic[OF 1 mon dupe] have one: "A * D + B * H = U" "B = 0 \ degree B < degree D" by auto from mon have D0: "D \ 0" by auto from uniqueness_poly_equality[OF cop one(2) other(2) D0, of A A', unfolded other, OF one(1)] show "A' = A" "B' = B" by auto qed context ring_ops begin lemma poly_rel_dupe_monic_i: assumes mon: "monic D" and rel: "poly_rel d D" "poly_rel h H" "poly_rel s S" "poly_rel t T" "poly_rel u U" shows "rel_prod poly_rel poly_rel (dupe_monic_i ops d h s t u) (dupe_monic D H S T U)" proof - note defs = dupe_monic_i_def dupe_monic_def note [transfer_rule] = rel have [transfer_rule]: "rel_prod poly_rel poly_rel (pdivmod_monic_i ops (times_poly_i ops t u) d) (pdivmod_monic (T * U) D)" by (rule poly_rel_pdivmod_monic[OF mon], transfer_prover+) show ?thesis unfolding defs by transfer_prover qed end context mod_ring_gen begin lemma monic_of_int_poly: "monic D \ monic (of_int_poly (Mp D) :: 'a mod_ring poly)" using Mp_f_representative Mp_to_int_poly monic_Mp by auto lemma dupe_monic_i: assumes dupe_i: "dupe_monic_i ff_ops d h s t u = (a,b)" and 1: "D*S + H*T =m 1" and mon: "monic D" and A: "A = to_int_poly_i ff_ops a" and B: "B = to_int_poly_i ff_ops b" and d: "Mp_rel_i d D" and h: "Mp_rel_i h H" and s: "Mp_rel_i s S" and t: "Mp_rel_i t T" and u: "Mp_rel_i u U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp_rel_i a A" "Mp_rel_i b B" proof - let ?I = "\ f. of_int_poly (Mp f) :: 'a mod_ring poly" let ?i = "to_int_poly_i ff_ops" note dd = Mp_rel_iD[OF d] note hh = Mp_rel_iD[OF h] note ss = Mp_rel_iD[OF s] note tt = Mp_rel_iD[OF t] note uu = Mp_rel_iD[OF u] obtain A' B' where dupe: "dupe_monic (?I D) (?I H) (?I S) (?I T) (?I U) = (A',B')" by force from poly_rel_dupe_monic_i[OF monic_of_int_poly[OF mon] dd(1) hh(1) ss(1) tt(1) uu(1), unfolded dupe_i dupe] have a: "poly_rel a A'" and b: "poly_rel b B'" by auto show aa: "Mp_rel_i a A" by (rule Mp_rel_iI'[OF a, folded A]) show bb: "Mp_rel_i b B" by (rule Mp_rel_iI'[OF b, folded B]) note Aa = Mp_rel_iD[OF aa] note Bb = Mp_rel_iD[OF bb] from poly_rel_inj[OF a Aa(1)] A have A: "A' = ?I A" by simp from poly_rel_inj[OF b Bb(1)] B have B: "B' = ?I B" by simp note Mp = dd(2) hh(2) ss(2) tt(2) uu(2) note [transfer_rule] = Mp have "(=) (D * S + H * T =m 1) (?I D * ?I S + ?I H * ?I T = 1)" by transfer_prover with 1 have 11: "?I D * ?I S + ?I H * ?I T = 1" by simp from dupe_monic[OF 11 monic_of_int_poly[OF mon] dupe, unfolded A B] have res: "?I A * ?I D + ?I B * ?I H = ?I U" "?I B = 0 \ degree (?I B) < degree (?I D)" by auto note [transfer_rule] = Aa(2) Bb(2) have "(=) (A * D + B * H =m U) (?I A * ?I D + ?I B * ?I H = ?I U)" "(=) (B =m 0 \ degree_m B < degree_m D) (?I B = 0 \ degree (?I B) < degree (?I D))" by transfer_prover+ with res have *: "A * D + B * H =m U" "B =m 0 \ degree_m B < degree_m D" by auto show "A * D + B * H =m U" by fact have B: "Mp B = B" using Mp_rel_i_Mp_to_int_poly_i assms(5) bb by blast from *(2) show "B = 0 \ degree B < degree D" unfolding B using degree_m_le[of D] by auto qed lemma Mp_rel_i_of_int_poly_i: assumes "Mp F = F" shows "Mp_rel_i (of_int_poly_i ff_ops F) F" by (metis Mp_f_representative Mp_rel_iI' assms poly_rel_of_int_poly to_int_poly_i) lemma dupe_monic_i_int: assumes dupe_i: "dupe_monic_i_int ff_ops D H S T U = (A,B)" and 1: "D*S + H*T =m 1" and mon: "monic D" and norm: "Mp D = D" "Mp H = H" "Mp S = S" "Mp T = T" "Mp U = U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" proof - let ?oi = "of_int_poly_i ff_ops" let ?ti = "to_int_poly_i ff_ops" note rel = norm[THEN Mp_rel_i_of_int_poly_i] obtain a b where dupe: "dupe_monic_i ff_ops (?oi D) (?oi H) (?oi S) (?oi T) (?oi U) = (a,b)" by force from dupe_i[unfolded dupe_monic_i_int_def this Let_def] have AB: "A = ?ti a" "B = ?ti b" by auto from dupe_monic_i[OF dupe 1 mon AB rel] Mp_rel_i_Mp_to_int_poly_i show "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" unfolding AB by auto qed end definition dupe_monic_dynamic :: "int \ int poly \ int poly \ int poly \ int poly \ int poly \ int poly \ int poly" where "dupe_monic_dynamic p = ( if p \ 65535 then dupe_monic_i_int (finite_field_ops32 (uint32_of_int p)) else if p \ 4294967295 then dupe_monic_i_int (finite_field_ops64 (uint64_of_int p)) else dupe_monic_i_int (finite_field_ops_integer (integer_of_int p)))" context poly_mod_2 begin lemma dupe_monic_i_int_finite_field_ops_integer: assumes dupe_i: "dupe_monic_i_int (finite_field_ops_integer (integer_of_int m)) D H S T U = (A,B)" and 1: "D*S + H*T =m 1" and mon: "monic D" and norm: "Mp D = D" "Mp H = H" "Mp S = S" "Mp T = T" "Mp U = U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" using m1 mod_ring_gen.dupe_monic_i_int[OF mod_ring_locale.mod_ring_finite_field_ops_integer[unfolded mod_ring_locale_def], internalize_sort "'a :: nontriv", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF _ assms] by auto lemma dupe_monic_i_int_finite_field_ops32: assumes m: "m \ 65535" and dupe_i: "dupe_monic_i_int (finite_field_ops32 (uint32_of_int m)) D H S T U = (A,B)" and 1: "D*S + H*T =m 1" and mon: "monic D" and norm: "Mp D = D" "Mp H = H" "Mp S = S" "Mp T = T" "Mp U = U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" using m1 mod_ring_gen.dupe_monic_i_int[OF mod_ring_locale.mod_ring_finite_field_ops32[unfolded mod_ring_locale_def], internalize_sort "'a :: nontriv", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF _ assms] by auto lemma dupe_monic_i_int_finite_field_ops64: assumes m: "m \ 4294967295" and dupe_i: "dupe_monic_i_int (finite_field_ops64 (uint64_of_int m)) D H S T U = (A,B)" and 1: "D*S + H*T =m 1" and mon: "monic D" and norm: "Mp D = D" "Mp H = H" "Mp S = S" "Mp T = T" "Mp U = U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" using m1 mod_ring_gen.dupe_monic_i_int[OF mod_ring_locale.mod_ring_finite_field_ops64[unfolded mod_ring_locale_def], internalize_sort "'a :: nontriv", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF _ assms] by auto lemma dupe_monic_dynamic: assumes dupe: "dupe_monic_dynamic m D H S T U = (A,B)" and 1: "D*S + H*T =m 1" and mon: "monic D" and norm: "Mp D = D" "Mp H = H" "Mp S = S" "Mp T = T" "Mp U = U" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" using dupe dupe_monic_i_int_finite_field_ops32[OF _ _ 1 mon norm, of A B] dupe_monic_i_int_finite_field_ops64[OF _ _ 1 mon norm, of A B] dupe_monic_i_int_finite_field_ops_integer[OF _ 1 mon norm, of A B] unfolding dupe_monic_dynamic_def by (auto split: if_splits) end context poly_mod begin definition dupe_monic_int :: "int poly \ int poly \ int poly \ int poly \ int poly \ int poly * int poly" where "dupe_monic_int D H S T U = (case pdivmod_monic (Mp (T * U)) D of (q,r) \ (Mp (S * U + H * q), Mp r))" end declare poly_mod.dupe_monic_int_def[code] text \Old direct proof on int poly. It does not permit to change implementation. This proof is still present, since we did not export the uniqueness part from the type-based uniqueness result @{thm dupe_monic_unique} via the various relations.\ lemma (in poly_mod_2) dupe_monic_int: assumes 1: "D*S + H*T =m 1" and mon: "monic D" and dupe: "dupe_monic_int D H S T U = (A,B)" shows "A * D + B * H =m U" "B = 0 \ degree B < degree D" "Mp A = A" "Mp B = B" "coprime_m D H \ A' * D + B' * H =m U \ B' = 0 \ degree B' < degree D \ Mp D = D \ Mp A' = A' \ Mp B' = B' \ prime m \ A' = A \ B' = B" proof - obtain Q R where div: "pdivmod_monic (Mp (T * U)) D = (Q,R)" by force from dupe[unfolded dupe_monic_int_def div split] have A: "A = Mp (S * U + H * Q)" and B: "B = Mp R" by auto from pdivmod_monic[OF mon div] have TU: "Mp (T * U) = D * Q + R" and deg: "R = 0 \ degree R < degree D" by auto hence "Mp R = Mp (Mp (T * U) - D * Q)" by simp also have "\ = Mp (T * U - Mp (Mp (Mp D * Q)))" unfolding Mp_Mp unfolding minus_Mp using minus_Mp mult_Mp by metis also have "\ = Mp (T * U - D * Q)" by simp finally have r: "Mp R = Mp (T * U - D * Q)" by simp have "Mp (A * D + B * H) = Mp (Mp (A * D) + Mp (B * H))" by simp also have "Mp (A * D) = Mp ((S * U + H * Q) * D)" unfolding A by simp also have "Mp (B * H) = Mp (Mp R * Mp H)" unfolding B by simp also have "\ = Mp ((T * U - D * Q) * H)" unfolding r by simp also have "Mp (Mp ((S * U + H * Q) * D) + Mp ((T * U - D * Q) * H)) = Mp ((S * U + H * Q) * D + (T * U - D * Q) * H)" by simp also have "(S * U + H * Q) * D + (T * U - D * Q) * H = (D * S + H * T) * U" by (simp add: field_simps) also have "Mp \ = Mp (Mp (D * S + H * T) * U)" by simp also have "Mp (D * S + H * T) = 1" using 1 by simp finally show eq: "A * D + B * H =m U" by simp have id: "degree_m (Mp R) = degree_m R" by simp have id': "degree D = degree_m D" using mon by simp show degB: "B = 0 \ degree B < degree D" using deg unfolding B id id' using degree_m_le[of R] by (cases "R = 0", auto) show Mp: "Mp A = A" "Mp B = B" unfolding A B by auto assume another: "A' * D + B' * H =m U" and degB': "B' = 0 \ degree B' < degree D" and norm: "Mp A' = A'" "Mp B' = B'" and cop: "coprime_m D H" and D: "Mp D = D" and prime: "prime m" from degB Mp D have degB: "B =m 0 \ degree_m B < degree_m D" by auto from degB' Mp D norm have degB': "B' =m 0 \ degree_m B' < degree_m D" by auto from mon D have D0: "\ (D =m 0)" by auto from prime interpret poly_mod_prime m by unfold_locales from another eq have "A' * D + B' * H =m A * D + B * H" by simp from uniqueness_poly_equality[OF cop degB' degB D0 this] show "A' = A \ B' = B" unfolding norm Mp by auto qed lemma coprime_bezout_coefficients: assumes cop: "coprime f g" and ext: "bezout_coefficients f g = (a, b)" shows "a * f + b * g = 1" using assms bezout_coefficients [of f g a b] by simp lemma (in poly_mod_prime_type) bezout_coefficients_mod_int: assumes f: "(F :: 'a mod_ring poly) = of_int_poly f" and g: "(G :: 'a mod_ring poly) = of_int_poly g" and cop: "coprime_m f g" and fact: "bezout_coefficients F G = (A,B)" and a: "a = to_int_poly A" and b: "b = to_int_poly B" shows "f * a + g * b =m 1" proof - have f[transfer_rule]: "MP_Rel f F" unfolding f MP_Rel_def by (simp add: Mp_f_representative) have g[transfer_rule]: "MP_Rel g G" unfolding g MP_Rel_def by (simp add: Mp_f_representative) have [transfer_rule]: "MP_Rel a A" unfolding a MP_Rel_def by (rule Mp_to_int_poly) have [transfer_rule]: "MP_Rel b B" unfolding b MP_Rel_def by (rule Mp_to_int_poly) from cop have "coprime F G" using coprime_MP_Rel[unfolded rel_fun_def] f g by auto from coprime_bezout_coefficients [OF this fact] have "A * F + B * G = 1" . from this [untransferred] show ?thesis by (simp add: ac_simps) qed definition bezout_coefficients_i :: "'i arith_ops_record \ 'i list \ 'i list \ 'i list \ 'i list" where "bezout_coefficients_i ff_ops f g = fst (euclid_ext_poly_i ff_ops f g)" definition euclid_ext_poly_mod_main :: "int \ 'a arith_ops_record \ int poly \ int poly \ int poly \ int poly" where "euclid_ext_poly_mod_main p ff_ops f g = (case bezout_coefficients_i ff_ops (of_int_poly_i ff_ops f) (of_int_poly_i ff_ops g) of (a,b) \ (to_int_poly_i ff_ops a, to_int_poly_i ff_ops b))" definition euclid_ext_poly_dynamic :: "int \ int poly \ int poly \ int poly \ int poly" where "euclid_ext_poly_dynamic p = ( if p \ 65535 then euclid_ext_poly_mod_main p (finite_field_ops32 (uint32_of_int p)) else if p \ 4294967295 then euclid_ext_poly_mod_main p (finite_field_ops64 (uint64_of_int p)) else euclid_ext_poly_mod_main p (finite_field_ops_integer (integer_of_int p)))" context prime_field_gen begin lemma bezout_coefficients_i[transfer_rule]: "(poly_rel ===> poly_rel ===> rel_prod poly_rel poly_rel) (bezout_coefficients_i ff_ops) bezout_coefficients" unfolding bezout_coefficients_i_def bezout_coefficients_def by transfer_prover lemma bezout_coefficients_i_sound: assumes f: "f' = of_int_poly_i ff_ops f" "Mp f = f" and g: "g' = of_int_poly_i ff_ops g" "Mp g = g" and cop: "coprime_m f g" and res: "bezout_coefficients_i ff_ops f' g' = (a',b')" and a: "a = to_int_poly_i ff_ops a'" and b: "b = to_int_poly_i ff_ops b'" shows "f * a + g * b =m 1" "Mp a = a" "Mp b = b" proof - from f have f': "f' = of_int_poly_i ff_ops (Mp f)" by simp define f'' where "f'' \ of_int_poly (Mp f) :: 'a mod_ring poly" have f'': "f'' = of_int_poly f" unfolding f''_def f by simp have rel_f[transfer_rule]: "poly_rel f' f''" by (rule poly_rel_of_int_poly[OF f'], simp add: f'' f) from g have g': "g' = of_int_poly_i ff_ops (Mp g)" by simp define g'' where "g'' \ of_int_poly (Mp g) :: 'a mod_ring poly" have g'': "g'' = of_int_poly g" unfolding g''_def g by simp have rel_g[transfer_rule]: "poly_rel g' g''" by (rule poly_rel_of_int_poly[OF g'], simp add: g'' g) obtain a'' b'' where eucl: "bezout_coefficients f'' g'' = (a'',b'')" by force from bezout_coefficients_i[unfolded rel_fun_def rel_prod_conv, rule_format, OF rel_f rel_g, unfolded res split eucl] have rel[transfer_rule]: "poly_rel a' a''" "poly_rel b' b''" by auto with to_int_poly_i have a: "a = to_int_poly a''" and b: "b = to_int_poly b''" unfolding a b by auto from bezout_coefficients_mod_int [OF f'' g'' cop eucl a b] show "f * a + g * b =m 1" . show "Mp a = a" "Mp b = b" unfolding a b by (auto simp: Mp_to_int_poly) qed lemma euclid_ext_poly_mod_main: assumes cop: "coprime_m f g" and f: "Mp f = f" and g: "Mp g = g" and res: "euclid_ext_poly_mod_main m ff_ops f g = (a,b)" shows "f * a + g * b =m 1" "Mp a = a" "Mp b = b" proof - obtain a' b' where res': "bezout_coefficients_i ff_ops (of_int_poly_i ff_ops f) (of_int_poly_i ff_ops g) = (a', b')" by force show "f * a + g * b =m 1" "Mp a = a" "Mp b = b" by (insert bezout_coefficients_i_sound[OF refl f refl g cop res'] res [unfolded euclid_ext_poly_mod_main_def res'], auto) qed end context poly_mod_prime begin lemmas euclid_ext_poly_mod_integer = prime_field_gen.euclid_ext_poly_mod_main [OF prime_field.prime_field_finite_field_ops_integer, unfolded prime_field_def mod_ring_locale_def poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas euclid_ext_poly_mod_uint32 = prime_field_gen.euclid_ext_poly_mod_main [OF prime_field.prime_field_finite_field_ops32, unfolded prime_field_def mod_ring_locale_def poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas euclid_ext_poly_mod_uint64 = prime_field_gen.euclid_ext_poly_mod_main[OF prime_field.prime_field_finite_field_ops64, unfolded prime_field_def mod_ring_locale_def poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemma euclid_ext_poly_dynamic: assumes cop: "coprime_m f g" and f: "Mp f = f" and g: "Mp g = g" and res: "euclid_ext_poly_dynamic p f g = (a,b)" shows "f * a + g * b =m 1" "Mp a = a" "Mp b = b" using euclid_ext_poly_mod_integer[OF cop f g, of p a b] euclid_ext_poly_mod_uint32[OF _ cop f g, of p a b] euclid_ext_poly_mod_uint64[OF _ cop f g, of p a b] res[unfolded euclid_ext_poly_dynamic_def] by (auto split: if_splits) end lemma range_sum_prod: assumes xy: "x \ {0.. {0.. {0..

{0 ..< q} \ 0 \ x \ x < q" by auto } note id = this from xy have 0: "0 \ x + q * y" by auto have "x + q * y \ q - 1 + q * y" using xy by simp also have "q * y \ q * (p - 1)" using xy by auto finally have "x + q * y \ q - 1 + q * (p - 1)" by auto also have "\ = p * q - 1" by (simp add: field_simps) finally show ?thesis using 0 by auto qed context fixes C :: "int poly" begin context fixes p :: int and S T D1 H1 :: "int poly" begin (* The linear lifting is implemented for ease of provability. Aim: show uniqueness of factorization *) fun linear_hensel_main where "linear_hensel_main (Suc 0) = (D1,H1)" | "linear_hensel_main (Suc n) = ( let (D,H) = linear_hensel_main n; q = p ^ n; U = poly_mod.Mp p (sdiv_poly (C - D * H) q); \ \\H2 + H3\\ (A,B) = poly_mod.dupe_monic_int p D1 H1 S T U in (D + smult q B, H + smult q A)) \ \\H4\\" | "linear_hensel_main 0 = (D1,H1)" lemma linear_hensel_main: assumes 1: "poly_mod.eq_m p (D1 * S + H1 * T) 1" and equiv: "poly_mod.eq_m p (D1 * H1) C" and monD1: "monic D1" and normDH1: "poly_mod.Mp p D1 = D1" "poly_mod.Mp p H1 = H1" and res: "linear_hensel_main n = (D,H)" and n: "n \ 0" and prime: "prime p" \ \\p > 1\ suffices if one does not need uniqueness\ and cop: "poly_mod.coprime_m p D1 H1" shows "poly_mod.eq_m (p^n) (D * H) C \ monic D \ poly_mod.eq_m p D D1 \ poly_mod.eq_m p H H1 \ poly_mod.Mp (p^n) D = D \ poly_mod.Mp (p^n) H = H \ (poly_mod.eq_m (p^n) (D' * H') C \ poly_mod.eq_m p D' D1 \ poly_mod.eq_m p H' H1 \ poly_mod.Mp (p^n) D' = D' \ poly_mod.Mp (p^n) H' = H' \ monic D' \ D' = D \ H' = H) " using res n proof (induct n arbitrary: D H D' H') case (Suc n D' H' D'' H'') show ?case proof (cases "n = 0") case True with Suc equiv monD1 normDH1 show ?thesis by auto next case False hence n: "n \ 0" by auto let ?q = "p^n" let ?pq = "p * p^n" from prime have p: "p > 1" using prime_gt_1_int by force from n p have q: "?q > 1" by auto from n p have pq: "?pq > 1" by (metis power_gt1_lemma) interpret p: poly_mod_2 p using p unfolding poly_mod_2_def . interpret q: poly_mod_2 ?q using q unfolding poly_mod_2_def . interpret pq: poly_mod_2 ?pq using pq unfolding poly_mod_2_def . obtain D H where rec: "linear_hensel_main n = (D,H)" by force obtain V where V: "sdiv_poly (C - D * H) ?q = V" by force obtain U where U: "p.Mp (sdiv_poly (C - D * H) ?q) = U" by auto obtain A B where dupe: "p.dupe_monic_int D1 H1 S T U = (A,B)" by force note IH = Suc(1)[OF rec n] from IH have CDH: "q.eq_m (D * H) C" and monD: "monic D" and p_eq: "p.eq_m D D1" "p.eq_m H H1" and norm: "q.Mp D = D" "q.Mp H = H" by auto from n obtain k where n: "n = Suc k" by (cases n, auto) have qq: "?q * ?q = ?pq * p^k" unfolding n by simp from Suc(2)[unfolded n linear_hensel_main.simps, folded n, unfolded rec split Let_def U dupe] have D': "D' = D + smult ?q B" and H': "H' = H + smult ?q A" by auto note dupe = p.dupe_monic_int[OF 1 monD1 dupe] from CDH have "q.Mp C - q.Mp (D * H) = 0" by simp hence "q.Mp (q.Mp C - q.Mp (D * H)) = 0" by simp hence "q.Mp (C - D*H) = 0" by simp from q.Mp_0_smult_sdiv_poly[OF this] have CDHq: "smult ?q (sdiv_poly (C - D * H) ?q) = C - D * H" . have ADBHU: "p.eq_m (A * D + B * H) U" using p_eq dupe(1) by (metis (mono_tags, lifting) p.mult_Mp(2) poly_mod.plus_Mp) have "pq.Mp (D' * H') = pq.Mp ((D + smult ?q B) * (H + smult ?q A))" unfolding D' H' by simp also have "(D + smult ?q B) * (H + smult ?q A) = (D * H + smult ?q (A * D + B * H)) + smult (?q * ?q) (A * B)" by (simp add: field_simps smult_distribs) also have "pq.Mp \ = pq.Mp (D * H + pq.Mp (smult ?q (A * D + B * H)) + pq.Mp (smult (?q * ?q) (A * B)))" using pq.plus_Mp by metis also have "pq.Mp (smult (?q * ?q) (A * B)) = 0" unfolding qq by (metis pq.Mp_smult_m_0 smult_smult) finally have DH': "pq.Mp (D' * H') = pq.Mp (D * H + pq.Mp (smult ?q (A * D + B * H)))" by simp also have "pq.Mp (smult ?q (A * D + B * H)) = pq.Mp (smult ?q U)" using p.Mp_lift_modulus[OF ADBHU, of ?q] by simp also have "\ = pq.Mp (C - D * H)" unfolding arg_cong[OF CDHq, of pq.Mp, symmetric] U[symmetric] V by (rule p.Mp_lift_modulus[of _ _ ?q], auto) also have "pq.Mp (D * H + pq.Mp (C - D * H)) = pq.Mp C" by simp finally have CDH: "pq.eq_m C (D' * H')" by simp have deg: "degree D1 = degree D" using p_eq(1) monD1 monD by (metis p.monic_degree_m) have mon: "monic D'" unfolding D' using dupe(2) monD unfolding deg by (rule monic_smult_add_small) have normD': "pq.Mp D' = D'" unfolding D' pq.Mp_ident_iff poly_mod.Mp_coeff plus_poly.rep_eq coeff_smult proof fix i from norm(1) dupe(4) have "coeff D i \ {0.. {0.. {0..< ?pq}" by (rule range_sum_prod) qed have normH': "pq.Mp H' = H'" unfolding H' pq.Mp_ident_iff poly_mod.Mp_coeff plus_poly.rep_eq coeff_smult proof fix i from norm(2) dupe(3) have "coeff H i \ {0.. {0.. {0..< ?pq}" by (rule range_sum_prod) qed have eq: "p.eq_m D D'" "p.eq_m H H'" unfolding D' H' n poly_eq_iff p.Mp_coeff p.M_def by (auto simp: field_simps) with p_eq have eq: "p.eq_m D' D1" "p.eq_m H' H1" by auto { assume CDH'': "pq.eq_m C (D'' * H'')" and DH1'': "p.eq_m D1 D''" "p.eq_m H1 H''" and norm'': "pq.Mp D'' = D''" "pq.Mp H'' = H''" and monD'': "monic D''" from q.Dp_Mp_eq[of D''] obtain d B' where D'': "D'' = q.Mp d + smult ?q B'" by auto from q.Dp_Mp_eq[of H''] obtain h A' where H'': "H'' = q.Mp h + smult ?q A'" by auto { fix A B assume *: "pq.Mp (q.Mp A + smult ?q B) = q.Mp A + smult ?q B" have "p.Mp B = B" unfolding p.Mp_ident_iff proof fix i from arg_cong[OF *, of "\ f. coeff f i", unfolded pq.Mp_coeff pq.M_def] have "coeff (q.Mp A + smult ?q B) i \ {0 ..< ?pq}" using "*" pq.Mp_ident_iff by blast hence sum: "coeff (q.Mp A) i + ?q * coeff B i \ {0 ..< ?pq}" by auto have "q.Mp (q.Mp A) = q.Mp A" by auto from this[unfolded q.Mp_ident_iff] have A: "coeff (q.Mp A) i \ {0 ..< p^n}" by auto { assume "coeff B i < 0" hence "coeff B i \ -1" by auto from mult_left_mono[OF this, of ?q] q.m1 have "?q * coeff B i \ -?q" by simp with A sum have False by auto } hence "coeff B i \ 0" by force moreover { assume "coeff B i \ p" from mult_left_mono[OF this, of ?q] q.m1 have "?q * coeff B i \ ?pq" by simp with A sum have False by auto } hence "coeff B i < p" by force ultimately show "coeff B i \ {0 ..< p}" by auto qed } note norm_convert = this from norm_convert[OF norm''(1)[unfolded D'']] have normB': "p.Mp B' = B'" . from norm_convert[OF norm''(2)[unfolded H'']] have normA': "p.Mp A' = A'" . let ?d = "q.Mp d" let ?h = "q.Mp h" { assume lt: "degree ?d < degree B'" hence eq: "degree D'' = degree B'" unfolding D'' using q.m1 p.m1 by (subst degree_add_eq_right, auto) from lt have [simp]: "coeff ?d (degree B') = 0" by (rule coeff_eq_0) from monD''[unfolded eq, unfolded D'', simplified] False q.m1 lt have False by (metis mod_mult_self1_is_0 poly_mod.M_def q.M_1 zero_neq_one) } hence deg_dB': "degree ?d \ degree B'" by presburger { assume eq: "degree ?d = degree B'" and B': "B' \ 0" let ?B = "coeff B' (degree B')" from normB'[unfolded p.Mp_ident_iff, rule_format, of "degree B'"] B' have "?B \ {0.. 0" "?B < p" by auto have degD'': "degree D'' \ degree ?d" unfolding D'' using eq by (simp add: degree_add_le) have "?q * ?B \ 1 * 1" by (rule mult_mono, insert q.m1 bnds, auto) moreover have "coeff D'' (degree ?d) = 1 + ?q * ?B" using monD'' unfolding D'' using eq by (metis D'' coeff_smult monD'' plus_poly.rep_eq poly_mod.Dp_Mp_eq poly_mod.degree_m_eq_monic poly_mod.plus_Mp(1) q.Mp_smult_m_0 q.m1 q.monic_Mp q.plus_Mp(2)) ultimately have gt: "coeff D'' (degree ?d) > 1" by auto hence "coeff D'' (degree ?d) \ 0" by auto hence "degree D'' \ degree ?d" by (rule le_degree) with degree_add_le_max[of ?d "smult ?q B'", folded D''] eq have deg: "degree D'' = degree ?d" using degD'' by linarith from gt[folded this] have "\ monic D''" by auto with monD'' have False by auto } with deg_dB' have deg_dB2: "B' = 0 \ degree B' < degree ?d" by fastforce have d: "q.Mp D'' = ?d" unfolding D'' by (metis add.right_neutral poly_mod.Mp_smult_m_0 poly_mod.plus_Mp) have h: "q.Mp H'' = ?h" unfolding H'' by (metis add.right_neutral poly_mod.Mp_smult_m_0 poly_mod.plus_Mp) from CDH'' have "pq.Mp C = pq.Mp (D'' * H'')" by simp from arg_cong[OF this, of q.Mp] have "q.Mp C = q.Mp (D'' * H'')" using p.m1 q.Mp_product_modulus by auto also have "\ = q.Mp (q.Mp D'' * q.Mp H'')" by simp also have "\ = q.Mp (?d * ?h)" unfolding d h by simp finally have eqC: "q.eq_m (?d * ?h) C" by auto have d1: "p.eq_m ?d D1" unfolding d[symmetric] using DH1'' using assms(4) n p.Mp_product_modulus p.m1 by auto have h1: "p.eq_m ?h H1" unfolding h[symmetric] using DH1'' using assms(5) n p.Mp_product_modulus p.m1 by auto have mond: "monic (q.Mp d)" using monD'' deg_dB2 unfolding D'' using d q.monic_Mp[OF monD''] by simp from eqC d1 h1 mond IH[of "q.Mp d" "q.Mp h"] have IH: "?d = D" "?h = H" by auto from deg_dB2[unfolded IH] have degB': "B' = 0 \ degree B' < degree D" by auto from IH have D'': "D'' = D + smult ?q B'" and H'': "H'' = H + smult ?q A'" unfolding D'' H'' by auto have "pq.Mp (D'' * H'') = pq.Mp (D' * H')" using CDH'' CDH by simp also have "pq.Mp (D'' * H'') = pq.Mp ((D + smult ?q B') * (H + smult ?q A'))" unfolding D'' H'' by simp also have "(D + smult ?q B') * (H + smult ?q A') = (D * H + smult ?q (A' * D + B' * H)) + smult (?q * ?q) (A' * B')" by (simp add: field_simps smult_distribs) also have "pq.Mp \ = pq.Mp (D * H + pq.Mp (smult ?q (A' * D + B' * H)) + pq.Mp (smult (?q * ?q) (A' * B')))" using pq.plus_Mp by metis also have "pq.Mp (smult (?q * ?q) (A' * B')) = 0" unfolding qq by (metis pq.Mp_smult_m_0 smult_smult) finally have "pq.Mp (D * H + pq.Mp (smult ?q (A' * D + B' * H))) = pq.Mp (D * H + pq.Mp (smult ?q (A * D + B * H)))" unfolding DH' by simp hence "pq.Mp (smult ?q (A' * D + B' * H)) = pq.Mp (smult ?q (A * D + B * H))" by (metis (no_types, lifting) add_diff_cancel_left' poly_mod.minus_Mp(1) poly_mod.plus_Mp(2)) hence "p.Mp (A' * D + B' * H) = p.Mp (A * D + B * H)" unfolding poly_eq_iff p.Mp_coeff pq.Mp_coeff coeff_smult by (insert p, auto simp: p.M_def pq.M_def) hence "p.Mp (A' * D1 + B' * H1) = p.Mp (A * D1 + B * H1)" using p_eq by (metis p.mult_Mp(2) poly_mod.plus_Mp) hence eq: "p.eq_m (A' * D1 + B' * H1) U" using dupe(1) by auto have "degree D = degree D1" using monD monD1 arg_cong[OF p_eq(1), of degree] p.degree_m_eq_monic[OF _ p.m1] by auto hence "B' = 0 \ degree B' < degree D1" using degB' by simp from dupe(5)[OF cop eq this normDH1(1) normA' normB' prime] have "A' = A" "B' = B" by auto hence "D'' = D'" "H'' = H'" unfolding D'' H'' D' H' by auto } thus ?thesis using normD' normH' CDH mon eq by simp qed qed simp end end definition linear_hensel_binary :: "int \ nat \ int poly \ int poly \ int poly \ int poly \ int poly" where "linear_hensel_binary p n C D H = (let (S,T) = euclid_ext_poly_dynamic p D H in linear_hensel_main C p S T D H n)" lemma (in poly_mod_prime) unique_hensel_binary: assumes prime: "prime p" and cop: "coprime_m D H" and eq: "eq_m (D * H) C" and normalized_input: "Mp D = D" "Mp H = H" and monic_input: "monic D" and n: "n \ 0" shows "\! (D',H'). \ \\D'\, \H'\ are computed via \linear_hensel_binary\\ poly_mod.eq_m (p^n) (D' * H') C \ \the main result: equivalence mod \p^n\\ \ monic D' \ \monic output\ \ eq_m D D' \ eq_m H H' \ \apply \`mod p`\ on \D'\ and \H'\ yields \D\ and \H\ again\ \ poly_mod.Mp (p^n) D' = D' \ poly_mod.Mp (p^n) H' = H' \ \output is normalized\" proof - obtain D' H' where hensel_result: "linear_hensel_binary p n C D H = (D',H')" by force from m1 have p: "p > 1" . obtain S T where ext: "euclid_ext_poly_dynamic p D H = (S,T)" by force obtain D1 H1 where main: "linear_hensel_main C p S T D H n = (D1,H1)" by force from hensel_result[unfolded linear_hensel_binary_def ext split Let_def main] have id: "D1 = D'" "H1 = H'" by auto note eucl = euclid_ext_poly_dynamic [OF cop normalized_input ext] from linear_hensel_main [OF eucl(1) eq monic_input normalized_input main [unfolded id] n prime cop] show ?thesis by (intro ex1I, auto) qed (* The quadratic lifting is implemented more efficienty. Aim: compute factorization *) context fixes C :: "int poly" begin lemma hensel_step_main: assumes one_q: "poly_mod.eq_m q (D * S + H * T) 1" and one_p: "poly_mod.eq_m p (D1 * S1 + H1 * T1) 1" and CDHq: "poly_mod.eq_m q C (D * H)" and D1D: "poly_mod.eq_m p D1 D" and H1H: "poly_mod.eq_m p H1 H" and S1S: "poly_mod.eq_m p S1 S" and T1T: "poly_mod.eq_m p T1 T" and mon: "monic D" and mon1: "monic D1" and q: "q > 1" and p: "p > 1" and D1: "poly_mod.Mp p D1 = D1" and H1: "poly_mod.Mp p H1 = H1" and S1: "poly_mod.Mp p S1 = S1" and T1: "poly_mod.Mp p T1 = T1" and D: "poly_mod.Mp q D = D" and H: "poly_mod.Mp q H = H" and S: "poly_mod.Mp q S = S" and T: "poly_mod.Mp q T = T" and U1: "U1 = poly_mod.Mp p (sdiv_poly (C - D * H) q)" and dupe1: "dupe_monic_dynamic p D1 H1 S1 T1 U1 = (A,B)" and D': "D' = D + smult q B" and H': "H' = H + smult q A" and U2: "U2 = poly_mod.Mp q (sdiv_poly (S*D' + T*H' - 1) p)" and dupe2: "dupe_monic_dynamic q D H S T U2 = (A',B')" and rq: "r = p * q" and pq: "p dvd q" and S': "S' = poly_mod.Mp r (S - smult p A')" and T': "T' = poly_mod.Mp r (T - smult p B')" shows "poly_mod.eq_m r C (D' * H')" "poly_mod.Mp r D' = D'" "poly_mod.Mp r H' = H'" "poly_mod.Mp r S' = S'" "poly_mod.Mp r T' = T'" "poly_mod.eq_m r (D' * S' + H' * T') 1" "monic D'" unfolding rq proof - from pq obtain k where qp: "q = p * k" unfolding dvd_def by auto from arg_cong[OF qp, of sgn] q p have k0: "k > 0" unfolding sgn_mult by (auto simp: sgn_1_pos) from qp have qq: "q * q = p * q * k" by auto let ?r = "p * q" interpret poly_mod_2 p by (standard, insert p, auto) interpret q: poly_mod_2 q by (standard, insert q, auto) from p q have r: "?r > 1" by (simp add: less_1_mult) interpret r: poly_mod_2 ?r using r unfolding poly_mod_2_def . have Mp_conv: "Mp (q.Mp x) = Mp x" for x unfolding qp by (rule Mp_product_modulus[OF refl k0]) from arg_cong[OF CDHq, of Mp, unfolded Mp_conv] have "Mp C = Mp (Mp D * Mp H)" by simp also have "Mp D = Mp D1" using D1D by simp also have "Mp H = Mp H1" using H1H by simp finally have CDHp: "eq_m C (D1 * H1)" by simp have "Mp U1 = U1" unfolding U1 by simp note dupe1 = dupe_monic_dynamic[OF dupe1 one_p mon1 D1 H1 S1 T1 this] have "q.Mp U2 = U2" unfolding U2 by simp note dupe2 = q.dupe_monic_dynamic[OF dupe2 one_q mon D H S T this] from CDHq have "q.Mp C - q.Mp (D * H) = 0" by simp hence "q.Mp (q.Mp C - q.Mp (D * H)) = 0" by simp hence "q.Mp (C - D*H) = 0" by simp from q.Mp_0_smult_sdiv_poly[OF this] have CDHq: "smult q (sdiv_poly (C - D * H) q) = C - D * H" . { fix A B have "Mp (A * D1 + B * H1) = Mp (Mp (A * D1) + Mp (B * H1))" by simp also have "Mp (A * D1) = Mp (A * Mp D1)" by simp also have "\ = Mp (A * D)" unfolding D1D by simp also have "Mp (B * H1) = Mp (B * Mp H1)" by simp also have "\ = Mp (B * H)" unfolding H1H by simp finally have "Mp (A * D1 + B * H1) = Mp (A * D + B * H)" by simp } note D1H1 = this have "r.Mp (D' * H') = r.Mp ((D + smult q B) * (H + smult q A))" unfolding D' H' by simp also have "(D + smult q B) * (H + smult q A) = (D * H + smult q (A * D + B * H)) + smult (q * q) (A * B)" by (simp add: field_simps smult_distribs) also have "r.Mp \ = r.Mp (D * H + r.Mp (smult q (A * D + B * H)) + r.Mp (smult (q * q) (A * B)))" using r.plus_Mp by metis also have "r.Mp (smult (q * q) (A * B)) = 0" unfolding qq by (metis r.Mp_smult_m_0 smult_smult) also have "r.Mp (smult q (A * D + B * H)) = r.Mp (smult q U1)" proof (rule Mp_lift_modulus[of _ _ q]) show "Mp (A * D + B * H) = Mp U1" using dupe1(1) unfolding D1H1 by simp qed also have "\ = r.Mp (C - D * H)" unfolding arg_cong[OF CDHq, of r.Mp, symmetric] using Mp_lift_modulus[of U1 "sdiv_poly (C - D * H) q" q] unfolding U1 by simp also have "r.Mp (D * H + r.Mp (C - D * H) + 0) = r.Mp C" by simp finally show CDH: "r.eq_m C (D' * H')" by simp have "degree D1 = degree (Mp D1)" using mon1 by simp also have "\ = degree D" unfolding D1D using mon by simp finally have deg_eq: "degree D1 = degree D" by simp show mon: "monic D'" unfolding D' using dupe1(2) mon unfolding deg_eq by (rule monic_smult_add_small) have "Mp (S * D' + T * H' - 1) = Mp (Mp (D * S + H * T) + (smult q (S * B + T * A) - 1))" unfolding D' H' plus_Mp by (simp add: field_simps smult_distribs) also have "Mp (D * S + H * T) = Mp (Mp (D1 * Mp S) + Mp (H1 * Mp T))" using D1H1[of S T] by (simp add: ac_simps) also have "\ = 1" using one_p unfolding S1S[symmetric] T1T[symmetric] by simp also have "Mp (1 + (smult q (S * B + T * A) - 1)) = Mp (smult q (S * B + T * A))" by simp also have "\ = 0" unfolding qp by (metis Mp_smult_m_0 smult_smult) finally have "Mp (S * D' + T * H' - 1) = 0" . from Mp_0_smult_sdiv_poly[OF this] have SDTH: "smult p (sdiv_poly (S * D' + T * H' - 1) p) = S * D' + T * H' - 1" . have swap: "q * p = p * q" by simp have "r.Mp (D' * S' + H' * T') = r.Mp ((D + smult q B) * (S - smult p A') + (H + smult q A) * (T - smult p B'))" unfolding D' S' H' T' rq using r.plus_Mp r.mult_Mp by metis also have "\ = r.Mp ((D * S + H * T + smult q (B * S + A * T)) - smult p (A' * D + B' * H) - smult ?r (A * B' + B * A'))" by (simp add: field_simps smult_distribs) also have "\ = r.Mp ((D * S + H * T + smult q (B * S + A * T)) - r.Mp (smult p (A' * D + B' * H)) - r.Mp (smult ?r (A * B' + B * A')))" using r.plus_Mp r.minus_Mp by metis also have "r.Mp (smult ?r (A * B' + B * A')) = 0" by simp also have "r.Mp (smult p (A' * D + B' * H)) = r.Mp (smult p U2)" using q.Mp_lift_modulus[OF dupe2(1), of p] unfolding swap . also have "\ = r.Mp (S * D' + T * H' - 1)" unfolding arg_cong[OF SDTH, of r.Mp, symmetric] using q.Mp_lift_modulus[of U2 "sdiv_poly (S * D' + T * H' - 1) p" p] unfolding U2 swap by simp also have "S * D' + T * H' - 1 = S * D + T * H + smult q (B * S + A * T) - 1" unfolding D' H' by (simp add: field_simps smult_distribs) also have "r.Mp (D * S + H * T + smult q (B * S + A * T) - r.Mp (S * D + T * H + smult q (B * S + A * T) - 1) - 0) = 1" by simp finally show 1: "r.eq_m (D' * S' + H' * T') 1" by simp show D': "r.Mp D' = D'" unfolding D' r.Mp_ident_iff poly_mod.Mp_coeff plus_poly.rep_eq coeff_smult proof fix n from D dupe1(4) have "coeff D n \ {0.. {0.. {0.. {0.. {0.. {0.. \\Z2 and Z3\\ (A,B) = dupe_monic_dynamic p D1 H1 S1 T1 U; D' = D + smult q B; \ \\Z4\\ H' = H + smult q A; U' = poly_mod.Mp q (sdiv_poly (S*D' + T*H' - 1) p); \ \\Z5 + Z6\\ (A',B') = dupe_monic_dynamic q D H S T U'; q' = p * q; S' = poly_mod.Mp q' (S - smult p A'); \ \\Z7\\ T' = poly_mod.Mp q' (T - smult p B') in (S',T',D',H'))" definition "quadratic_hensel_step q S T D H = hensel_step q q S T D H S T D H" lemma quadratic_hensel_step_code[code]: "quadratic_hensel_step q S T D H = (let dupe = dupe_monic_dynamic q D H S T; \ \this will share the conversions of \D H S T\\ U = poly_mod.Mp q (sdiv_poly (C - D * H) q); (A, B) = dupe U; D' = D + Polynomial.smult q B; H' = H + Polynomial.smult q A; U' = poly_mod.Mp q (sdiv_poly (S * D' + T * H' - 1) q); (A', B') = dupe U'; q' = q * q; S' = poly_mod.Mp q' (S - Polynomial.smult q A'); T' = poly_mod.Mp q' (T - Polynomial.smult q B') in (S', T', D', H'))" unfolding quadratic_hensel_step_def[unfolded hensel_step_def] Let_def .. definition simple_quadratic_hensel_step where \ \do not compute new values \S'\ and \T'\\ "simple_quadratic_hensel_step q S T D H = ( let U = poly_mod.Mp q (sdiv_poly (C - D * H) q); \ \\Z2 + Z3\\ (A,B) = dupe_monic_dynamic q D H S T U; D' = D + smult q B; \ \\Z4\\ H' = H + smult q A in (D',H'))" lemma hensel_step: assumes step: "hensel_step p q S1 T1 D1 H1 S T D H = (S', T', D', H')" and one_p: "poly_mod.eq_m p (D1 * S1 + H1 * T1) 1" and mon1: "monic D1" and p: "p > 1" and CDHq: "poly_mod.eq_m q C (D * H)" and one_q: "poly_mod.eq_m q (D * S + H * T) 1" and D1D: "poly_mod.eq_m p D1 D" and H1H: "poly_mod.eq_m p H1 H" and S1S: "poly_mod.eq_m p S1 S" and T1T: "poly_mod.eq_m p T1 T" and mon: "monic D" and q: "q > 1" and D1: "poly_mod.Mp p D1 = D1" and H1: "poly_mod.Mp p H1 = H1" and S1: "poly_mod.Mp p S1 = S1" and T1: "poly_mod.Mp p T1 = T1" and D: "poly_mod.Mp q D = D" and H: "poly_mod.Mp q H = H" and S: "poly_mod.Mp q S = S" and T: "poly_mod.Mp q T = T" and rq: "r = p * q" and pq: "p dvd q" shows "poly_mod.eq_m r C (D' * H')" "poly_mod.eq_m r (D' * S' + H' * T') 1" "poly_mod.Mp r D' = D'" "poly_mod.Mp r H' = H'" "poly_mod.Mp r S' = S'" "poly_mod.Mp r T' = T'" "poly_mod.Mp p D1 = poly_mod.Mp p D'" "poly_mod.Mp p H1 = poly_mod.Mp p H'" "poly_mod.Mp p S1 = poly_mod.Mp p S'" "poly_mod.Mp p T1 = poly_mod.Mp p T'" "monic D'" proof - define U where U: "U = poly_mod.Mp p (sdiv_poly (C - D * H) q)" note step = step[unfolded hensel_step_def Let_def, folded U] obtain A B where dupe1: "dupe_monic_dynamic p D1 H1 S1 T1 U = (A,B)" by force note step = step[unfolded dupe1 split] from step have D': "D' = D + smult q B" and H': "H' = H + smult q A" by (auto split: prod.splits) define U' where U': "U' = poly_mod.Mp q (sdiv_poly (S * D' + T * H' - 1) p)" obtain A' B' where dupe2: "dupe_monic_dynamic q D H S T U' = (A',B')" by force from step[folded D' H', folded U', unfolded dupe2 split, folded rq] have S': "S' = poly_mod.Mp r (S - Polynomial.smult p A')" and T': "T' = poly_mod.Mp r (T - Polynomial.smult p B')" by auto from hensel_step_main[OF one_q one_p CDHq D1D H1H S1S T1T mon mon1 q p D1 H1 S1 T1 D H S T U dupe1 D' H' U' dupe2 rq pq S' T'] show "poly_mod.eq_m r (D' * S' + H' * T') 1" "poly_mod.eq_m r C (D' * H')" "poly_mod.Mp r D' = D'" "poly_mod.Mp r H' = H'" "poly_mod.Mp r S' = S'" "poly_mod.Mp r T' = T'" "monic D'" by auto from pq obtain s where q: "q = p * s" by (metis dvdE) show "poly_mod.Mp p D1 = poly_mod.Mp p D'" "poly_mod.Mp p H1 = poly_mod.Mp p H'" unfolding q D' D1D H' H1H by (metis add.right_neutral poly_mod.Mp_smult_m_0 poly_mod.plus_Mp(2) smult_smult)+ from \q > 1\ have q0: "q > 0" by auto show "poly_mod.Mp p S1 = poly_mod.Mp p S'" "poly_mod.Mp p T1 = poly_mod.Mp p T'" unfolding S' S1S T' T1T poly_mod_2.Mp_product_modulus[OF poly_mod_2.intro[OF \p > 1\] rq q0] by (metis group_add_class.diff_0_right poly_mod.Mp_smult_m_0 poly_mod.minus_Mp(2))+ qed lemma quadratic_hensel_step: assumes step: "quadratic_hensel_step q S T D H = (S', T', D', H')" and CDH: "poly_mod.eq_m q C (D * H)" and one: "poly_mod.eq_m q (D * S + H * T) 1" and D: "poly_mod.Mp q D = D" and H: "poly_mod.Mp q H = H" and S: "poly_mod.Mp q S = S" and T: "poly_mod.Mp q T = T" and mon: "monic D" and q: "q > 1" and rq: "r = q * q" shows "poly_mod.eq_m r C (D' * H')" "poly_mod.eq_m r (D' * S' + H' * T') 1" "poly_mod.Mp r D' = D'" "poly_mod.Mp r H' = H'" "poly_mod.Mp r S' = S'" "poly_mod.Mp r T' = T'" "poly_mod.Mp q D = poly_mod.Mp q D'" "poly_mod.Mp q H = poly_mod.Mp q H'" "poly_mod.Mp q S = poly_mod.Mp q S'" "poly_mod.Mp q T = poly_mod.Mp q T'" "monic D'" proof (atomize(full), goal_cases) case 1 from hensel_step[OF step[unfolded quadratic_hensel_step_def] one mon q CDH one refl refl refl refl mon q D H S T D H S T rq] show ?case by auto qed context fixes p :: int and S1 T1 D1 H1 :: "int poly" begin private lemma decrease[termination_simp]: "\ j \ 1 \ odd j \ Suc (j div 2) < j" by presburger fun quadratic_hensel_loop where "quadratic_hensel_loop (j :: nat) = ( if j \ 1 then (p, S1, T1, D1, H1) else if even j then (case quadratic_hensel_loop (j div 2) of (q, S, T, D, H) \ let qq = q * q in (case quadratic_hensel_step q S T D H of \ \quadratic step\ (S', T', D', H') \ (qq, S', T', D', H'))) else \ \odd \j\\ (case quadratic_hensel_loop (j div 2 + 1) of (q, S, T, D, H) \ (case quadratic_hensel_step q S T D H of \ \quadratic step\ (S', T', D', H') \ let qq = q * q; pj = qq div p; down = poly_mod.Mp pj in (pj, down S', down T', down D', down H'))))" definition "quadratic_hensel_main j = (case quadratic_hensel_loop j of (qq, S, T, D, H) \ (D, H))" declare quadratic_hensel_loop.simps[simp del] \ \unroll the definition of \hensel_loop\ so that in outermost iteration we can use \simple_hensel_step\\ lemma quadratic_hensel_main_code[code]: "quadratic_hensel_main j = ( if j \ 1 then (D1, H1) else if even j then (case quadratic_hensel_loop (j div 2) of (q, S, T, D, H) \ simple_quadratic_hensel_step q S T D H) else (case quadratic_hensel_loop (j div 2 + 1) of (q, S, T, D, H) \ (case simple_quadratic_hensel_step q S T D H of (D', H') \ let down = poly_mod.Mp (q * q div p) in (down D', down H'))))" unfolding quadratic_hensel_loop.simps[of j] quadratic_hensel_main_def Let_def by (simp split: if_splits prod.splits option.splits sum.splits add: quadratic_hensel_step_code simple_quadratic_hensel_step_def Let_def) context fixes j :: nat assumes 1: "poly_mod.eq_m p (D1 * S1 + H1 * T1) 1" and CDH1: "poly_mod.eq_m p C (D1 * H1)" and mon1: "monic D1" and p: "p > 1" and D1: "poly_mod.Mp p D1 = D1" and H1: "poly_mod.Mp p H1 = H1" and S1: "poly_mod.Mp p S1 = S1" and T1: "poly_mod.Mp p T1 = T1" and j: "j \ 1" begin lemma quadratic_hensel_loop: assumes "quadratic_hensel_loop j = (q, S, T, D, H)" shows "(poly_mod.eq_m q C (D * H) \ monic D \ poly_mod.eq_m p D1 D \ poly_mod.eq_m p H1 H \ poly_mod.eq_m q (D * S + H * T) 1 \ poly_mod.Mp q D = D \ poly_mod.Mp q H = H \ poly_mod.Mp q S = S \ poly_mod.Mp q T = T \ q = p^j)" using j assms proof (induct j arbitrary: q S T D H rule: less_induct) case (less j q' S' T' D' H') note res = less(3) interpret poly_mod_2 p using p by (rule poly_mod_2.intro) let ?hens = "quadratic_hensel_loop" note simp[simp] = quadratic_hensel_loop.simps[of j] show ?case proof (cases "j = 1") case True show ?thesis using res simp unfolding True using CDH1 1 mon1 D1 H1 S1 T1 by auto next case False with less(2) have False: "(j \ 1) = False" by auto have mod_2: "k \ 1 \ poly_mod_2 (p^k)" for k by (intro poly_mod_2.intro, insert p, auto) { fix k D assume *: "k \ 1" "k \ j" "poly_mod.Mp (p ^ k) D = D" from *(2) have "{0..

{0..

?j2" by auto obtain q S T D H where rec: "?hens ?j2 = (q, S, T, D, H)" by (cases "?hens ?j2", auto) note IH = less(1)[OF lt rec] from IH have *: "poly_mod.eq_m q C (D * H)" "poly_mod.eq_m q (D * S + H * T) 1" "monic D" "eq_m D1 D" "eq_m H1 H" "poly_mod.Mp q D = D" "poly_mod.Mp q H = H" "poly_mod.Mp q S = S" "poly_mod.Mp q T = T" "q = p ^ ?j2" by auto hence norm: "poly_mod.Mp (p ^ j) D = D" "poly_mod.Mp (p ^ j) H = H" "poly_mod.Mp (p ^ j) S = S" "poly_mod.Mp (p ^ j) T = T" using lift_norm[OF lt(2)] by auto from lt p have q: "q > 1" unfolding * by simp let ?step = "quadratic_hensel_step q S T D H" obtain S2 T2 D2 H2 where step_res: "?step = (S2, T2, D2, H2)" by (cases ?step, auto) note step = quadratic_hensel_step[OF step_res *(1,2,6-9,3) q refl] let ?qq = "q * q" { fix D D2 assume "poly_mod.Mp q D = poly_mod.Mp q D2" from arg_cong[OF this, of Mp] Mp_Mp_pow_is_Mp[of ?j2, OF _ p, folded *(10)] lt have "Mp D = Mp D2" by simp } note shrink = this have **: "poly_mod.eq_m ?qq C (D2 * H2)" "poly_mod.eq_m ?qq (D2 * S2 + H2 * T2) 1" "monic D2" "eq_m D1 D2" "eq_m H1 H2" "poly_mod.Mp ?qq D2 = D2" "poly_mod.Mp ?qq H2 = H2" "poly_mod.Mp ?qq S2 = S2" "poly_mod.Mp ?qq T2 = T2" using step shrink[of H H2] shrink[of D D2] *(4-7) by auto note simp = simp False if_False rec split Let_def step_res option.simps from True have j: "p ^ j = p ^ (2 * ?j2)" by auto with *(10) have qq: "q * q = p ^ j" by (simp add: power_mult_distrib semiring_normalization_rules(30-)) from res[unfolded simp] True have id': "q' = ?qq" "S' = S2" "T' = T2" "D' = D2" "H' = H2" by auto show ?thesis unfolding id' using ** by (auto simp: qq) next case odd: False hence False': "(even j) = False" by auto let ?j2 = "j div 2 + 1" from False odd have lt: "?j2 < j" "1 \ ?j2" by presburger+ obtain q S T D H where rec: "?hens ?j2 = (q, S, T, D, H)" by (cases "?hens ?j2", auto) note IH = less(1)[OF lt rec] note simp = simp False if_False rec sum.simps split Let_def False' option.simps from IH have *: "poly_mod.eq_m q C (D * H)" "poly_mod.eq_m q (D * S + H * T) 1" "monic D" "eq_m D1 D" "eq_m H1 H" "poly_mod.Mp q D = D" "poly_mod.Mp q H = H" "poly_mod.Mp q S = S" "poly_mod.Mp q T = T" "q = p ^ ?j2" by auto hence norm: "poly_mod.Mp (p ^ j) D = D" "poly_mod.Mp (p ^ j) H = H" using lift_norm[OF lt(2)] lt by auto from lt p have q: "q > 1" unfolding * using mod_2 poly_mod_2.m1 by blast let ?step = "quadratic_hensel_step q S T D H" obtain S2 T2 D2 H2 where step_res: "?step = (S2, T2, D2, H2)" by (cases ?step, auto) have dvd: "q dvd q" by auto note step = quadratic_hensel_step[OF step_res *(1,2,6-9,3) q refl] let ?qq = "q * q" { fix D D2 assume "poly_mod.Mp q D = poly_mod.Mp q D2" from arg_cong[OF this, of Mp] Mp_Mp_pow_is_Mp[of ?j2, OF _ p, folded *(10)] lt have "Mp D = Mp D2" by simp } note shrink = this have **: "poly_mod.eq_m ?qq C (D2 * H2)" "poly_mod.eq_m ?qq (D2 * S2 + H2 * T2) 1" "monic D2" "eq_m D1 D2" "eq_m H1 H2" "poly_mod.Mp ?qq D2 = D2" "poly_mod.Mp ?qq H2 = H2" "poly_mod.Mp ?qq S2 = S2" "poly_mod.Mp ?qq T2 = T2" using step shrink[of H H2] shrink[of D D2] *(4-7) by auto note simp = simp False if_False rec split Let_def step_res option.simps from odd have j: "Suc j = 2 * ?j2" by auto from arg_cong[OF this, of "\ j. p ^ j div p"] have pj: "p ^ j = q * q div p" and qq: "q * q = p ^ j * p" unfolding *(10) using p by (simp add: power_mult_distrib semiring_normalization_rules(30-))+ let ?pj = "p ^ j" from res[unfolded simp] pj have id: "q' = p^j" "S' = poly_mod.Mp ?pj S2" "T' = poly_mod.Mp ?pj T2" "D' = poly_mod.Mp ?pj D2" "H' = poly_mod.Mp ?pj H2" by auto interpret pj: poly_mod_2 ?pj by (rule mod_2[OF \1 \ j\]) have norm: "pj.Mp D' = D'" "pj.Mp H' = H'" unfolding id by (auto simp: poly_mod.Mp_Mp) have mon: "monic D'" using pj.monic_Mp[OF step(11)] unfolding id . have id': "Mp (pj.Mp D) = Mp D" for D using \1 \ j\ by (simp add: Mp_Mp_pow_is_Mp p) have eq: "eq_m D1 D2 \ eq_m D1 (pj.Mp D2)" for D1 D2 unfolding id' by auto have id'': "pj.Mp (poly_mod.Mp (q * q) D) = pj.Mp D" for D unfolding qq by (rule pj.Mp_product_modulus[OF refl], insert p, auto) { fix D1 D2 assume "poly_mod.eq_m (q * q) D1 D2" hence "poly_mod.Mp (q * q) D1 = poly_mod.Mp (q * q) D2" by simp from arg_cong[OF this, of pj.Mp] have "pj.Mp D1 = pj.Mp D2" unfolding id'' . } note eq' = this from eq'[OF step(1)] have eq1: "pj.eq_m C (D' * H')" unfolding id by simp from eq'[OF step(2)] have eq2: "pj.eq_m (D' * S' + H' * T') 1" unfolding id by (metis pj.mult_Mp pj.plus_Mp) from **(4-5) have eq3: "eq_m D1 D'" "eq_m H1 H'" unfolding id by (auto intro: eq) from norm mon eq1 eq2 eq3 show ?thesis unfolding id by simp qed qed qed lemma quadratic_hensel_main: assumes res: "quadratic_hensel_main j = (D,H)" shows "poly_mod.eq_m (p^j) C (D * H)" "monic D" "poly_mod.eq_m p D1 D" "poly_mod.eq_m p H1 H" "poly_mod.Mp (p^j) D = D" "poly_mod.Mp (p^j) H = H" proof (atomize(full), goal_cases) case 1 let ?hen = "quadratic_hensel_loop j" from res obtain q S T where hen: "?hen = (q, S, T, D, H)" by (cases ?hen, auto simp: quadratic_hensel_main_def) from quadratic_hensel_loop[OF hen] show ?case by auto qed end end end datatype 'a factor_tree = Factor_Leaf 'a "int poly" | Factor_Node 'a "'a factor_tree" "'a factor_tree" fun factor_node_info :: "'a factor_tree \ 'a" where "factor_node_info (Factor_Leaf i x) = i" | "factor_node_info (Factor_Node i l r) = i" fun factors_of_factor_tree :: "'a factor_tree \ int poly multiset" where "factors_of_factor_tree (Factor_Leaf i x) = {#x#}" | "factors_of_factor_tree (Factor_Node i l r) = factors_of_factor_tree l + factors_of_factor_tree r" fun product_factor_tree :: "int \ 'a factor_tree \ int poly factor_tree" where "product_factor_tree p (Factor_Leaf i x) = (Factor_Leaf x x)" | "product_factor_tree p (Factor_Node i l r) = (let L = product_factor_tree p l; R = product_factor_tree p r; f = factor_node_info L; g = factor_node_info R; fg = poly_mod.Mp p (f * g) in Factor_Node fg L R)" fun sub_trees :: "'a factor_tree \ 'a factor_tree set" where "sub_trees (Factor_Leaf i x) = {Factor_Leaf i x}" | "sub_trees (Factor_Node i l r) = insert (Factor_Node i l r) (sub_trees l \ sub_trees r)" lemma sub_trees_refl[simp]: "t \ sub_trees t" by (cases t, auto) lemma product_factor_tree: assumes "\ x. x \# factors_of_factor_tree t \ poly_mod.Mp p x = x" shows "u \ sub_trees (product_factor_tree p t) \ factor_node_info u = f \ poly_mod.Mp p f = f \ f = poly_mod.Mp p (prod_mset (factors_of_factor_tree u)) \ factors_of_factor_tree (product_factor_tree p t) = factors_of_factor_tree t" using assms proof (induct t arbitrary: u f) case (Factor_Node i l r u f) interpret poly_mod p . let ?L = "product_factor_tree p l" let ?R = "product_factor_tree p r" let ?f = "factor_node_info ?L" let ?g = "factor_node_info ?R" let ?fg = "Mp (?f * ?g)" have "Mp ?f = ?f \ ?f = Mp (prod_mset (factors_of_factor_tree ?L)) \ (factors_of_factor_tree ?L) = (factors_of_factor_tree l)" by (rule Factor_Node(1)[OF sub_trees_refl refl], insert Factor_Node(5), auto) hence IH1: "?f = Mp (prod_mset (factors_of_factor_tree ?L))" "(factors_of_factor_tree ?L) = (factors_of_factor_tree l)" by blast+ have "Mp ?g = ?g \ ?g = Mp (prod_mset (factors_of_factor_tree ?R)) \ (factors_of_factor_tree ?R) = (factors_of_factor_tree r)" by (rule Factor_Node(2)[OF sub_trees_refl refl], insert Factor_Node(5), auto) hence IH2: "?g = Mp (prod_mset (factors_of_factor_tree ?R))" "(factors_of_factor_tree ?R) = (factors_of_factor_tree r)" by blast+ have id: "(factors_of_factor_tree (product_factor_tree p (Factor_Node i l r))) = (factors_of_factor_tree (Factor_Node i l r))" by (simp add: Let_def IH1 IH2) from Factor_Node(3) consider (root) "u = Factor_Node ?fg ?L ?R" | (l) "u \ sub_trees ?L" | (r) "u \ sub_trees ?R" by (auto simp: Let_def) thus ?case proof cases case root with Factor_Node have f: "f = ?fg" by auto show ?thesis unfolding f root id by (simp add: Let_def ac_simps IH1 IH2) next case l have "Mp f = f \ f = Mp (prod_mset (factors_of_factor_tree u))" using Factor_Node(1)[OF l Factor_Node(4)] Factor_Node(5) by auto thus ?thesis unfolding id by blast next case r have "Mp f = f \ f = Mp (prod_mset (factors_of_factor_tree u))" using Factor_Node(2)[OF r Factor_Node(4)] Factor_Node(5) by auto thus ?thesis unfolding id by blast qed qed auto fun create_factor_tree_simple :: "int poly list \ unit factor_tree" where "create_factor_tree_simple xs = (let n = length xs in if n \ 1 then Factor_Leaf () (hd xs) else let i = n div 2; xs1 = take i xs; xs2 = drop i xs in Factor_Node () (create_factor_tree_simple xs1) (create_factor_tree_simple xs2) )" declare create_factor_tree_simple.simps[simp del] lemma create_factor_tree_simple: "xs \ [] \ factors_of_factor_tree (create_factor_tree_simple xs) = mset xs" proof (induct xs rule: wf_induct[OF wf_measure[of length]]) case (1 xs) from 1(2) have xs: "length xs \ 0" by auto then consider (base) "length xs = 1" | (step) "length xs > 1" by linarith thus ?case proof cases case base then obtain x where xs: "xs = [x]" by (cases xs; cases "tl xs"; auto) thus ?thesis by (auto simp: create_factor_tree_simple.simps) next case step let ?i = "length xs div 2" let ?xs1 = "take ?i xs" let ?xs2 = "drop ?i xs" from step have xs1: "(?xs1, xs) \ measure length" "?xs1 \ []" by auto from step have xs2: "(?xs2, xs) \ measure length" "?xs2 \ []" by auto from step have id: "create_factor_tree_simple xs = Factor_Node () (create_factor_tree_simple (take ?i xs)) (create_factor_tree_simple (drop ?i xs))" unfolding create_factor_tree_simple.simps[of xs] Let_def by auto have xs: "xs = ?xs1 @ ?xs2" by auto show ?thesis unfolding id arg_cong[OF xs, of mset] mset_append using 1(1)[rule_format, OF xs1] 1(1)[rule_format, OF xs2] by auto qed qed text \We define a better factorization tree which balances the trees according to their degree., cf. Modern Computer Algebra, Chapter 15.5 on Multifactor Hensel lifting.\ fun partition_factors_main :: "nat \ ('a \ nat) list \ ('a \ nat) list \ ('a \ nat) list" where "partition_factors_main s [] = ([], [])" | "partition_factors_main s ((f,d) # xs) = (if d \ s then case partition_factors_main (s - d) xs of (l,r) \ ((f,d) # l, r) else case partition_factors_main d xs of (l,r) \ (l, (f,d) # r))" lemma partition_factors_main: "partition_factors_main s xs = (a,b) \ mset xs = mset a + mset b" by (induct s xs arbitrary: a b rule: partition_factors_main.induct, auto split: if_splits prod.splits) definition partition_factors :: "('a \ nat) list \ ('a \ nat) list \ ('a \ nat) list" where "partition_factors xs = (let n = sum_list (map snd xs) div 2 in case partition_factors_main n xs of ([], x # y # ys) \ ([x], y # ys) | (x # y # ys, []) \ ([x], y # ys) | pair \ pair)" lemma partition_factors: "partition_factors xs = (a,b) \ mset xs = mset a + mset b" unfolding partition_factors_def Let_def by (cases "partition_factors_main (sum_list (map snd xs) div 2) xs", auto split: list.splits simp: partition_factors_main) lemma partition_factors_length: assumes "\ length xs \ 1" "(a,b) = partition_factors xs" shows [termination_simp]: "length a < length xs" "length b < length xs" and "a \ []" "b \ []" proof - obtain ys zs where main: "partition_factors_main (sum_list (map snd xs) div 2) xs = (ys,zs)" by force note res = assms(2)[unfolded partition_factors_def Let_def main split] from arg_cong[OF partition_factors_main[OF main], of size] have len: "length xs = length ys + length zs" by auto with assms(1) have len2: "length ys + length zs \ 2" by auto from res len2 have "length a < length xs \ length b < length xs \ a \ [] \ b \ []" unfolding len by (cases ys; cases zs; cases "tl ys"; cases "tl zs"; auto) thus "length a < length xs" "length b < length xs" "a \ []" "b \ []" by blast+ qed fun create_factor_tree_balanced :: "(int poly \ nat)list \ unit factor_tree" where "create_factor_tree_balanced xs = (if length xs \ 1 then Factor_Leaf () (fst (hd xs)) else case partition_factors xs of (l,r) \ Factor_Node () (create_factor_tree_balanced l) (create_factor_tree_balanced r))" definition create_factor_tree :: "int poly list \ unit factor_tree" where "create_factor_tree xs = (let ys = map (\ f. (f, degree f)) xs; zs = rev (sort_key snd ys) in create_factor_tree_balanced zs)" lemma create_factor_tree_balanced: "xs \ [] \ factors_of_factor_tree (create_factor_tree_balanced xs) = mset (map fst xs)" proof (induct xs rule: create_factor_tree_balanced.induct) case (1 xs) show ?case proof (cases "length xs \ 1") case True with 1(3) obtain x where xs: "xs = [x]" by (cases xs; cases "tl xs", auto) show ?thesis unfolding xs by auto next case False obtain a b where part: "partition_factors xs = (a,b)" by force note abp = this[symmetric] note nonempty = partition_factors_length(3-4)[OF False abp] note IH = 1(1)[OF False abp nonempty(1)] 1(2)[OF False abp nonempty(2)] show ?thesis unfolding create_factor_tree_balanced.simps[of xs] part split using False IH partition_factors[OF part] by auto qed qed lemma create_factor_tree: assumes "xs \ []" shows "factors_of_factor_tree (create_factor_tree xs) = mset xs" proof - let ?xs = "rev (sort_key snd (map (\f. (f, degree f)) xs))" from assms have "set xs \ {}" by auto hence "set ?xs \ {}" by auto hence xs: "?xs \ []" by blast show ?thesis unfolding create_factor_tree_def Let_def create_factor_tree_balanced[OF xs] by (auto, induct xs, auto) qed context fixes p :: int and n :: nat begin definition quadratic_hensel_binary :: "int poly \ int poly \ int poly \ int poly \ int poly" where "quadratic_hensel_binary C D H = ( case euclid_ext_poly_dynamic p D H of (S,T) \ quadratic_hensel_main C p S T D H n)" fun hensel_lifting_main :: "int poly \ int poly factor_tree \ int poly list" where "hensel_lifting_main U (Factor_Leaf _ _) = [U]" | "hensel_lifting_main U (Factor_Node _ l r) = (let v = factor_node_info l; w = factor_node_info r; (V,W) = quadratic_hensel_binary U v w in hensel_lifting_main V l @ hensel_lifting_main W r)" definition hensel_lifting_monic :: "int poly \ int poly list \ int poly list" where "hensel_lifting_monic u vs = (if vs = [] then [] else let pn = p^n; C = poly_mod.Mp pn u; tree = product_factor_tree p (create_factor_tree vs) in hensel_lifting_main C tree)" definition hensel_lifting :: "int poly \ int poly list \ int poly list" where "hensel_lifting f gs = (let lc = lead_coeff f; ilc = inverse_mod lc (p^n); g = smult ilc f in hensel_lifting_monic g gs)" end context poly_mod_prime begin context fixes n :: nat assumes n: "n \ 0" begin abbreviation "hensel_binary \ quadratic_hensel_binary p n" abbreviation "hensel_main \ hensel_lifting_main p n" lemma hensel_binary: assumes cop: "coprime_m D H" and eq: "eq_m C (D * H)" and normalized_input: "Mp D = D" "Mp H = H" and monic_input: "monic D" and hensel_result: "hensel_binary C D H = (D',H')" shows "poly_mod.eq_m (p^n) C (D' * H') \ \the main result: equivalence mod \p^n\\ \ monic D' \ \monic output\ \ eq_m D D' \ eq_m H H' \ \apply \`mod p`\ on \D'\ and \H'\ yields \D\ and \H\ again\ \ poly_mod.Mp (p^n) D' = D' \ poly_mod.Mp (p^n) H' = H' \ \output is normalized\" proof - from m1 have p: "p > 1" . obtain S T where ext: "euclid_ext_poly_dynamic p D H = (S,T)" by force obtain D1 H1 where main: "quadratic_hensel_main C p S T D H n = (D1,H1)" by force note hen = hensel_result[unfolded quadratic_hensel_binary_def ext split Let_def main] from n have n: "n \ 1" by simp note eucl = euclid_ext_poly_dynamic[OF cop normalized_input ext] note main = quadratic_hensel_main[OF eucl(1) eq monic_input p normalized_input eucl(2-) n main] show ?thesis using hen main by auto qed lemma hensel_main: assumes eq: "eq_m C (prod_mset (factors_of_factor_tree Fs))" and "\ F. F \# factors_of_factor_tree Fs \ Mp F = F \ monic F" and hensel_result: "hensel_main C Fs = Gs" and C: "monic C" "poly_mod.Mp (p^n) C = C" and sf: "square_free_m C" and "\ f t. t \ sub_trees Fs \ factor_node_info t = f \ f = Mp (prod_mset (factors_of_factor_tree t))" shows "poly_mod.eq_m (p^n) C (prod_list Gs) \ \the main result: equivalence mod \p^n\\ \ factors_of_factor_tree Fs = mset (map Mp Gs) \ (\ G. G \ set Gs \ monic G \ poly_mod.Mp (p^n) G = G)" using assms proof (induct Fs arbitrary: C Gs) case (Factor_Leaf f fs C Gs) thus ?case by auto next case (Factor_Node f l r C Gs) note * = this note simps = hensel_lifting_main.simps note IH1 = *(1)[rule_format] note IH2 = *(2)[rule_format] note res = *(5)[unfolded simps Let_def] note eq = *(3) note Fs = *(4) note C = *(6,7) note sf = *(8) note inv = *(9) interpret pn: poly_mod_2 "p^n" apply (unfold_locales) using m1 n by auto let ?Mp = "pn.Mp" define D where "D \ prod_mset (factors_of_factor_tree l)" define H where "H \ prod_mset (factors_of_factor_tree r)" let ?D = "Mp D" let ?H = "Mp H" let ?D' = "factor_node_info l" let ?H' = "factor_node_info r" obtain A B where hen: "hensel_binary C ?D' ?H' = (A,B)" by force note res = res[unfolded hen split] obtain AD where AD': "AD = hensel_main A l" by auto obtain BH where BH': "BH = hensel_main B r" by auto from inv[of l, OF _ refl] have D': "?D' = ?D" unfolding D_def by auto from inv[of r, OF _ refl] have H': "?H' = ?H" unfolding H_def by auto from eq[simplified] have eq': "Mp C = Mp (?D * ?H)" unfolding D_def H_def by simp from square_free_m_cong[OF sf, of "?D * ?H", OF eq'] have sf': "square_free_m (?D * ?H)" . from poly_mod_prime.square_free_m_prod_imp_coprime_m[OF _ this] have cop': "coprime_m ?D ?H" unfolding poly_mod_prime_def using prime . from eq' have eq': "eq_m C (?D * ?H)" by simp have monD: "monic D" unfolding D_def by (rule monic_prod_mset, insert Fs, auto) from hensel_binary[OF _ _ _ _ _ hen, unfolded D' H', OF cop' eq' Mp_Mp Mp_Mp monic_Mp[OF monD]] have step: "poly_mod.eq_m (p ^ n) C (A * B) \ monic A \ eq_m ?D A \ eq_m ?H B \ ?Mp A = A \ ?Mp B = B" . from res have Gs: "Gs = AD @ BH" by (simp add: AD' BH') have AD: "eq_m A ?D" "?Mp A = A" "eq_m A (prod_mset (factors_of_factor_tree l))" and monA: "monic A" using step by (auto simp: D_def) note sf_fact = square_free_m_factor[OF sf'] from square_free_m_cong[OF sf_fact(1)] AD have sfA: "square_free_m A" by auto have IH1: "poly_mod.eq_m (p ^ n) A (prod_list AD) \ factors_of_factor_tree l = mset (map Mp AD) \ (\G. G \ set AD \ monic G \ ?Mp G = G)" by (rule IH1[OF AD(3) Fs AD'[symmetric] monA AD(2) sfA inv], auto) have BH: "eq_m B ?H" "pn.Mp B = B" "eq_m B (prod_mset (factors_of_factor_tree r))" using step by (auto simp: H_def) from step have "pn.eq_m C (A * B)" by simp hence "?Mp C = ?Mp (A * B)" by simp with C AD(2) have "pn.Mp C = pn.Mp (A * pn.Mp B)" by simp from arg_cong[OF this, of lead_coeff] C have "monic (pn.Mp (A * B))" by simp then have "lead_coeff (pn.Mp A) * lead_coeff (pn.Mp B) = 1" by (metis lead_coeff_mult leading_coeff_neq_0 local.step mult_cancel_right2 pn.degree_m_eq pn.m1 poly_mod.M_def poly_mod.Mp_coeff) with monA AD(2) BH(2) have monB: "monic B" by simp from square_free_m_cong[OF sf_fact(2)] BH have sfB: "square_free_m B" by auto have IH2: "poly_mod.eq_m (p ^ n) B (prod_list BH) \ factors_of_factor_tree r = mset (map Mp BH) \ (\G. G \ set BH \ monic G \ ?Mp G = G)" by (rule IH2[OF BH(3) Fs BH'[symmetric] monB BH(2) sfB inv], auto) from step have "?Mp C = ?Mp (?Mp A * ?Mp B)" by auto also have "?Mp A = ?Mp (prod_list AD)" using IH1 by auto also have "?Mp B = ?Mp (prod_list BH)" using IH2 by auto finally have "poly_mod.eq_m (p ^ n) C (prod_list AD * prod_list BH)" by (auto simp: poly_mod.mult_Mp) thus ?case unfolding Gs using IH1 IH2 by auto qed lemma hensel_lifting_monic: assumes eq: "poly_mod.eq_m p C (prod_list Fs)" and Fs: "\ F. F \ set Fs \ poly_mod.Mp p F = F \ monic F" and res: "hensel_lifting_monic p n C Fs = Gs" and mon: "monic (poly_mod.Mp (p^n) C)" and sf: "poly_mod.square_free_m p C" shows "poly_mod.eq_m (p^n) C (prod_list Gs)" "mset (map (poly_mod.Mp p) Gs) = mset Fs" "G \ set Gs \ monic G \ poly_mod.Mp (p^n) G = G" proof - note res = res[unfolded hensel_lifting_monic_def Let_def] let ?Mp = "poly_mod.Mp (p ^ n)" let ?C = "?Mp C" interpret poly_mod_prime p by (unfold_locales, insert n prime, auto) interpret pn: poly_mod_2 "p^n" using m1 n poly_mod_2.intro by auto from eq n have eq: "eq_m (?Mp C) (prod_list Fs)" using Mp_Mp_pow_is_Mp eq m1 n by force have "poly_mod.eq_m (p^n) C (prod_list Gs) \ mset (map (poly_mod.Mp p) Gs) = mset Fs \ (G \ set Gs \ monic G \ poly_mod.Mp (p^n) G = G)" proof (cases "Fs = []") case True with res have Gs: "Gs = []" by auto from eq have "Mp ?C = 1" unfolding True by simp hence "degree (Mp ?C) = 0" by simp with degree_m_eq_monic[OF mon m1] have "degree ?C = 0" by simp with mon have "?C = 1" using monic_degree_0 by blast thus ?thesis unfolding True Gs by auto next case False let ?t = "create_factor_tree Fs" note tree = create_factor_tree[OF False] from False res have hen: "hensel_main ?C (product_factor_tree p ?t) = Gs" by auto have tree1: "x \# factors_of_factor_tree ?t \ Mp x = x" for x unfolding tree using Fs by auto from product_factor_tree[OF tree1 sub_trees_refl refl, of ?t] have id: "(factors_of_factor_tree (product_factor_tree p ?t)) = (factors_of_factor_tree ?t)" by auto have eq: "eq_m ?C (prod_mset (factors_of_factor_tree (product_factor_tree p ?t)))" unfolding id tree using eq by auto have id': "Mp C = Mp ?C" using n by (simp add: Mp_Mp_pow_is_Mp m1) have "pn.eq_m ?C (prod_list Gs) \ mset Fs = mset (map Mp Gs) \ (\G. G \ set Gs \ monic G \ pn.Mp G = G)" by (rule hensel_main[OF eq Fs hen mon pn.Mp_Mp square_free_m_cong[OF sf id'], unfolded id tree], insert product_factor_tree[OF tree1], auto) thus ?thesis by auto qed thus "poly_mod.eq_m (p^n) C (prod_list Gs)" "mset (map (poly_mod.Mp p) Gs) = mset Fs" "G \ set Gs \ monic G \ poly_mod.Mp (p^n) G = G" by blast+ qed lemma hensel_lifting: assumes res: "hensel_lifting p n f fs = gs" \ \result of hensel is fact. \gs\\ and cop: "coprime (lead_coeff f) p" and sf: "poly_mod.square_free_m p f" and fact: "poly_mod.factorization_m p f (c, mset fs)" \ \input is fact. \fs mod p\\ and c: "c \ {0..fi\set fs. set (coeffs fi) \ {0.. \factorization mod \p^n\\" "sort (map degree fs) = sort (map degree gs) \ \degrees stay the same\" "\ g. g \ set gs \ monic g \ poly_mod.Mp (p^n) g = g \ \ \monic and normalized\ irreducible_m g \ \ \irreducibility even mod \p\\ degree_m g = degree g \ \mod \p\ does not change degree of \g\\" proof - interpret poly_mod_prime p using prime by unfold_locales interpret q: poly_mod_2 "p^n" using m1 n unfolding poly_mod_2_def by auto from fact have eq: "eq_m f (smult c (prod_list fs))" and mon_fs: "(\fi\set fs. monic (Mp fi) \ irreducible\<^sub>d_m fi)" unfolding factorization_m_def by auto { fix f assume "f \ set fs" with mon_fs norm have "set (coeffs f) \ {0.. f. Mp (q.Mp f) = Mp f" by (simp add: Mp_Mp_pow_is_Mp m1 n) let ?lc = "lead_coeff f" let ?q = "p ^ n" define ilc where "ilc \ inverse_mod ?lc ?q" define F where "F \ smult ilc f" from res[unfolded hensel_lifting_def Let_def] have hen: "hensel_lifting_monic p n F fs = gs" unfolding ilc_def F_def . from m1 n cop have inv: "q.M (ilc * ?lc) = 1" by (auto simp add: q.M_def inverse_mod_pow ilc_def) hence ilc0: "ilc \ 0" by (cases "ilc = 0", auto) { fix q assume "ilc * ?lc = ?q * q" from arg_cong[OF this, of q.M] have "q.M (ilc * ?lc) = 0" unfolding q.M_def by auto with inv have False by auto } note not_dvd = this have mon: "monic (q.Mp F)" unfolding F_def q.Mp_coeff coeff_smult by (subst q.degree_m_eq [OF _ q.m1]) (auto simp: inv ilc0 [symmetric] intro: not_dvd) have "q.Mp f = q.Mp (smult (q.M (?lc * ilc)) f)" using inv by (simp add: ac_simps) also have "\ = q.Mp (smult ?lc F)" by (simp add: F_def) finally have f: "q.Mp f = q.Mp (smult ?lc F)" . from arg_cong[OF f, of Mp] have f_p: "Mp f = Mp (smult ?lc F)" by (simp add: Mp_Mp_pow_is_Mp n m1) from arg_cong[OF this, of square_free_m, unfolded Mp_square_free_m] sf have "square_free_m (smult ?lc F)" by simp from square_free_m_smultD[OF this] have sf: "square_free_m F" . define c' where "c' \ M (c * ilc)" from factorization_m_smult[OF fact, of ilc, folded F_def] have fact: "factorization_m F (c', mset fs)" unfolding c'_def factorization_m_def by auto hence eq: "eq_m F (smult c' (prod_list fs))" unfolding factorization_m_def by auto from factorization_m_lead_coeff[OF fact] monic_Mp[OF mon, unfolded Mp_id] have "M c' = 1" by auto hence c': "c' = 1" unfolding c'_def by auto with eq have eq: "eq_m F (prod_list fs)" by auto { fix f assume "f \ set fs" with mon_fs' norm have "Mp f = f \ monic f" unfolding Mp_ident_iff' by auto } note fs = this note hen = hensel_lifting_monic[OF eq fs hen mon sf] from hen(2) have gs_fs: "mset (map Mp gs) = mset fs" by auto have eq: "q.eq_m f (smult ?lc (prod_list gs))" unfolding f using arg_cong[OF hen(1), of "\ f. q.Mp (smult ?lc f)"] by simp { fix g assume g: "g \ set gs" from hen(3)[OF _ g] have mon_g: "monic g" and Mp_g: "q.Mp g = g" by auto from g have "Mp g \# mset (map Mp gs)" by auto from this[unfolded gs_fs] obtain f where f: "f \ set fs" and fg: "eq_m f g" by auto from mon_fs f fs have irr_f: "irreducible\<^sub>d_m f" and mon_f: "monic f" and Mp_f: "Mp f = f" by auto have deg: "degree_m g = degree g" by (rule degree_m_eq_monic[OF mon_g m1]) from irr_f fg have irr_g: "irreducible\<^sub>d_m g" unfolding irreducible\<^sub>d_m_def dvdm_def by simp have "q.irreducible\<^sub>d_m g" by (rule irreducible\<^sub>d_lifting[OF n _ irr_g], unfold deg, rule q.degree_m_eq_monic[OF mon_g q.m1]) note mon_g Mp_g deg irr_g this } note g = this { fix g assume "g \ set gs" from g[OF this] show "monic g \ q.Mp g = g \ irreducible_m g \ degree_m g = degree g" by auto } show "sort (map degree fs) = sort (map degree gs)" proof (rule sort_key_eq_sort_key) have "mset (map degree fs) = image_mset degree (mset fs)" by auto also have "\ = image_mset degree (mset (map Mp gs))" unfolding gs_fs .. also have "\ = mset (map degree (map Mp gs))" unfolding mset_map .. also have "map degree (map Mp gs) = map degree_m gs" by auto also have "\ = map degree gs" using g(3) by auto finally show "mset (map degree fs) = mset (map degree gs)" . qed auto show "q.factorization_m f (lead_coeff f, mset gs)" using eq g unfolding q.factorization_m_def by auto qed end end end diff --git a/thys/Berlekamp_Zassenhaus/Square_Free_Factorization_Int.thy b/thys/Berlekamp_Zassenhaus/Square_Free_Factorization_Int.thy --- a/thys/Berlekamp_Zassenhaus/Square_Free_Factorization_Int.thy +++ b/thys/Berlekamp_Zassenhaus/Square_Free_Factorization_Int.thy @@ -1,702 +1,702 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada *) theory Square_Free_Factorization_Int imports Square_Free_Int_To_Square_Free_GFp Suitable_Prime Code_Abort_Gcd Gcd_Finite_Field_Impl begin definition yun_wrel :: "int poly \ rat \ rat poly \ bool" where "yun_wrel F c f = (map_poly rat_of_int F = smult c f)" definition yun_rel :: "int poly \ rat \ rat poly \ bool" where "yun_rel F c f = (yun_wrel F c f \ content F = 1 \ lead_coeff F > 0 \ monic f)" definition yun_erel :: "int poly \ rat poly \ bool" where "yun_erel F f = (\ c. yun_rel F c f)" lemma yun_wrelD: assumes "yun_wrel F c f" shows "map_poly rat_of_int F = smult c f" using assms unfolding yun_wrel_def by auto lemma yun_relD: assumes "yun_rel F c f" shows "yun_wrel F c f" "map_poly rat_of_int F = smult c f" "degree F = degree f" "F \ 0" "lead_coeff F > 0" "monic f" "f = 1 \ F = 1" "content F = 1" proof - note * = assms[unfolded yun_rel_def yun_wrel_def, simplified] then have "degree (map_poly rat_of_int F) = degree f" by auto then show deg: "degree F = degree f" by simp show "F \ 0" "lead_coeff F > 0" "monic f" "content F = 1" "map_poly rat_of_int F = smult c f" "yun_wrel F c f" using * by (auto simp: yun_wrel_def) { assume "f = 1" with deg have "degree F = 0" by auto from degree0_coeffs[OF this] obtain c where F: "F = [:c:]" and c: "c = lead_coeff F" by auto from c * have c0: "c > 0" by auto hence cF: "content F = c" unfolding F content_def by auto with * have "c = 1" by auto with F have "F = 1" by simp } moreover { assume "F = 1" with deg have "degree f = 0" by auto with \monic f\ have "f = 1" using monic_degree_0 by blast } ultimately show "(f = 1) \ (F = 1)" by auto qed lemma yun_erel_1_eq: assumes "yun_erel F f" shows "(F = 1) \ (f = 1)" proof - from assms[unfolded yun_erel_def] obtain c where "yun_rel F c f" by auto from yun_relD[OF this] show ?thesis by simp qed lemma yun_rel_1[simp]: "yun_rel 1 1 1" by (auto simp: yun_rel_def yun_wrel_def content_def) lemma yun_erel_1[simp]: "yun_erel 1 1" unfolding yun_erel_def using yun_rel_1 by blast lemma yun_rel_mult: "yun_rel F c f \ yun_rel G d g \ yun_rel (F * G) (c * d) (f * g)" unfolding yun_rel_def yun_wrel_def content_mult lead_coeff_mult by (auto simp: monic_mult hom_distribs) lemma yun_erel_mult: "yun_erel F f \ yun_erel G g \ yun_erel (F * G) (f * g)" unfolding yun_erel_def using yun_rel_mult[of F _ f G _ g] by blast lemma yun_rel_pow: assumes "yun_rel F c f" shows "yun_rel (F^n) (c^n) (f^n)" by (induct n, insert assms yun_rel_mult, auto) lemma yun_erel_pow: "yun_erel F f \ yun_erel (F^n) (f^n)" using yun_rel_pow unfolding yun_erel_def by blast lemma yun_wrel_pderiv: assumes "yun_wrel F c f" shows "yun_wrel (pderiv F) c (pderiv f)" by (unfold yun_wrel_def, simp add: yun_wrelD[OF assms] pderiv_smult hom_distribs) lemma yun_wrel_minus: assumes "yun_wrel F c f" "yun_wrel G c g" shows "yun_wrel (F - G) c (f - g)" using assms unfolding yun_wrel_def by (auto simp: smult_diff_right hom_distribs) lemma gcd_smult_left: assumes "c \ 0" - shows "gcd (smult c f) g = gcd f (g :: 'b :: {field, factorial_ring_gcd} poly)" + shows "gcd (smult c f) g = gcd f (g :: 'b :: {field_gcd} poly)" proof - from assms have "normalize c = 1" by (meson dvd_field_iff is_unit_normalize) then show ?thesis by (metis (no_types) Polynomial.normalize_smult gcd.commute gcd.left_commute gcd_left_idem gcd_self smult_1_left) qed -lemma gcd_smult_right: "c \ 0 \ gcd f (smult c g) = gcd f (g :: 'b :: {field, factorial_ring_gcd} poly)" +lemma gcd_smult_right: "c \ 0 \ gcd f (smult c g) = gcd f (g :: 'b :: {field_gcd} poly)" using gcd_smult_left[of c g f] by (simp add: gcd.commute) lemma gcd_rat_to_gcd_int: "gcd (of_int_poly f :: rat poly) (of_int_poly g) = smult (inverse (of_int (lead_coeff (gcd f g)))) (of_int_poly (gcd f g))" proof (cases "f = 0 \ g = 0") case True thus ?thesis by simp next case False let ?r = rat_of_int let ?rp = "map_poly ?r" from False have gcd0: "gcd f g \ 0" by auto hence lc0: "lead_coeff (gcd f g) \ 0" by auto hence inv: "inverse (?r (lead_coeff (gcd f g))) \ 0" by auto show ?thesis proof (rule sym, rule gcdI, goal_cases) case 1 have "gcd f g dvd f" by auto then obtain h where f: "f = gcd f g * h" unfolding dvd_def by auto show ?case by (rule smult_dvd[OF _ inv], insert arg_cong[OF f, of ?rp], simp add: hom_distribs) next case 2 have "gcd f g dvd g" by auto then obtain h where g: "g = gcd f g * h" unfolding dvd_def by auto show ?case by (rule smult_dvd[OF _ inv], insert arg_cong[OF g, of ?rp], simp add: hom_distribs) next case (3 h) show ?case proof (rule dvd_smult) obtain ch ph where h: "rat_to_normalized_int_poly h = (ch, ph)" by force from 3 obtain ff where f: "?rp f = h * ff" unfolding dvd_def by auto from 3 obtain gg where g: "?rp g = h * gg" unfolding dvd_def by auto from rat_to_int_factor_explicit[OF f h] obtain f' where f: "f = ph * f'" by blast from rat_to_int_factor_explicit[OF g h] obtain g' where g: "g = ph * g'" by blast from f g have "ph dvd gcd f g" by auto then obtain gg where gcd: "gcd f g = ph * gg" unfolding dvd_def by auto note * = rat_to_normalized_int_poly[OF h] show "h dvd ?rp (gcd f g)" unfolding gcd *(1) by (rule smult_dvd, insert *(2), auto) qed next case 4 show ?case unfolding normalize_poly_def by (rule poly_eqI) (simp add: one_poly_def [symmetric]) qed qed lemma yun_wrel_div: assumes f: "yun_wrel F c f" and g: "yun_wrel G d g" and dvd: "G dvd F" "g dvd f" and G0: "G \ 0" shows "yun_wrel (F div G) (c / d) (f div g)" proof - let ?r = "rat_of_int" let ?rp = "map_poly ?r" from dvd obtain H h where fgh: "F = G * H" "f = g * h" unfolding dvd_def by auto from G0 yun_wrelD[OF g] have g0: "g \ 0" and d0: "d \ 0" by auto from arg_cong[OF fgh(1), of "\ x. x div G"] have H: "H = F div G" using G0 by simp from arg_cong[OF fgh(1), of ?rp] have "?rp F = ?rp G * ?rp H" by (auto simp: hom_distribs) from arg_cong[OF this, of "\ x. x div ?rp G"] G0 have id: "?rp H = ?rp F div ?rp G" by auto have "?rp (F div G) = ?rp F div ?rp G" unfolding H[symmetric] id by simp also have "\ = smult c f div smult d g" using f g unfolding yun_wrel_def by auto also have "\ = smult (c / d) (f div g)" unfolding div_smult_right[OF d0] div_smult_left by (simp add: field_simps) finally show ?thesis unfolding yun_wrel_def by simp qed lemma yun_rel_div: assumes f: "yun_rel F c f" and g: "yun_rel G d g" and dvd: "G dvd F" "g dvd f" shows "yun_rel (F div G) (c / d) (f div g)" proof - note ff = yun_relD[OF f] note gg = yun_relD[OF g] show ?thesis unfolding yun_rel_def proof (intro conjI) from yun_wrel_div[OF ff(1) gg(1) dvd gg(4)] show "yun_wrel (F div G) (c / d) (f div g)" by auto from dvd have fg: "f = g * (f div g)" by auto from arg_cong[OF fg, of monic] ff(6) gg(6) show "monic (f div g)" using monic_factor by blast from dvd have FG: "F = G * (F div G)" by auto from arg_cong[OF FG, of content, unfolded content_mult] ff(8) gg(8) show "content (F div G) = 1" by simp from arg_cong[OF FG, of lead_coeff, unfolded lead_coeff_mult] ff(5) gg(5) show "lead_coeff (F div G) > 0" by (simp add: zero_less_mult_iff) qed qed lemma yun_wrel_gcd: assumes "yun_wrel F c' f" "yun_wrel G c g" and c: "c' \ 0" "c \ 0" and d: "d = rat_of_int (lead_coeff (gcd F G))" "d \ 0" shows "yun_wrel (gcd F G) d (gcd f g)" proof - let ?r = "rat_of_int" let ?rp = "map_poly ?r" have "smult d (gcd f g) = smult d (gcd (smult c' f) (smult c g))" by (simp add: c gcd_smult_left gcd_smult_right) also have "\ = smult d (gcd (?rp F) (?rp G))" using assms(1-2)[unfolded yun_wrel_def] by simp also have "\ = smult (d * inverse d) (?rp (gcd F G))" unfolding gcd_rat_to_gcd_int d by simp also have "d * inverse d = 1" using d by auto finally show ?thesis unfolding yun_wrel_def by simp qed lemma yun_rel_gcd: assumes f: "yun_rel F c f" and g: "yun_wrel G c' g" and c': "c' \ 0" and d: "d = rat_of_int (lead_coeff (gcd F G))" shows "yun_rel (gcd F G) d (gcd f g)" unfolding yun_rel_def proof (intro conjI) note ff = yun_relD[OF f] from ff have c0: "c \ 0" by auto from ff d have d0: "d \ 0" by auto from yun_wrel_gcd[OF ff(1) g c0 c' d d0] show "yun_wrel (gcd F G) d (gcd f g)" by auto from ff have "gcd f g \ 0" by auto thus "monic (gcd f g)" by (simp add: poly_gcd_monic) obtain H where H: "gcd F G = H" by auto obtain lc where lc: "coeff H (degree H) = lc" by auto from ff have "gcd F G \ 0" by auto hence "H \ 0" "lc \ 0" unfolding H[symmetric] lc[symmetric] by auto thus "0 < lead_coeff (gcd F G)" unfolding arg_cong[OF normalize_gcd[of F G], of lead_coeff, symmetric] unfolding normalize_poly_eq_map_poly H by (auto, subst Polynomial.coeff_map_poly, auto, subst Polynomial.degree_map_poly, auto simp: sgn_if) have "H dvd F" unfolding H[symmetric] by auto then obtain K where F: "F = H * K" unfolding dvd_def by auto from arg_cong[OF this, of content, unfolded content_mult ff(8)] content_ge_0_int[of H] have "content H = 1" by (auto simp add: zmult_eq_1_iff) thus "content (gcd F G) = 1" unfolding H . qed lemma yun_factorization_main_int: assumes f: "f = p div gcd p (pderiv p)" and "g = pderiv p div gcd p (pderiv p)" "monic p" and "yun_gcd.yun_factorization_main gcd f g i hs = res" and "yun_gcd.yun_factorization_main gcd F G i Hs = Res" and "yun_rel F c f" "yun_wrel G c g" "list_all2 (rel_prod yun_erel (=)) Hs hs" shows "list_all2 (rel_prod yun_erel (=)) Res res" proof - let ?P = "\ f g. \ i hs res F G Hs Res c. yun_gcd.yun_factorization_main gcd f g i hs = res \ yun_gcd.yun_factorization_main gcd F G i Hs = Res \ yun_rel F c f \ yun_wrel G c g \ list_all2 (rel_prod yun_erel (=)) Hs hs \ list_all2 (rel_prod yun_erel (=)) Res res" note simps = yun_gcd.yun_factorization_main.simps note rel = yun_relD let ?rel = "\ F f. map_poly rat_of_int F = smult (rat_of_int (lead_coeff F)) f" show ?thesis proof (induct rule: yun_factorization_induct[of ?P, rule_format, OF _ _ assms]) case (1 f g i hs res F G Hs Res c) from rel[OF 1(4)] 1(1) have "f = 1" "F = 1" by auto from 1(2-3)[unfolded simps[of _ 1] this] have "res = hs" "Res = Hs" by auto with 1(6) show ?case by simp next case (2 f g i hs res F G Hs Res c) define d where "d = g - pderiv f" define a where "a = gcd f d" define D where "D = G - pderiv F" define A where "A = gcd F D" note f = 2(5) note g = 2(6) note hs = 2(7) note f1 = 2(1) from f1 rel[OF f] have *: "(f = 1) = False" "(F = 1) = False" and c: "c \ 0" by auto note res = 2(3)[unfolded simps[of _ f] * if_False Let_def, folded d_def a_def] note Res = 2(4)[unfolded simps[of _ F] * if_False Let_def, folded D_def A_def] note IH = 2(2)[folded d_def a_def, OF res Res] obtain c' where c': "c' = rat_of_int (lead_coeff (gcd F D))" by auto show ?case proof (rule IH) from yun_wrel_minus[OF g yun_wrel_pderiv[OF rel(1)[OF f]]] have d: "yun_wrel D c d" unfolding D_def d_def . have a: "yun_rel A c' a" unfolding A_def a_def by (rule yun_rel_gcd[OF f d c c']) hence "yun_erel A a" unfolding yun_erel_def by auto thus "list_all2 (rel_prod yun_erel (=)) ((A, i) # Hs) ((a, i) # hs)" using hs by auto have A0: "A \ 0" by (rule rel(4)[OF a]) have "A dvd D" "a dvd d" unfolding A_def a_def by auto from yun_wrel_div[OF d rel(1)[OF a] this A0] show "yun_wrel (D div A) (c / c') (d div a)" . have "A dvd F" "a dvd f" unfolding A_def a_def by auto from yun_rel_div[OF f a this] show "yun_rel (F div A) (c / c') (f div a)" . qed qed qed lemma yun_monic_factorization_int_yun_rel: assumes res: "yun_gcd.yun_monic_factorization gcd f = res" and Res: "yun_gcd.yun_monic_factorization gcd F = Res" and f: "yun_rel F c f" shows "list_all2 (rel_prod yun_erel (=)) Res res" proof - note ff = yun_relD[OF f] let ?g = "gcd f (pderiv f)" let ?yf = "yun_gcd.yun_factorization_main gcd (f div ?g) (pderiv f div ?g) 0 []" let ?G = "gcd F (pderiv F)" let ?yF = "yun_gcd.yun_factorization_main gcd (F div ?G) (pderiv F div ?G) 0 []" obtain r R where r: "?yf = r" and R: "?yF = R" by blast from res[unfolded yun_gcd.yun_monic_factorization_def Let_def r] have res: "res = [(a, i)\r . a \ 1]" by simp from Res[unfolded yun_gcd.yun_monic_factorization_def Let_def R] have Res: "Res = [(A, i)\R . A \ 1]" by simp from yun_wrel_pderiv[OF ff(1)] have f': "yun_wrel (pderiv F) c (pderiv f)" . from ff have c: "c \ 0" by auto from yun_rel_gcd[OF f f' c refl] obtain d where g: "yun_rel ?G d ?g" .. from yun_rel_div[OF f g] have 1: "yun_rel (F div ?G) (c / d) (f div ?g)" by auto from yun_wrel_div[OF f' yun_relD(1)[OF g] _ _ yun_relD(4)[OF g]] have 2: "yun_wrel (pderiv F div ?G) (c / d) (pderiv f div ?g)" by auto from yun_factorization_main_int[OF refl refl ff(6) r R 1 2] have "list_all2 (rel_prod yun_erel (=)) R r" by simp thus ?thesis unfolding res Res by (induct R r rule: list_all2_induct, auto dest: yun_erel_1_eq) qed lemma yun_rel_same_right: assumes "yun_rel f c G" "yun_rel g d G" shows "f = g" proof - note f = yun_relD[OF assms(1)] note g = yun_relD[OF assms(2)] let ?r = "rat_of_int" let ?rp = "map_poly ?r" from g have d: "d \ 0" by auto obtain a b where quot: "quotient_of (c / d) = (a,b)" by force from quotient_of_nonzero[of "c/d", unfolded quot] have b: "b \ 0" by simp note f(2) also have "smult c G = smult (c / d) (smult d G)" using d by (auto simp: field_simps) also have "smult d G = ?rp g" using g(2) by simp also have cd: "c / d = (?r a / ?r b)" using quotient_of_div[OF quot] . finally have fg: "?rp f = smult (?r a / ?r b) (?rp g)" by simp from f have "c \ 0" by auto with cd d have a: "a \ 0" by auto from arg_cong[OF fg, of "\ x. smult (?r b) x"] have "smult (?r b) (?rp f) = smult (?r a) (?rp g)" using b by auto hence "?rp (smult b f) = ?rp (smult a g)" by (auto simp: hom_distribs) then have fg: "[:b:] * f = [:a:] * g" by auto from arg_cong[OF this, of content, unfolded content_mult f(8) g(8)] have "content [: b :] = content [: a :]" by simp hence abs: "abs a = abs b" unfolding content_def using b a by auto from arg_cong[OF fg, of "\ x. lead_coeff x > 0", unfolded lead_coeff_mult] f(5) g(5) a b have "(a > 0) = (b > 0)" by (simp add: zero_less_mult_iff) with a b abs have "a = b" by auto with arg_cong[OF fg, of "\ x. x div [:b:]"] b show ?thesis by (metis nonzero_mult_div_cancel_left pCons_eq_0_iff) qed definition square_free_factorization_int_main :: "int poly \ (int poly \ nat) list" where "square_free_factorization_int_main f = (case square_free_heuristic f of None \ yun_gcd.yun_monic_factorization gcd f | Some p \ [(f,0)])" lemma square_free_factorization_int_main: assumes res: "square_free_factorization_int_main f = fs" and ct: "content f = 1" and lc: "lead_coeff f > 0" and deg: "degree f \ 0" shows "square_free_factorization f (1,fs) \ (\ fi i. (fi, i) \ set fs \ content fi = 1 \ lead_coeff fi > 0) \ distinct (map snd fs)" proof (cases "square_free_heuristic f") case None from lc have f0: "f \ 0" by auto from res None have fs: "yun_gcd.yun_monic_factorization gcd f = fs" unfolding square_free_factorization_int_main_def by auto let ?r = "rat_of_int" let ?rp = "map_poly ?r" define G where "G = smult (inverse (lead_coeff (?rp f))) (?rp f)" have "?rp f \ 0" using f0 by auto hence mon: "monic G" unfolding G_def coeff_smult by simp obtain Fs where Fs: "yun_gcd.yun_monic_factorization gcd G = Fs" by blast from lc have lg: "lead_coeff (?rp f) \ 0" by auto let ?c = "lead_coeff (?rp f)" define c where "c = ?c" have rp: "?rp f = smult c G" unfolding G_def c_def by (simp add: field_simps) have in_rel: "yun_rel f c G" unfolding yun_rel_def yun_wrel_def using rp mon lc ct by auto from yun_monic_factorization_int_yun_rel[OF Fs fs in_rel] have out_rel: "list_all2 (rel_prod yun_erel (=)) fs Fs" by auto from yun_monic_factorization[OF Fs mon] have "square_free_factorization G (1, Fs)" and dist: "distinct (map snd Fs)" by auto note sff = square_free_factorizationD[OF this(1)] from out_rel have "map snd fs = map snd Fs" by (induct fs Fs rule: list_all2_induct, auto) with dist have dist': "distinct (map snd fs)" by auto have main: "square_free_factorization f (1, fs) \ (\ fi i. (fi, i) \ set fs \ content fi = 1 \ lead_coeff fi > 0)" unfolding square_free_factorization_def split proof (intro conjI allI impI) from ct have "f \ 0" by auto thus "f = 0 \ 1 = 0" "f = 0 \ fs = []" by auto from dist' show "distinct fs" by (simp add: distinct_map) { fix a i assume a: "(a,i) \ set fs" with out_rel obtain bj where "bj \ set Fs" and "rel_prod yun_erel (=) (a,i) bj" unfolding list_all2_conv_all_nth set_conv_nth by fastforce then obtain b where b: "(b,i) \ set Fs" and ab: "yun_erel a b" by (cases bj, auto simp: rel_prod.simps) from sff(2)[OF b] have b': "square_free b" "degree b \ 0" by auto from ab obtain c where rel: "yun_rel a c b" unfolding yun_erel_def by auto note aa = yun_relD[OF this] from aa have c0: "c \ 0" by auto from b' aa(3) show "degree a > 0" by simp from square_free_smult[OF c0 b'(1), folded aa(2)] show "square_free a" unfolding square_free_def by (force simp: dvd_def hom_distribs) show cnt: "content a = 1" and lc: "lead_coeff a > 0" using aa by auto fix A I assume A: "(A,I) \ set fs" and diff: "(a,i) \ (A,I)" from a[unfolded set_conv_nth] obtain k where k: "fs ! k = (a,i)" "k < length fs" by auto from A[unfolded set_conv_nth] obtain K where K: "fs ! K = (A,I)" "K < length fs" by auto from diff k K have kK: "k \ K" by auto from dist'[unfolded distinct_conv_nth length_map, rule_format, OF k(2) K(2) kK] have iI: "i \ I" using k K by simp from A out_rel obtain Bj where "Bj \ set Fs" and "rel_prod yun_erel (=) (A,I) Bj" unfolding list_all2_conv_all_nth set_conv_nth by fastforce then obtain B where B: "(B,I) \ set Fs" and AB: "yun_erel A B" by (cases Bj, auto simp: rel_prod.simps) then obtain C where Rel: "yun_rel A C B" unfolding yun_erel_def by auto note AA = yun_relD[OF this] from iI have "(b,i) \ (B,I)" by auto from sff(3)[OF b B this] have cop: "coprime b B" by simp from AA have C: "C \ 0" by auto from yun_rel_gcd[OF rel AA(1) C refl] obtain c where "yun_rel (gcd a A) c (gcd b B)" by auto note rel = yun_relD[OF this] from rel(2) cop have "?rp (gcd a A) = [: c :]" by simp from arg_cong[OF this, of degree] have "degree (gcd a A) = 0" by simp from degree0_coeffs[OF this] obtain c where gcd: "gcd a A = [: c :]" by auto from rel(8) rel(5) show "Rings.coprime a A" by (auto intro!: gcd_eq_1_imp_coprime simp add: gcd) } let ?prod = "\ fs. (\(a, i)\set fs. a ^ Suc i)" let ?pr = "\ fs. (\(a, i)\fs. a ^ Suc i)" define pr where "pr = ?prod fs" from \distinct fs\ have pfs: "?prod fs = ?pr fs" by (rule prod.distinct_set_conv_list) from \distinct Fs\ have pFs: "?prod Fs = ?pr Fs" by (rule prod.distinct_set_conv_list) from out_rel have "yun_erel (?prod fs) (?prod Fs)" unfolding pfs pFs proof (induct fs Fs rule: list_all2_induct) case (Cons ai fs Ai Fs) obtain a i where ai: "ai = (a,i)" by force from Cons(1) ai obtain A where Ai: "Ai = (A,i)" and rel: "yun_erel a A" by (cases Ai, auto simp: rel_prod.simps) show ?case unfolding ai Ai using yun_erel_mult[OF yun_erel_pow[OF rel, of "Suc i"] Cons(3)] by auto qed simp also have "?prod Fs = G" using sff(1) by simp finally obtain d where rel: "yun_rel pr d G" unfolding yun_erel_def pr_def by auto with in_rel have "f = pr" by (rule yun_rel_same_right) thus "f = smult 1 (?prod fs)" unfolding pr_def by simp qed from main dist' show ?thesis by auto next case (Some p) from res[unfolded square_free_factorization_int_main_def Some] have fs: "fs = [(f,0)]" by auto from lc have f0: "f \ 0" by auto from square_free_heuristic[OF Some] poly_mod_prime.separable_impl(1)[of p f] square_free_mod_imp_square_free[of p f] deg show ?thesis unfolding fs by (auto simp: ct lc square_free_factorization_def f0 poly_mod_prime_def) qed definition square_free_factorization_int' :: "int poly \ int \ (int poly \ nat)list" where "square_free_factorization_int' f = (if degree f = 0 then (lead_coeff f,[]) else (let \ \content factorization\ c = content f; d = (sgn (lead_coeff f) * c); g = sdiv_poly f d \ \and \square_free\ factorization\ in (d, square_free_factorization_int_main g)))" lemma square_free_factorization_int': assumes res: "square_free_factorization_int' f = (d, fs)" shows "square_free_factorization f (d,fs)" "(fi, i) \ set fs \ content fi = 1 \ lead_coeff fi > 0" "distinct (map snd fs)" proof - note res = res[unfolded square_free_factorization_int'_def Let_def] have "square_free_factorization f (d,fs) \ ((fi, i) \ set fs \ content fi = 1 \ lead_coeff fi > 0) \ distinct (map snd fs)" proof (cases "degree f = 0") case True from degree0_coeffs[OF True] obtain c where f: "f = [: c :]" by auto thus ?thesis using res by (simp add: square_free_factorization_def) next case False let ?s = "sgn (lead_coeff f)" have s: "?s \ {-1,1}" using False unfolding sgn_if by auto define g where "g = smult ?s f" let ?d = "?s * content f" have "content g = content ([:?s:] * f)" unfolding g_def by simp also have "\ = content [:?s:] * content f" unfolding content_mult by simp also have "content [:?s:] = 1" using s by (auto simp: content_def) finally have cg: "content g = content f" by simp from False res have d: "d = ?d" and fs: "fs = square_free_factorization_int_main (sdiv_poly f ?d)" by auto let ?g = "primitive_part g" define ng where "ng = primitive_part g" note fs also have "sdiv_poly f ?d = sdiv_poly g (content g)" unfolding cg unfolding g_def by (rule poly_eqI, unfold coeff_sdiv_poly coeff_smult, insert s, auto simp: div_minus_right) finally have fs: "square_free_factorization_int_main ng = fs" unfolding primitive_part_alt_def ng_def by simp have "lead_coeff f \ 0" using False by auto hence lg: "lead_coeff g > 0" unfolding g_def lead_coeff_smult by (meson linorder_neqE_linordered_idom sgn_greater sgn_less zero_less_mult_iff) hence g0: "g \ 0" by auto from g0 have "content g \ 0" by simp from arg_cong[OF content_times_primitive_part[of g], of lead_coeff, unfolded lead_coeff_smult] lg content_ge_0_int[of g] have lg': "lead_coeff ng > 0" unfolding ng_def by (metis \content g \ 0\ dual_order.antisym dual_order.strict_implies_order zero_less_mult_iff) from content_primitive_part[OF g0] have c_ng: "content ng = 1" unfolding ng_def . have "degree ng = degree f" using \content [:sgn (lead_coeff f):] = 1\ g_def ng_def by (auto simp add: sgn_eq_0_iff) with False have "degree ng \ 0" by auto note main = square_free_factorization_int_main[OF fs c_ng lg' this] show ?thesis proof (intro conjI impI) { assume "(fi, i) \ set fs" with main show "content fi = 1" "0 < lead_coeff fi" by auto } have d0: "d \ 0" using \content [:?s:] = 1\ d by (auto simp:sgn_eq_0_iff) have "smult d ng = smult ?s (smult (content g) (primitive_part g))" unfolding ng_def d cg by simp also have "smult (content g) (primitive_part g) = g" using content_times_primitive_part . also have "smult ?s g = f" unfolding g_def using s by auto finally have id: "smult d ng = f" . from main have "square_free_factorization ng (1, fs)" by auto from square_free_factorization_smult[OF d0 this] show "square_free_factorization f (d,fs)" unfolding id by simp show "distinct (map snd fs)" using main by auto qed qed thus "square_free_factorization f (d,fs)" "(fi, i) \ set fs \ content fi = 1 \ lead_coeff fi > 0" "distinct (map snd fs)" by auto qed definition x_split :: "'a :: semiring_0 poly \ nat \ 'a poly" where "x_split f = (let fs = coeffs f; zs = takeWhile ((=) 0) fs in case zs of [] \ (0,f) | _ \ (length zs, poly_of_list (dropWhile ((=) 0) fs)))" lemma x_split: assumes "x_split f = (n, g)" shows "f = monom 1 n * g" "n \ 0 \ f \ 0 \ \ monom 1 1 dvd g" proof - define zs where "zs = takeWhile ((=) 0) (coeffs f)" note res = assms[unfolded zs_def[symmetric] x_split_def Let_def] have "f = monom 1 n * g \ ((n \ 0 \ f \ 0) \ \ (monom 1 1 dvd g))" (is "_ \ (_ \ \ (?x dvd _))") proof (cases "f = 0") case True with res have "n = 0" "g = 0" unfolding zs_def by auto thus ?thesis using True by auto next case False note f = this show ?thesis proof (cases "zs = []") case True hence choice: "coeff f 0 \ 0" using f unfolding zs_def coeff_f_0_code poly_compare_0_code by (cases "coeffs f", auto) have dvd: "?x dvd h \ coeff h 0 = 0" for h by (simp add: monom_1_dvd_iff') from True choice res f show ?thesis unfolding dvd by auto next case False define ys where "ys = dropWhile ((=) 0) (coeffs f)" have dvd: "?x dvd h \ coeff h 0 = 0" for h by (simp add: monom_1_dvd_iff') from res False have n: "n = length zs" and g: "g = poly_of_list ys" unfolding ys_def by (cases zs, auto)+ obtain xx where xx: "coeffs f = xx" by auto have "coeffs f = zs @ ys" unfolding zs_def ys_def by auto also have "zs = replicate n 0" unfolding zs_def n xx by (induct xx, auto) finally have ff: "coeffs f = replicate n 0 @ ys" by auto from f have "lead_coeff f \ 0" by auto then have nz: "coeffs f \ []" "last (coeffs f) \ 0" by (simp_all add: last_coeffs_eq_coeff_degree) have ys: "ys \ []" using nz[unfolded ff] by auto with ys_def have hd: "hd ys \ 0" by (metis (full_types) hd_dropWhile) hence "coeff (poly_of_list ys) 0 \ 0" unfolding poly_of_list_def coeff_Poly using ys by (cases ys, auto) moreover have "coeffs (Poly ys) = ys" by (simp add: ys_def strip_while_dropWhile_commute) then have "coeffs (monom_mult n (Poly ys)) = replicate n 0 @ ys" by (simp add: coeffs_eq_iff monom_mult_def [symmetric] ff ys monom_mult_code) ultimately show ?thesis unfolding dvd g by (auto simp add: coeffs_eq_iff monom_mult_def [symmetric] ff) qed qed thus "f = monom 1 n * g" "n \ 0 \ f \ 0 \ \ monom 1 1 dvd g" by auto qed definition square_free_factorization_int :: "int poly \ int \ (int poly \ nat)list" where "square_free_factorization_int f = (case x_split f of (n,g) \ \extract \x^n\\ \ case square_free_factorization_int' g of (d,fs) \ if n = 0 then (d,fs) else (d, (monom 1 1, n - 1) # fs))" lemma square_free_factorization_int: assumes res: "square_free_factorization_int f = (d, fs)" shows "square_free_factorization f (d,fs)" "(fi, i) \ set fs \ primitive fi \ lead_coeff fi > 0" proof - obtain n g where xs: "x_split f = (n,g)" by force obtain c hs where sf: "square_free_factorization_int' g = (c,hs)" by force from res[unfolded square_free_factorization_int_def xs sf split] have d: "d = c" and fs: "fs = (if n = 0 then hs else (monom 1 1, n - 1) # hs)" by (cases n, auto) note sff = square_free_factorization_int'(1-2)[OF sf] note xs = x_split[OF xs] let ?x = "monom 1 1 :: int poly" have x: "primitive ?x \ lead_coeff ?x = 1 \ degree ?x = 1" by (auto simp add: degree_monom_eq content_def monom_Suc) thus "(fi, i) \ set fs \ primitive fi \ lead_coeff fi > 0" using sff(2) unfolding fs by (cases n, auto) show "square_free_factorization f (d,fs)" proof (cases n) case 0 with d fs sff xs show ?thesis by auto next case (Suc m) with xs have fg: "f = monom 1 (Suc m) * g" and dvd: "\ ?x dvd g" by auto from Suc have fs: "fs = (?x,m) # hs" unfolding fs by auto have degx: "degree ?x = 1" by code_simp from irreducible\<^sub>d_square_free[OF linear_irreducible\<^sub>d[OF this]] have sfx: "square_free ?x" by auto have fg: "f = ?x ^ n * g" unfolding fg Suc by (metis x_pow_n) have eq0: "?x ^ n * g = 0 \ g = 0" by simp note sf = square_free_factorizationD[OF sff(1)] { fix a i assume ai: "(a,i) \ set hs" with sf(4) have g0: "g \ 0" by auto from split_list[OF ai] obtain ys zs where hs: "hs = ys @ (a,i) # zs" by auto have "a dvd g" unfolding square_free_factorization_prod_list[OF sff(1)] hs by (rule dvd_smult, simp add: ac_simps) moreover have "\ ?x dvd g" using xs[unfolded Suc] by auto ultimately have dvd: "\ ?x dvd a" using dvd_trans by blast from sf(2)[OF ai] have "a \ 0" by auto have "1 = gcd ?x a" proof (rule gcdI) fix d assume d: "d dvd ?x" "d dvd a" from content_dvd_contentI[OF d(1)] x have cnt: "is_unit (content d)" by auto show "is_unit d" proof (cases "degree d = 1") case False with divides_degree[OF d(1), unfolded degx] have "degree d = 0" by auto from degree0_coeffs[OF this] obtain c where dc: "d = [:c:]" by auto from cnt[unfolded dc] have "is_unit c" by (auto simp: content_def, cases "c = 0", auto) hence "d * d = 1" unfolding dc by (cases "c = -1"; cases "c = 1", auto) thus "is_unit d" by (metis dvd_triv_right) next case True from d(1) obtain e where xde: "?x = d * e" unfolding dvd_def by auto from arg_cong[OF this, of degree] degx have "degree d + degree e = 1" by (metis True add.right_neutral degree_0 degree_mult_eq one_neq_zero) with True have "degree e = 0" by auto from degree0_coeffs[OF this] xde obtain e where xde: "?x = [:e:] * d" by auto from arg_cong[OF this, of content, unfolded content_mult] x have "content [:e:] * content d = 1" by auto also have "content [:e :] = abs e" by (auto simp: content_def, cases "e = 0", auto) finally have "\e\ * content d = 1" . from pos_zmult_eq_1_iff_lemma[OF this] have "e * e = 1" by (cases "e = 1"; cases "e = -1", auto) with arg_cong[OF xde, of "smult e"] have "d = ?x * [:e:]" by auto hence "?x dvd d" unfolding dvd_def by blast with d(2) have "?x dvd a" by (metis dvd_trans) with dvd show ?thesis by auto qed qed auto hence "coprime ?x a" by (simp add: gcd_eq_1_imp_coprime) note this dvd } note hs_dvd_x = this from hs_dvd_x[of ?x m] have nmem: "(?x,m) \ set hs" by auto hence eq: "?x ^ n * g = smult c (\(a, i)\set fs. a ^ Suc i)" unfolding sf(1) unfolding fs Suc by simp show ?thesis unfolding fg d unfolding square_free_factorization_def split eq0 unfolding eq proof (intro conjI allI impI, rule refl) fix a i assume ai: "(a,i) \ set fs" thus "square_free a" "degree a > 0" using sf(2) sfx degx unfolding fs by auto fix b j assume bj: "(b,j) \ set fs" and diff: "(a,i) \ (b,j)" consider (hs_hs) "(a,i) \ set hs" "(b,j) \ set hs" | (hs_x) "(a,i) \ set hs" "b = ?x" | (x_hs) "(b,j) \ set hs" "a = ?x" using ai bj diff unfolding fs by auto then show "Rings.coprime a b" proof cases case hs_hs from sf(3)[OF this diff] show ?thesis . next case hs_x from hs_dvd_x(1)[OF hs_x(1)] show ?thesis unfolding hs_x(2) by (simp add: ac_simps) next case x_hs from hs_dvd_x(1)[OF x_hs(1)] show ?thesis unfolding x_hs(2) by simp qed next show "g = 0 \ c = 0" using sf(4) by auto show "g = 0 \ fs = []" using sf(4) xs Suc by auto show "distinct fs" using sf(5) nmem unfolding fs by auto qed qed qed end diff --git a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy --- a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy +++ b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy @@ -1,962 +1,970 @@ theory Unique_Factorization imports Polynomial_Interpolation.Ring_Hom_Poly Polynomial_Factorization.Polynomial_Divisibility "HOL-Library.Permutations" "HOL-Computational_Algebra.Euclidean_Algorithm" Containers.Containers_Auxiliary (* only for a lemma *) Missing_Multiset2 "HOL-Algebra.Divisibility" begin hide_const(open) Divisibility.prime Divisibility.irreducible hide_fact(open) Divisibility.irreducible_def Divisibility.irreducibleI Divisibility.irreducibleD Divisibility.irreducibleE hide_const (open) Rings.coprime lemma irreducible_uminus [simp]: fixes a::"'a::idom" shows "irreducible (-a) \ irreducible a" using irreducible_mult_unit_left[of "-1::'a"] by auto context comm_monoid_mult begin definition coprime :: "'a \ 'a \ bool" where coprime_def': "coprime p q \ \r. r dvd p \ r dvd q \ r dvd 1" lemma coprimeI: assumes "\r. r dvd p \ r dvd q \ r dvd 1" shows "coprime p q" using assms by (auto simp: coprime_def') lemma coprimeE: assumes "coprime p q" and "(\r. r dvd p \ r dvd q \ r dvd 1) \ thesis" shows thesis using assms by (auto simp: coprime_def') lemma coprime_commute [ac_simps]: "coprime p q \ coprime q p" by (auto simp add: coprime_def') lemma not_coprime_iff_common_factor: "\ coprime p q \ (\r. r dvd p \ r dvd q \ \ r dvd 1)" by (auto simp add: coprime_def') end lemma (in algebraic_semidom) coprime_iff_coprime [simp, code]: "coprime = Rings.coprime" by (simp add: fun_eq_iff coprime_def coprime_def') lemma (in comm_semiring_1) coprime_0 [simp]: "coprime p 0 \ p dvd 1" "coprime 0 p \ p dvd 1" by (auto intro: coprimeI elim: coprimeE dest: dvd_trans) (**** until here ****) (* TODO: move or...? *) lemma dvd_rewrites: "dvd.dvd ((*)) = (dvd)" by (unfold dvd.dvd_def dvd_def, rule) subsection \Interfacing UFD properties\ hide_const (open) Divisibility.irreducible context comm_monoid_mult_isom begin lemma coprime_hom[simp]: "coprime (hom x) y' \ coprime x (Hilbert_Choice.inv hom y')" proof- show ?thesis by (unfold coprime_def', fold ball_UNIV, subst surj[symmetric], simp) qed lemma coprime_inv_hom[simp]: "coprime (Hilbert_Choice.inv hom x') y \ coprime x' (hom y)" proof- interpret inv: comm_monoid_mult_isom "Hilbert_Choice.inv hom".. show ?thesis by simp qed end subsubsection \Original part\ lemma dvd_dvd_imp_smult: fixes p q :: "'a :: idom poly" assumes pq: "p dvd q" and qp: "q dvd p" shows "\c. p = smult c q" proof (cases "p = 0") case True then show ?thesis by auto next case False from qp obtain r where r: "p = q * r" by (elim dvdE, auto) with False qp have r0: "r \ 0" and q0: "q \ 0" by auto with divides_degree[OF pq] divides_degree[OF qp] False have "degree p = degree q" by auto with r degree_mult_eq[OF q0 r0] have "degree r = 0" by auto from degree_0_id[OF this] obtain c where "r = [:c:]" by metis from r[unfolded this] show ?thesis by auto qed lemma dvd_const: assumes pq: "(p::'a::semidom poly) dvd q" and q0: "q \ 0" and degq: "degree q = 0" shows "degree p = 0" proof- from dvdE[OF pq] obtain r where *: "q = p * r". with q0 have "p \ 0" "r \ 0" by auto from degree_mult_eq[OF this] degq * show "degree p = 0" by auto qed context Rings.dvd begin abbreviation ddvd (infix "ddvd" 40) where "x ddvd y \ x dvd y \ y dvd x" lemma ddvd_sym[sym]: "x ddvd y \ y ddvd x" by auto end context comm_monoid_mult begin lemma ddvd_trans[trans]: "x ddvd y \ y ddvd z \ x ddvd z" using dvd_trans by auto lemma ddvd_transp: "transp (ddvd)" by (intro transpI, fact ddvd_trans) end context comm_semiring_1 begin definition mset_factors where "mset_factors F p \ F \ {#} \ (\f. f \# F \ irreducible f) \ p = prod_mset F" lemma mset_factorsI[intro!]: assumes "\f. f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" shows "mset_factors F p" unfolding mset_factors_def using assms by auto lemma mset_factorsD: assumes "mset_factors F p" shows "f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" using assms[unfolded mset_factors_def] by auto lemma mset_factorsE[elim]: assumes "mset_factors F p" and "(\f. f \# F \ irreducible f) \ F \ {#} \ prod_mset F = p \ thesis" shows thesis using assms[unfolded mset_factors_def] by auto lemma mset_factors_imp_not_is_unit: assumes "mset_factors F p" shows "\ p dvd 1" proof(cases F) case empty with assms show ?thesis by auto next case (add f F) with assms have "\ f dvd 1" "p = f * prod_mset F" by (auto intro!: irreducible_not_unit) then show ?thesis by auto qed definition primitive_poly where "primitive_poly f \ \d. (\i. d dvd coeff f i) \ d dvd 1" end lemma(in semidom) mset_factors_imp_nonzero: assumes "mset_factors F p" shows "p \ 0" proof assume "p = 0" moreover from assms have "prod_mset F = p" by auto ultimately obtain f where "f \# F" "f = 0" by auto with assms show False by auto qed class ufd = idom + assumes mset_factors_exist: "\x. x \ 0 \ \ x dvd 1 \ \F. mset_factors F x" and mset_factors_unique: "\x F G. mset_factors F x \ mset_factors G x \ rel_mset (ddvd) F G" subsubsection \Connecting to HOL/Divisibility\ context comm_semiring_1 begin abbreviation "mk_monoid \ \carrier = UNIV - {0}, mult = (*), one = 1\" lemma carrier_0[simp]: "x \ carrier mk_monoid \ x \ 0" by auto lemmas mk_monoid_simps = carrier_0 monoid.simps abbreviation irred where "irred \ Divisibility.irreducible mk_monoid" abbreviation factor where "factor \ Divisibility.factor mk_monoid" abbreviation factors where "factors \ Divisibility.factors mk_monoid" abbreviation properfactor where "properfactor \ Divisibility.properfactor mk_monoid" lemma factors: "factors fs y \ prod_list fs = y \ Ball (set fs) irred" proof - have "prod_list fs = foldr (*) fs 1" by (induct fs, auto) thus ?thesis unfolding factors_def by auto qed lemma factor: "factor x y \ (\z. z \ 0 \ x * z = y)" unfolding factor_def by auto lemma properfactor_nz: shows "(y :: 'a) \ 0 \ properfactor x y \ x dvd y \ \ y dvd x" by (auto simp: properfactor_def factor_def dvd_def) lemma mem_Units[simp]: "y \ Units mk_monoid \ y dvd 1" unfolding dvd_def Units_def by (auto simp: ac_simps) end context idom begin lemma irred_0[simp]: "irred (0::'a)" by (unfold Divisibility.irreducible_def, auto simp: factor properfactor_def) lemma factor_idom[simp]: "factor (x::'a) y \ (if y = 0 then x = 0 else x dvd y)" by (cases "y = 0"; auto intro: exI[of _ 1] elim: dvdE simp: factor) lemma associated_connect[simp]: "(\\<^bsub>mk_monoid\<^esub>) = (ddvd)" by (intro ext, unfold associated_def, auto) lemma essentially_equal_connect[simp]: "essentially_equal mk_monoid fs gs \ rel_mset (ddvd) (mset fs) (mset gs)" by (auto simp: essentially_equal_def rel_mset_via_perm) lemma irred_idom_nz: assumes x0: "(x::'a) \ 0" shows "irred x \ irreducible x" using x0 by (auto simp: irreducible_altdef Divisibility.irreducible_def properfactor_nz) lemma dvd_dvd_imp_unit_mult: assumes xy: "x dvd y" and yx: "y dvd x" shows "\z. z dvd 1 \ y = x * z" proof(cases "x = 0") case True with xy show ?thesis by (auto intro: exI[of _ 1]) next case x0: False from xy obtain z where z: "y = x * z" by (elim dvdE, auto) from yx obtain w where w: "x = y * w" by (elim dvdE, auto) from z w have "x * (z * w) = x" by (auto simp: ac_simps) then have "z * w = 1" using x0 by auto with z show ?thesis by (auto intro: exI[of _ z]) qed lemma irred_inner_nz: assumes x0: "x \ 0" shows "(\b. b dvd x \ \ x dvd b \ b dvd 1) \ (\a b. x = a * b \ a dvd 1 \ b dvd 1)" (is "?l \ ?r") proof (intro iffI allI impI) assume l: ?l fix a b assume xab: "x = a * b" then have ax: "a dvd x" and bx: "b dvd x" by auto { assume a1: "\ a dvd 1" with l ax have xa: "x dvd a" by auto from dvd_dvd_imp_unit_mult[OF ax xa] obtain z where z1: "z dvd 1" and xaz: "x = a * z" by auto from xab x0 have "a \ 0" by auto with xab xaz have "b = z" by auto with z1 have "b dvd 1" by auto } then show "a dvd 1 \ b dvd 1" by auto next assume r: ?r fix b assume bx: "b dvd x" and xb: "\ x dvd b" then obtain a where xab: "x = a * b" by (elim dvdE, auto simp: ac_simps) with r consider "a dvd 1" | "b dvd 1" by auto then show "b dvd 1" proof(cases) case 2 then show ?thesis by auto next case 1 then obtain c where ac1: "a * c = 1" by (elim dvdE, auto) from xab have "x * c = b * (a * c)" by (auto simp: ac_simps) with ac1 have "x * c = b" by auto then have "x dvd b" by auto with xb show ?thesis by auto qed qed lemma irred_idom[simp]: "irred x \ x = 0 \ irreducible x" by (cases "x = 0"; simp add: irred_idom_nz irred_inner_nz irreducible_def) lemma assumes "x \ 0" and "factors fs x" and "f \ set fs" shows "f \ 0" using assms by (auto simp: factors) lemma factors_as_mset_factors: assumes x0: "x \ 0" and x1: "x \ 1" shows "factors fs x \ mset_factors (mset fs) x" using assms by (auto simp: factors prod_mset_prod_list) end context ufd begin interpretation comm_monoid_cancel: comm_monoid_cancel "mk_monoid::'a monoid" apply (unfold_locales) apply simp_all using mult_left_cancel apply (auto simp: ac_simps) done lemma factors_exist: assumes "a \ 0" and "\ a dvd 1" shows "\fs. set fs \ UNIV - {0} \ factors fs a" proof- from mset_factors_exist[OF assms] obtain F where "mset_factors F a" by auto also from ex_mset obtain fs where "F = mset fs" by metis finally have fs: "mset_factors (mset fs) a". then have "factors fs a" using assms by (subst factors_as_mset_factors, auto) moreover have "set fs \ UNIV - {0}" using fs by (auto elim!: mset_factorsE) ultimately show ?thesis by auto qed lemma factors_unique: assumes fs: "factors fs a" and gs: "factors gs a" and a0: "a \ 0" and a1: "\ a dvd 1" shows "rel_mset (ddvd) (mset fs) (mset gs)" proof- from a1 have "a \ 1" by auto with a0 fs gs have "mset_factors (mset fs) a" "mset_factors (mset gs) a" by (unfold factors_as_mset_factors) from mset_factors_unique[OF this] show ?thesis. qed lemma factorial_monoid: "factorial_monoid (mk_monoid :: 'a monoid)" by (unfold_locales; auto simp add: factors_exist factors_unique) end lemma (in idom) factorial_monoid_imp_ufd: assumes "factorial_monoid (mk_monoid :: 'a monoid)" shows "class.ufd ((*) :: 'a \ _) 1 (+) 0 (-) uminus" proof (unfold_locales) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact assms) { fix x assume x: "x \ 0" "\ x dvd 1" note * = factors_exist[simplified, OF this] with x show "\F. mset_factors F x" by (subst(asm) factors_as_mset_factors, auto) } fix x F G assume FG: "mset_factors F x" "mset_factors G x" with mset_factors_imp_not_is_unit have x1: "\ x dvd 1" by auto from FG(1) have x0: "x \ 0" by (rule mset_factors_imp_nonzero) obtain fs gs where fsgs: "F = mset fs" "G = mset gs" using ex_mset by metis note FG = FG[unfolded this] then have 0: "0 \ set fs" "0 \ set gs" by (auto elim!: mset_factorsE) from x1 have "x \ 1" by auto note FG[folded factors_as_mset_factors[OF x0 this]] from factors_unique[OF this, simplified, OF x0 x1, folded fsgs] 0 show "rel_mset (ddvd) F G" by auto qed subsection \Preservation of Irreducibility\ locale comm_semiring_1_hom = comm_monoid_mult_hom hom + zero_hom hom for hom :: "'a :: comm_semiring_1 \ 'b :: comm_semiring_1" locale irreducibility_hom = comm_semiring_1_hom + assumes irreducible_imp_irreducible_hom: "irreducible a \ irreducible (hom a)" begin lemma hom_mset_factors: assumes F: "mset_factors F p" shows "mset_factors (image_mset hom F) (hom p)" proof (unfold mset_factors_def, intro conjI allI impI) from F show "hom p = prod_mset (image_mset hom F)" "image_mset hom F \ {#}" by (auto simp: hom_distribs) fix f' assume "f' \# image_mset hom F" then obtain f where f: "f \# F" and f'f: "f' = hom f" by auto with F irreducible_imp_irreducible_hom show "irreducible f'" unfolding f'f by auto qed end locale unit_preserving_hom = comm_semiring_1_hom + assumes is_unit_hom_if: "\x. hom x dvd 1 \ x dvd 1" begin lemma is_unit_hom_iff[simp]: "hom x dvd 1 \ x dvd 1" using is_unit_hom_if hom_dvd by force lemma irreducible_hom_imp_irreducible: assumes irr: "irreducible (hom a)" shows "irreducible a" proof (intro irreducibleI) from irr show "a \ 0" by auto from irr show "\ a dvd 1" by (auto dest: irreducible_not_unit) fix b c assume "a = b * c" then have "hom a = hom b * hom c" by (simp add: hom_distribs) with irr have "hom b dvd 1 \ hom c dvd 1" by (auto dest: irreducibleD) then show "b dvd 1 \ c dvd 1" by simp qed end locale factor_preserving_hom = unit_preserving_hom + irreducibility_hom begin lemma irreducible_hom[simp]: "irreducible (hom a) \ irreducible a" using irreducible_hom_imp_irreducible irreducible_imp_irreducible_hom by metis end lemma factor_preserving_hom_comp: assumes f: "factor_preserving_hom f" and g: "factor_preserving_hom g" shows "factor_preserving_hom (f o g)" proof- interpret f: factor_preserving_hom f by (rule f) interpret g: factor_preserving_hom g by (rule g) show ?thesis by (unfold_locales, auto simp: hom_distribs) qed context comm_semiring_isom begin sublocale unit_preserving_hom by (unfold_locales, auto) sublocale factor_preserving_hom proof (standard) fix a :: 'a assume "irreducible a" note a = this[unfolded irreducible_def] show "irreducible (hom a)" proof (rule ccontr) assume "\ irreducible (hom a)" from this[unfolded Factorial_Ring.irreducible_def,simplified] a obtain hb hc where eq: "hom a = hb * hc" and nu: "\ hb dvd 1" "\ hc dvd 1" by auto from bij obtain b where hb: "hb = hom b" by (elim bij_pointE) from bij obtain c where hc: "hc = hom c" by (elim bij_pointE) from eq[unfolded hb hc, folded hom_mult] have "a = b * c" by auto with nu hb hc have "a = b * c" "\ b dvd 1" "\ c dvd 1" by auto with a show False by auto qed qed end subsubsection\Back to divisibility\ lemma(in comm_semiring_1) mset_factors_mult: assumes F: "mset_factors F a" and G: "mset_factors G b" shows "mset_factors (F+G) (a*b)" proof(intro mset_factorsI) fix f assume "f \# F + G" then consider "f \# F" | "f \# G" by auto then show "irreducible f" by(cases, insert F G, auto) qed (insert F G, auto) lemma(in ufd) dvd_imp_subset_factors: assumes ab: "a dvd b" and F: "mset_factors F a" and G: "mset_factors G b" shows "\G'. G' \# G \ rel_mset (ddvd) F G'" proof- from F G have a0: "a \ 0" and b0: "b \ 0" by (simp_all add: mset_factors_imp_nonzero) from ab obtain c where c: "b = a * c" by (elim dvdE, auto) with b0 have c0: "c \ 0" by auto show ?thesis proof(cases "c dvd 1") case True show ?thesis proof(cases F) case empty with F show ?thesis by auto next case (add f F') with F have a: "f * prod_mset F' = a" and F': "\f. f \# F' \ irreducible f" and irrf: "irreducible f" by auto from irrf have f0: "f \ 0" and f1: "\f dvd 1" by (auto dest: irreducible_not_unit) from a c have "(f * c) * prod_mset F' = b" by (auto simp: ac_simps) moreover { have "irreducible (f * c)" using True irrf by (subst irreducible_mult_unit_right) with F' irrf have "\f'. f' \# F' + {#f * c#} \ irreducible f'" by auto } ultimately have "mset_factors (F' + {#f * c#}) b" by (intro mset_factorsI, auto) from mset_factors_unique[OF this G] have F'G: "rel_mset (ddvd) (F' + {#f * c#}) G". from True add have FF': "rel_mset (ddvd) F (F' + {#f * c#})" by (auto simp add: multiset.rel_refl intro!: rel_mset_Plus) have "rel_mset (ddvd) F G" apply(rule transpD[OF multiset.rel_transp[OF transpI] FF' F'G]) using ddvd_trans. then show ?thesis by auto qed next case False from mset_factors_exist[OF c0 this] obtain H where H: "mset_factors H c" by auto from c mset_factors_mult[OF F H] have "mset_factors (F + H) b" by auto note mset_factors_unique[OF this G] from rel_mset_split[OF this] obtain G1 G2 where "G = G1 + G2" "rel_mset (ddvd) F G1" "rel_mset (ddvd) H G2" by auto then show ?thesis by (intro exI[of _ "G1"], auto) qed qed lemma(in idom) irreducible_factor_singleton: assumes a: "irreducible a" shows "mset_factors F a \ F = {#a#}" proof(cases F) case empty with mset_factorsD show ?thesis by auto next case (add f F') show ?thesis proof assume F: "mset_factors F a" from add mset_factorsD[OF F] have *: "a = f * prod_mset F'" by auto then have fa: "f dvd a" by auto from * a have f0: "f \ 0" by auto from add have "f \# F" by auto with F have f: "irreducible f" by auto from add have "F' \# F" by auto then have unitemp: "prod_mset F' dvd 1 \ F' = {#}" proof(induct F') case empty then show ?case by auto next case (add f F') from add have "f \# F" by (simp add: mset_subset_eq_insertD) with F irreducible_not_unit have "\ f dvd 1" by auto then have "\ (prod_mset F' * f) dvd 1" by simp with add show ?case by auto qed show "F = {#a#}" proof(cases "a dvd f") case True then obtain r where "f = a * r" by (elim dvdE, auto) with * have "f = (r * prod_mset F') * f" by (auto simp: ac_simps) with f0 have "r * prod_mset F' = 1" by auto then have "prod_mset F' dvd 1" by (metis dvd_triv_right) with unitemp * add show ?thesis by auto next case False with fa a f show ?thesis by (auto simp: irreducible_altdef) qed qed (insert a, auto) qed lemma(in ufd) irreducible_dvd_imp_factor: assumes ab: "a dvd b" and a: "irreducible a" and G: "mset_factors G b" shows "\g \# G. a ddvd g" proof- from a have "mset_factors {#a#} a" by auto from dvd_imp_subset_factors[OF ab this G] obtain G' where G'G: "G' \# G" and rel: "rel_mset (ddvd) {#a#} G'" by auto with rel_mset_size size_1_singleton_mset size_single obtain g where gG': "G' = {#g#}" by fastforce from rel[unfolded this rel_mset_def] have "a ddvd g" by auto with gG' G'G show ?thesis by auto qed lemma(in idom) prod_mset_remove_units: "prod_mset F ddvd prod_mset {# f \# F. \f dvd 1 #}" proof(induct F) case (add f F) then show ?case by (cases "f = 0", auto) qed auto lemma(in comm_semiring_1) mset_factors_imp_dvd: assumes "mset_factors F x" and "f \# F" shows "f dvd x" using assms by (simp add: dvd_prod_mset mset_factors_def) lemma(in ufd) prime_elem_iff_irreducible[iff]: "prime_elem x \ irreducible x" proof (intro iffI, fact prime_elem_imp_irreducible, rule prime_elemI) assume r: "irreducible x" then show x0: "x \ 0" and x1: "\ x dvd 1" by (auto dest: irreducible_not_unit) from irreducible_factor_singleton[OF r] have *: "mset_factors {#x#} x" by auto fix a b assume "x dvd a * b" then obtain c where abxc: "a * b = x * c" by (elim dvdE, auto) show "x dvd a \ x dvd b" proof(cases "c = 0 \ a = 0 \ b = 0") case True with abxc show ?thesis by auto next case False then have a0: "a \ 0" and b0: "b \ 0" and c0: "c \ 0" by auto from x0 c0 have xc0: "x * c \ 0" by auto from x1 have xc1: "\ x * c dvd 1" by auto show ?thesis proof (cases "a dvd 1 \ b dvd 1") case False then have a1: "\ a dvd 1" and b1: "\ b dvd 1" by auto from mset_factors_exist[OF a0 a1] obtain F where Fa: "mset_factors F a" by auto then have F0: "F \ {#}" by auto from mset_factors_exist[OF b0 b1] obtain G where Gb: "mset_factors G b" by auto then have G0: "G \ {#}" by auto from mset_factors_mult[OF Fa Gb] have FGxc: "mset_factors (F + G) (x * c)" by (simp add: abxc) show ?thesis proof (cases "c dvd 1") case True from r irreducible_mult_unit_right[OF this] have "irreducible (x*c)" by simp note irreducible_factor_singleton[OF this] FGxc with F0 G0 have False by (cases F; cases G; auto) then show ?thesis by auto next case False from mset_factors_exist[OF c0 this] obtain H where "mset_factors H c" by auto with * have xHxc: "mset_factors (add_mset x H) (x * c)" by force note rel = mset_factors_unique[OF this FGxc] obtain hs where "mset hs = H" using ex_mset by auto then have "mset (x#hs) = add_mset x H" by auto from rel_mset_free[OF rel this] obtain jjs where jjsGH: "mset jjs = F + G" and rel: "list_all2 (ddvd) (x # hs) jjs" by auto then obtain j js where jjs: "jjs = j # js" by (cases jjs, auto) with rel have xj: "x ddvd j" by auto from jjs jjsGH have j: "j \ set_mset (F + G)" by (intro union_single_eq_member, auto) from j consider "j \# F" | "j \# G" by auto then show ?thesis proof(cases) case 1 with Fa have "j dvd a" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd a" by auto then show ?thesis by auto next case 2 with Gb have "j dvd b" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd b" by auto then show ?thesis by auto qed qed next case True then consider "a dvd 1" | "b dvd 1" by auto then show ?thesis proof(cases) case 1 then obtain d where ad: "a * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = a * d * b" by (auto simp: ac_simps) finally have "x dvd b" by (intro dvdI, auto simp: ad) then show ?thesis by auto next case 2 then obtain d where bd: "b * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = (b * d) * a" by (auto simp: ac_simps) finally have "x dvd a" by (intro dvdI, auto simp:bd) then show ?thesis by auto qed qed qed qed subsection\Results for GCDs etc.\ lemma prod_list_remove1: "(x :: 'b :: comm_monoid_mult) \ set xs \ prod_list (remove1 x xs) * x = prod_list xs" by (induct xs, auto simp: ac_simps) (* Isabelle 2015-style and generalized gcd-class without normalization and factors *) class comm_monoid_gcd = gcd + comm_semiring_1 + assumes gcd_dvd1[iff]: "gcd a b dvd a" and gcd_dvd2[iff]: "gcd a b dvd b" and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" begin lemma gcd_0_0[simp]: "gcd 0 0 = 0" using gcd_greatest[OF dvd_0_right dvd_0_right, of 0] by auto lemma gcd_zero_iff[simp]: "gcd a b = 0 \ a = 0 \ b = 0" proof assume "gcd a b = 0" from gcd_dvd1[of a b, unfolded this] gcd_dvd2[of a b, unfolded this] show "a = 0 \ b = 0" by auto qed auto lemma gcd_zero_iff'[simp]: "0 = gcd a b \ a = 0 \ b = 0" using gcd_zero_iff by metis lemma dvd_gcd_0_iff[simp]: shows "x dvd gcd 0 a \ x dvd a" (is ?g1) and "x dvd gcd a 0 \ x dvd a" (is ?g2) proof- have "a dvd gcd a 0" "a dvd gcd 0 a" by (auto intro: gcd_greatest) with dvd_refl show ?g1 ?g2 by (auto dest: dvd_trans) qed lemma gcd_dvd_1[simp]: "gcd a b dvd 1 \ coprime a b" using dvd_trans[OF gcd_greatest[of _ a b], of _ 1] by (cases "a = 0 \ b = 0") (auto intro!: coprimeI elim: coprimeE) lemma dvd_imp_gcd_dvd_gcd: "b dvd c \ gcd a b dvd gcd a c" by (meson gcd_dvd1 gcd_dvd2 gcd_greatest dvd_trans) definition listgcd :: "'a list \ 'a" where "listgcd xs = foldr gcd xs 0" lemma listgcd_simps[simp]: "listgcd [] = 0" "listgcd (x # xs) = gcd x (listgcd xs)" by (auto simp: listgcd_def) lemma listgcd: "x \ set xs \ listgcd xs dvd x" proof (induct xs) case (Cons y ys) show ?case proof (cases "x = y") case False with Cons have dvd: "listgcd ys dvd x" by auto thus ?thesis unfolding listgcd_simps using dvd_trans by blast next case True thus ?thesis unfolding listgcd_simps using dvd_trans by blast qed qed simp lemma listgcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd listgcd xs" by (induct xs arbitrary:y, auto intro: gcd_greatest) end context Rings.dvd begin definition "is_gcd x a b \ x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x)" definition "some_gcd a b \ SOME x. is_gcd x a b" lemma is_gcdI[intro!]: assumes "x dvd a" "x dvd b" "\y. y dvd a \ y dvd b \ y dvd x" shows "is_gcd x a b" by (insert assms, auto simp: is_gcd_def) lemma is_gcdE[elim!]: assumes "is_gcd x a b" and "x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x) \ thesis" shows thesis by (insert assms, auto simp: is_gcd_def) lemma is_gcd_some_gcdI: assumes "\x. is_gcd x a b" shows "is_gcd (some_gcd a b) a b" by (unfold some_gcd_def, rule someI_ex[OF assms]) end context comm_semiring_1 begin lemma some_gcd_0[intro!]: "is_gcd (some_gcd a 0) a 0" "is_gcd (some_gcd 0 b) 0 b" by (auto intro!: is_gcd_some_gcdI intro: exI[of _ a] exI[of _ b]) lemma some_gcd_0_dvd[intro!]: "some_gcd a 0 dvd a" "some_gcd 0 b dvd b" using some_gcd_0 by auto lemma dvd_some_gcd_0[intro!]: "a dvd some_gcd a 0" "b dvd some_gcd 0 b" using some_gcd_0[of a] some_gcd_0[of b] by auto end context idom begin lemma is_gcd_connect: assumes "a \ 0" "b \ 0" shows "isgcd mk_monoid x a b \ is_gcd x a b" using assms by (force simp: isgcd_def) lemma some_gcd_connect: assumes "a \ 0" and "b \ 0" shows "somegcd mk_monoid a b = some_gcd a b" using assms by (auto intro!: arg_cong[of _ _ Eps] simp: is_gcd_connect some_gcd_def somegcd_def) end context comm_monoid_gcd begin lemma is_gcd_gcd: "is_gcd (gcd a b) a b" using gcd_greatest by auto lemma is_gcd_some_gcd: "is_gcd (some_gcd a b) a b" by (insert is_gcd_gcd, auto intro!: is_gcd_some_gcdI) lemma gcd_dvd_some_gcd: "gcd a b dvd some_gcd a b" using is_gcd_some_gcd by auto lemma some_gcd_dvd_gcd: "some_gcd a b dvd gcd a b" using is_gcd_some_gcd by (auto intro: gcd_greatest) lemma some_gcd_ddvd_gcd: "some_gcd a b ddvd gcd a b" by (auto intro: gcd_dvd_some_gcd some_gcd_dvd_gcd) lemma some_gcd_dvd: "some_gcd a b dvd d \ gcd a b dvd d" "d dvd some_gcd a b \ d dvd gcd a b" using some_gcd_ddvd_gcd[of a b] by (auto dest:dvd_trans) end class idom_gcd = comm_monoid_gcd + idom begin interpretation raw: comm_monoid_cancel "mk_monoid :: 'a monoid" by (unfold_locales, auto intro: mult_commute mult_assoc) interpretation raw: gcd_condition_monoid "mk_monoid :: 'a monoid" by (unfold_locales, auto simp: is_gcd_connect intro!: exI[of _ "gcd _ _"] dest: gcd_greatest) lemma gcd_mult_ddvd: "d * gcd a b ddvd gcd (d * a) (d * b)" proof (cases "d = 0") case True then show ?thesis by auto next case d0: False show ?thesis proof (cases "a = 0 \ b = 0") case False note some_gcd_ddvd_gcd[of a b] with d0 have "d * gcd a b ddvd d * some_gcd a b" by auto also have "d * some_gcd a b ddvd some_gcd (d * a) (d * b)" using False d0 raw.gcd_mult by (simp add: some_gcd_connect) also note some_gcd_ddvd_gcd finally show ?thesis. next case True with d0 show ?thesis apply (elim disjE) apply (rule ddvd_trans[of _ "d * b"]; force) apply (rule ddvd_trans[of _ "d * a"]; force) done qed qed lemma gcd_greatest_mult: assumes cad: "c dvd a * d" and cbd: "c dvd b * d" shows "c dvd gcd a b * d" proof- from gcd_greatest[OF assms] have c: "c dvd gcd (d * a) (d * b)" by (auto simp: ac_simps) note gcd_mult_ddvd[of d a b] then have "gcd (d * a) (d * b) dvd gcd a b * d" by (auto simp: ac_simps) from dvd_trans[OF c this] show ?thesis . qed lemma listgcd_greatest_mult: "(\ x :: 'a. x \ set xs \ y dvd x * z) \ y dvd listgcd xs * z" proof (induct xs) case (Cons x xs) from Cons have "y dvd x * z" "y dvd listgcd xs * z" by auto thus ?case unfolding listgcd_simps by (rule gcd_greatest_mult) qed (simp) lemma dvd_factor_mult_gcd: assumes dvd: "k dvd p * q" "k dvd p * r" and q0: "q \ 0" and r0: "r \ 0" shows "k dvd p * gcd q r" proof - from dvd gcd_greatest[of k "p * q" "p * r"] have "k dvd gcd (p * q) (p * r)" by simp also from gcd_mult_ddvd[of p q r] have "... dvd (p * gcd q r)" by auto finally show ?thesis . qed lemma coprime_mult_cross_dvd: assumes coprime: "coprime p q" and eq: "p' * p = q' * q" shows "p dvd q'" (is ?g1) and "q dvd p'" (is ?g2) proof (atomize(full), cases "p = 0 \ q = 0") case True then show "?g1 \ ?g2" proof assume p0: "p = 0" with coprime have "q dvd 1" by auto with eq p0 show ?thesis by auto next assume q0: "q = 0" with coprime have "p dvd 1" by auto with eq q0 show ?thesis by auto qed next case False { fix p q r p' q' :: 'a assume cop: "coprime p q" and eq: "p' * p = q' * q" and p: "p \ 0" and q: "q \ 0" and r: "r dvd p" "r dvd q" let ?gcd = "gcd q p" from eq have "p' * p dvd q' * q" by auto hence d1: "p dvd q' * q" by (rule dvd_mult_right) have d2: "p dvd q' * p" by auto from dvd_factor_mult_gcd[OF d1 d2 q p] have 1: "p dvd q' * ?gcd" . from q p have 2: "?gcd dvd q" by auto from q p have 3: "?gcd dvd p" by auto from cop[unfolded coprime_def', rule_format, OF 3 2] have "?gcd dvd 1" . from 1 dvd_mult_unit_iff[OF this] have "p dvd q'" by auto } note main = this from main[OF coprime eq,of 1] False coprime coprime_commute main[OF _ eq[symmetric], of 1] show "?g1 \ ?g2" by auto qed end subclass (in ring_gcd) idom_gcd by (unfold_locales, auto) lemma coprime_rewrites: "comm_monoid_mult.coprime ((*)) 1 = coprime" apply (intro ext) apply (subst comm_monoid_mult.coprime_def') apply (unfold_locales) apply (unfold dvd_rewrites) apply (fold coprime_def') .. (* TODO: incorporate into the default class hierarchy *) locale gcd_condition = fixes ty :: "'a :: idom itself" assumes gcd_exists: "\a b :: 'a. \x. is_gcd x a b" begin sublocale idom_gcd "(*)" "1 :: 'a" "(+)" 0 "(-)" uminus some_gcd rewrites "dvd.dvd ((*)) = (dvd)" and "comm_monoid_mult.coprime ((*) ) 1 = Unique_Factorization.coprime" proof- have "is_gcd (some_gcd a b) a b" for a b :: 'a by (intro is_gcd_some_gcdI gcd_exists) from this[unfolded is_gcd_def] show "class.idom_gcd (*) (1 :: 'a) (+) 0 (-) uminus some_gcd" by (unfold_locales, auto simp: dvd_rewrites) qed (simp_all add: dvd_rewrites coprime_rewrites) end instance semiring_gcd \ comm_monoid_gcd by (intro_classes, auto) lemma listgcd_connect: "listgcd = gcd_list" proof (intro ext) fix xs :: "'a list" show "listgcd xs = gcd_list xs" by(induct xs, auto) qed interpretation some_gcd: gcd_condition "TYPE('a::ufd)" proof(unfold_locales, intro exI) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact factorial_monoid) note d = dvd.dvd_def some_gcd_def carrier_0 fix a b :: 'a show "is_gcd (some_gcd a b) a b" proof (cases "a = 0 \ b = 0") case True thus ?thesis using some_gcd_0 by auto next case False with gcdof_exists[of a b] show ?thesis by (auto intro!: is_gcd_some_gcdI simp add: is_gcd_connect some_gcd_connect) qed qed lemma some_gcd_listgcd_dvd_listgcd: "some_gcd.listgcd xs dvd listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) lemma listgcd_dvd_some_gcd_listgcd: "listgcd xs dvd some_gcd.listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) context factorial_ring_gcd begin text \Do not declare the following as subclass, to avoid conflict in \field \ gcd_condition\ vs. \factorial_ring_gcd \ gcd_condition\. \ sublocale as_ufd: ufd proof(unfold_locales, goal_cases) case (1 x) from prime_factorization_exists[OF \x \ 0\] - obtain F where f: "\f. f \# F \ prime_elem f" and Fx: "prod_mset F = normalize x" by auto - from Fx have x: "x = unit_factor x * prod_mset F" by auto + obtain F where f: "\f. f \# F \ prime_elem f" + and Fx: "normalize (prod_mset F) = normalize x" by auto + from associatedE2[OF Fx] obtain u where u: "is_unit u" "x = u * prod_mset F" + by blast from \\ is_unit x\ Fx have "F \ {#}" by auto then obtain g G where F: "F = add_mset g G" by (cases F, auto) then have "g \# F" by auto with f[OF this]prime_elem_iff_irreducible irreducible_mult_unit_left[OF unit_factor_is_unit[OF \x \ 0\]] - have g: "irreducible (unit_factor x * g)" by auto + have g: "irreducible (u * g)" using u(1) + by (subst irreducible_mult_unit_left) simp_all show ?case proof (intro exI conjI mset_factorsI) - from x show "prod_mset (add_mset (unit_factor x * g) G) = x" by (simp add: F ac_simps) - fix f assume "f \# add_mset (unit_factor x * g) G" + show "prod_mset (add_mset (u * g) G) = x" + using \x \ 0\ by (simp add: F ac_simps u) + fix f assume "f \# add_mset (u * g) G" with f[unfolded F] g prime_elem_iff_irreducible show "irreducible f" by auto qed auto next case (2 x F G) note transpD[OF multiset.rel_transp[OF ddvd_transp],trans] obtain fs where F: "F = mset fs" by (metis ex_mset) have "list_all2 (ddvd) fs (map normalize fs)" by (intro list_all2_all_nthI, auto) then have FH: "rel_mset (ddvd) F (image_mset normalize F)" by (unfold rel_mset_def F, force) also - have FG: "image_mset normalize F = image_mset normalize G" - proof (intro prime_factorization_unique') - from 2 have xF: "x = prod_mset F" and xG: "x = prod_mset G" by auto - from xF have "normalize x = prod_mset (image_mset normalize F)" by (simp add: local.normalize_prod_mset) - with xG have nFG: "\ = prod_mset (image_mset normalize G)" by (simp_all add: local.normalize_prod_mset) - then show "(\i\#image_mset normalize F. i) = (\i\#image_mset normalize G. i)" by auto - from 2 prime_elem_iff_irreducible have "f \# F \ prime_elem f" "g \# G \ prime_elem g" for f g - by (auto intro: prime_elemI) - then show " Multiset.Ball (image_mset normalize F) prime" - "Multiset.Ball (image_mset normalize G) prime" by auto - qed + have FG: "image_mset normalize F = image_mset normalize G" + proof (intro prime_factorization_unique'') + from 2 have xF: "x = prod_mset F" and xG: "x = prod_mset G" by auto + from xF have "normalize x = normalize (prod_mset (image_mset normalize F))" + by (simp add: normalize_prod_mset_normalize) + with xG have nFG: "\ = normalize (prod_mset (image_mset normalize G))" + by (simp_all add: normalize_prod_mset_normalize) + then show "normalize (\i\#image_mset normalize F. i) = + normalize (\i\#image_mset normalize G. i)" by auto + next + from 2 prime_elem_iff_irreducible have "f \# F \ prime_elem f" "g \# G \ prime_elem g" for f g + by (auto intro: prime_elemI) + then show " Multiset.Ball (image_mset normalize F) prime" + "Multiset.Ball (image_mset normalize G) prime" by auto + qed also obtain gs where G: "G = mset gs" by (metis ex_mset) have "list_all2 ((ddvd)\\) gs (map normalize gs)" by (intro list_all2_all_nthI, auto) then have "rel_mset (ddvd) (image_mset normalize G) G" by (subst multiset.rel_flip[symmetric], unfold rel_mset_def G, force) finally show ?case. qed end instance int :: ufd by (intro class.ufd.of_class.intro as_ufd.ufd_axioms) instance int :: idom_gcd by (intro_classes, auto) instance field \ ufd by (intro_classes, auto simp: dvd_field_iff) end diff --git a/thys/Berlekamp_Zassenhaus/Unique_Factorization_Poly.thy b/thys/Berlekamp_Zassenhaus/Unique_Factorization_Poly.thy --- a/thys/Berlekamp_Zassenhaus/Unique_Factorization_Poly.thy +++ b/thys/Berlekamp_Zassenhaus/Unique_Factorization_Poly.thy @@ -1,899 +1,911 @@ (* Author: RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Unique Factorization Domain for Polynomials\ text \In this theory we prove that the polynomials over a unique factorization domain (UFD) form a UFD.\ theory Unique_Factorization_Poly imports Unique_Factorization Polynomial_Factorization.Missing_Polynomial_Factorial Subresultants.More_Homomorphisms "HOL-Computational_Algebra.Field_as_Ring" begin hide_const (open) module.smult hide_const (open) Divisibility.irreducible instantiation fract :: (idom) "{normalization_euclidean_semiring, euclidean_ring}" begin definition [simp]: "normalize_fract \ (normalize_field :: 'a fract \ _)" definition [simp]: "unit_factor_fract = (unit_factor_field :: 'a fract \ _)" definition [simp]: "euclidean_size_fract = (euclidean_size_field :: 'a fract \ _)" definition [simp]: "modulo_fract = (mod_field :: 'a fract \ _)" instance by standard (simp_all add: dvd_field_iff divide_simps) end instantiation fract :: (idom) euclidean_ring_gcd begin definition gcd_fract :: "'a fract \ 'a fract \ 'a fract" where "gcd_fract \ Euclidean_Algorithm.gcd" definition lcm_fract :: "'a fract \ 'a fract \ 'a fract" where "lcm_fract \ Euclidean_Algorithm.lcm" definition Gcd_fract :: "'a fract set \ 'a fract" where "Gcd_fract \ Euclidean_Algorithm.Gcd" definition Lcm_fract :: "'a fract set \ 'a fract" where "Lcm_fract \ Euclidean_Algorithm.Lcm" instance by (standard, simp_all add: gcd_fract_def lcm_fract_def Gcd_fract_def Lcm_fract_def) end +(*field + unique_euclidean_ring + euclidean_ring_gcd + normalization_semidom_multiplicative*) + +instantiation fract :: (idom) unique_euclidean_ring +begin + +definition [simp]: "division_segment_fract (x :: 'a fract) = (1 :: 'a fract)" + +instance by standard (auto split: if_splits) +end + +instance fract :: (idom) field_gcd by standard auto + definition divides_ff :: "'a::idom fract \ 'a fract \ bool" where "divides_ff x y \ \ r. y = x * to_fract r" lemma ff_list_pairs: "\ xs. X = map (\ (x,y). Fraction_Field.Fract x y) xs \ 0 \ snd ` set xs" proof (induct X) case (Cons a X) from Cons(1) obtain xs where X: "X = map (\ (x,y). Fraction_Field.Fract x y) xs" and xs: "0 \ snd ` set xs" by auto obtain x y where a: "a = Fraction_Field.Fract x y" and y: "y \ 0" by (cases a, auto) show ?case unfolding X a using xs y by (intro exI[of _ "(x,y) # xs"], auto) qed auto lemma divides_ff_to_fract[simp]: "divides_ff (to_fract x) (to_fract y) \ x dvd y" unfolding divides_ff_def dvd_def by (simp add: to_fract_def eq_fract(1) mult.commute) lemma shows divides_ff_mult_cancel_left[simp]: "divides_ff (z * x) (z * y) \ z = 0 \ divides_ff x y" and divides_ff_mult_cancel_right[simp]: "divides_ff (x * z) (y * z) \ z = 0 \ divides_ff x y" unfolding divides_ff_def by auto definition gcd_ff_list :: "'a::ufd fract list \ 'a fract \ bool" where "gcd_ff_list X g = ( (\ x \ set X. divides_ff g x) \ (\ d. (\ x \ set X. divides_ff d x) \ divides_ff d g))" lemma gcd_ff_list_exists: "\ g. gcd_ff_list (X :: 'a::ufd fract list) g" proof - interpret some_gcd: idom_gcd "(*)" "1 :: 'a" "(+)" 0 "(-)" uminus some_gcd rewrites "dvd.dvd ((*)) = (dvd)" by (unfold_locales, auto simp: dvd_rewrites) from ff_list_pairs[of X] obtain xs where X: "X = map (\ (x,y). Fraction_Field.Fract x y) xs" and xs: "0 \ snd ` set xs" by auto define r where "r \ prod_list (map snd xs)" have r: "r \ 0" unfolding r_def prod_list_zero_iff using xs by auto define ys where "ys \ map (\ (x,y). x * prod_list (remove1 y (map snd xs))) xs" { fix i assume "i < length X" hence i: "i < length xs" unfolding X by auto obtain x y where xsi: "xs ! i = (x,y)" by force with i have "(x,y) \ set xs" unfolding set_conv_nth by force hence y_mem: "y \ set (map snd xs)" by force with xs have y: "y \ 0" by force from i have id1: "ys ! i = x * prod_list (remove1 y (map snd xs))" unfolding ys_def using xsi by auto from i xsi have id2: "X ! i = Fraction_Field.Fract x y" unfolding X by auto have lp: "prod_list (remove1 y (map snd xs)) * y = r" unfolding r_def by (rule prod_list_remove1[OF y_mem]) have "ys ! i \ set ys" using i unfolding ys_def by auto moreover have "to_fract (ys ! i) = to_fract r * (X ! i)" unfolding id1 id2 to_fract_def mult_fract by (subst eq_fract(1), force, force simp: y, simp add: lp) ultimately have "ys ! i \ set ys" "to_fract (ys ! i) = to_fract r * (X ! i)" . } note ys = this define G where "G \ some_gcd.listgcd ys" define g where "g \ to_fract G * Fraction_Field.Fract 1 r" have len: "length X = length ys" unfolding X ys_def by auto show ?thesis proof (rule exI[of _ g], unfold gcd_ff_list_def, intro ballI conjI impI allI) fix x assume "x \ set X" then obtain i where i: "i < length X" and x: "x = X ! i" unfolding set_conv_nth by auto from ys[OF i] have id: "to_fract (ys ! i) = to_fract r * x" and ysi: "ys ! i \ set ys" unfolding x by auto from some_gcd.listgcd[OF ysi] have "G dvd ys ! i" unfolding G_def . then obtain d where ysi: "ys ! i = G * d" unfolding dvd_def by auto have "to_fract d * (to_fract G * Fraction_Field.Fract 1 r) = x * (to_fract r * Fraction_Field.Fract 1 r)" using id[unfolded ysi] by (simp add: ac_simps) also have "\ = x" using r unfolding to_fract_def by (simp add: eq_fract One_fract_def) finally have "to_fract d * (to_fract G * Fraction_Field.Fract 1 r) = x" by simp thus "divides_ff g x" unfolding divides_ff_def g_def by (intro exI[of _ d], auto) next fix d assume "\x \ set X. divides_ff d x" hence "Ball ((\ x. to_fract r * x) ` set X) ( divides_ff (to_fract r * d))" by simp also have "(\ x. to_fract r * x) ` set X = to_fract ` set ys" unfolding set_conv_nth using ys len by force finally have dvd: "Ball (set ys) (\ y. divides_ff (to_fract r * d) (to_fract y))" by auto obtain nd dd where d: "d = Fraction_Field.Fract nd dd" and dd: "dd \ 0" by (cases d, auto) { fix y assume "y \ set ys" hence "divides_ff (to_fract r * d) (to_fract y)" using dvd by auto from this[unfolded divides_ff_def d to_fract_def mult_fract] obtain ra where "Fraction_Field.Fract y 1 = Fraction_Field.Fract (r * nd * ra) dd" by auto hence "y * dd = ra * (r * nd)" by (simp add: eq_fract dd) hence "r * nd dvd y * dd" by auto } hence "r * nd dvd some_gcd.listgcd ys * dd" by (rule some_gcd.listgcd_greatest_mult) hence "divides_ff (to_fract r * d) (to_fract G)" unfolding to_fract_def d mult_fract G_def divides_ff_def by (auto simp add: eq_fract dd dvd_def) also have "to_fract G = to_fract r * g" unfolding g_def using r by (auto simp: to_fract_def eq_fract) finally show "divides_ff d g" using r by simp qed qed definition some_gcd_ff_list :: "'a :: ufd fract list \ 'a fract" where "some_gcd_ff_list xs = (SOME g. gcd_ff_list xs g)" lemma some_gcd_ff_list: "gcd_ff_list xs (some_gcd_ff_list xs)" unfolding some_gcd_ff_list_def using gcd_ff_list_exists[of xs] by (rule someI_ex) lemma some_gcd_ff_list_divides: "x \ set xs \ divides_ff (some_gcd_ff_list xs) x" using some_gcd_ff_list[of xs] unfolding gcd_ff_list_def by auto lemma some_gcd_ff_list_greatest: "(\x \ set xs. divides_ff d x) \ divides_ff d (some_gcd_ff_list xs)" using some_gcd_ff_list[of xs] unfolding gcd_ff_list_def by auto lemma divides_ff_refl[simp]: "divides_ff x x" unfolding divides_ff_def by (rule exI[of _ 1], auto simp: to_fract_def One_fract_def) lemma divides_ff_trans: "divides_ff x y \ divides_ff y z \ divides_ff x z" unfolding divides_ff_def by (auto simp del: to_fract_hom.hom_mult simp add: to_fract_hom.hom_mult[symmetric]) lemma divides_ff_mult_right: "a \ 0 \ divides_ff (x * inverse a) y \ divides_ff x (a * y)" unfolding divides_ff_def divide_inverse[symmetric] by auto definition eq_dff :: "'a :: ufd fract \ 'a fract \ bool" (infix "=dff" 50) where "x =dff y \ divides_ff x y \ divides_ff y x" lemma eq_dffI[intro]: "divides_ff x y \ divides_ff y x \ x =dff y" unfolding eq_dff_def by auto lemma eq_dff_refl[simp]: "x =dff x" by (intro eq_dffI, auto) lemma eq_dff_sym: "x =dff y \ y =dff x" unfolding eq_dff_def by auto lemma eq_dff_trans[trans]: "x =dff y \ y =dff z \ x =dff z" unfolding eq_dff_def using divides_ff_trans by auto lemma eq_dff_cancel_right[simp]: "x * y =dff x * z \ x = 0 \ y =dff z" unfolding eq_dff_def by auto lemma eq_dff_mult_right_trans[trans]: "x =dff y * z \ z =dff u \ x =dff y * u" using eq_dff_trans by force lemma some_gcd_ff_list_smult: "a \ 0 \ some_gcd_ff_list (map ((*) a) xs) =dff a * some_gcd_ff_list xs" proof let ?g = "some_gcd_ff_list (map ((*) a) xs)" show "divides_ff (a * some_gcd_ff_list xs) ?g" by (rule some_gcd_ff_list_greatest, insert some_gcd_ff_list_divides[of _ xs], auto simp: divides_ff_def) assume a: "a \ 0" show "divides_ff ?g (a * some_gcd_ff_list xs)" proof (rule divides_ff_mult_right[OF a some_gcd_ff_list_greatest], intro ballI) fix x assume x: "x \ set xs" have "divides_ff (?g * inverse a) x = divides_ff (inverse a * ?g) (inverse a * (a * x))" using a by (simp add: field_simps) also have "\" using a x by (auto intro: some_gcd_ff_list_divides) finally show "divides_ff (?g * inverse a) x" . qed qed definition content_ff :: "'a::ufd fract poly \ 'a fract" where "content_ff p = some_gcd_ff_list (coeffs p)" lemma content_ff_iff: "divides_ff x (content_ff p) \ (\ c \ set (coeffs p). divides_ff x c)" (is "?l = ?r") proof assume ?l from divides_ff_trans[OF this, unfolded content_ff_def, OF some_gcd_ff_list_divides] show ?r .. next assume ?r thus ?l unfolding content_ff_def by (intro some_gcd_ff_list_greatest, auto) qed lemma content_ff_divides_ff: "x \ set (coeffs p) \ divides_ff (content_ff p) x" unfolding content_ff_def by (rule some_gcd_ff_list_divides) lemma content_ff_0[simp]: "content_ff 0 = 0" using content_ff_iff[of 0 0] by (auto simp: divides_ff_def) lemma content_ff_0_iff[simp]: "(content_ff p = 0) = (p = 0)" proof (cases "p = 0") case False define a where "a \ last (coeffs p)" define xs where "xs \ coeffs p" from False have mem: "a \ set (coeffs p)" and a: "a \ 0" unfolding a_def last_coeffs_eq_coeff_degree[OF False] coeffs_def by auto from content_ff_divides_ff[OF mem] have "divides_ff (content_ff p) a" . with a have "content_ff p \ 0" unfolding divides_ff_def by auto with False show ?thesis by auto qed auto lemma content_ff_eq_dff_nonzero: "content_ff p =dff x \ x \ 0 \ p \ 0" using divides_ff_def eq_dff_def by force lemma content_ff_smult: "content_ff (smult (a::'a::ufd fract) p) =dff a * content_ff p" proof (cases "a = 0") case False note a = this have id: "coeffs (smult a p) = map ((*) a) (coeffs p)" unfolding coeffs_smult using a by (simp add: Polynomial.coeffs_smult) show ?thesis unfolding content_ff_def id using some_gcd_ff_list_smult[OF a] . qed simp definition normalize_content_ff where "normalize_content_ff (p::'a::ufd fract poly) \ smult (inverse (content_ff p)) p" lemma smult_normalize_content_ff: "smult (content_ff p) (normalize_content_ff p) = p" unfolding normalize_content_ff_def by (cases "p = 0", auto) lemma content_ff_normalize_content_ff_1: assumes p0: "p \ 0" shows "content_ff (normalize_content_ff p) =dff 1" proof - have "content_ff p = content_ff (smult (content_ff p) (normalize_content_ff p))" unfolding smult_normalize_content_ff .. also have "\ =dff content_ff p * content_ff (normalize_content_ff p)" by (rule content_ff_smult) finally show ?thesis unfolding eq_dff_def divides_ff_def using p0 by auto qed lemma content_ff_to_fract: assumes "set (coeffs p) \ range to_fract" shows "content_ff p \ range to_fract" proof - have "divides_ff 1 (content_ff p)" using assms unfolding content_ff_iff unfolding divides_ff_def[abs_def] by auto thus ?thesis unfolding divides_ff_def by auto qed lemma content_ff_map_poly_to_fract: "content_ff (map_poly to_fract (p :: 'a :: ufd poly)) \ range to_fract" by (rule content_ff_to_fract, subst coeffs_map_poly, auto) lemma range_coeffs_to_fract: assumes "set (coeffs p) \ range to_fract" shows "\ m. coeff p i = to_fract m" proof - from assms(1) to_fract_0 have "coeff p i \ range to_fract" using range_coeff [of p] by auto (metis contra_subsetD to_fract_hom.hom_zero insertE range_eqI) thus ?thesis by auto qed lemma divides_ff_coeff: assumes "set (coeffs p) \ range to_fract" and "divides_ff (to_fract n) (coeff p i)" shows "\ m. coeff p i = to_fract n * to_fract m" proof - from range_coeffs_to_fract[OF assms(1)] obtain k where pi: "coeff p i = to_fract k" by auto from assms(2)[unfolded this] have "n dvd k" by simp then obtain j where k: "k = n * j" unfolding Rings.dvd_def by auto show ?thesis unfolding pi k by auto qed definition inv_embed :: "'a :: ufd fract \ 'a" where "inv_embed = the_inv to_fract" lemma inv_embed[simp]: "inv_embed (to_fract x) = x" unfolding inv_embed_def by (rule the_inv_f_f, auto simp: inj_on_def) lemma inv_embed_0[simp]: "inv_embed 0 = 0" unfolding to_fract_0[symmetric] inv_embed by simp lemma range_to_fract_embed_poly: assumes "set (coeffs p) \ range to_fract" shows "p = map_poly to_fract (map_poly inv_embed p)" proof - have "p = map_poly (to_fract o inv_embed) p" by (rule sym, rule map_poly_idI, insert assms, auto) also have "\ = map_poly to_fract (map_poly inv_embed p)" by (subst map_poly_map_poly, auto) finally show ?thesis . qed lemma content_ff_to_fract_coeffs_to_fract: assumes "content_ff p \ range to_fract" shows "set (coeffs p) \ range to_fract" proof fix x assume "x \ set (coeffs p)" from content_ff_divides_ff[OF this] assms[unfolded eq_dff_def] show "x \ range to_fract" unfolding divides_ff_def by (auto simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) qed lemma content_ff_1_coeffs_to_fract: assumes "content_ff p =dff 1" shows "set (coeffs p) \ range to_fract" proof fix x assume "x \ set (coeffs p)" from content_ff_divides_ff[OF this] assms[unfolded eq_dff_def] show "x \ range to_fract" unfolding divides_ff_def by (auto simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) qed lemma gauss_lemma: fixes p q :: "'a :: ufd fract poly" shows "content_ff (p * q) =dff content_ff p * content_ff q" proof (cases "p = 0 \ q = 0") case False hence p: "p \ 0" and q: "q \ 0" by auto let ?c = "content_ff :: 'a fract poly \ 'a fract" { fix p q :: "'a fract poly" assume cp1: "?c p =dff 1" and cq1: "?c q =dff 1" define ip where "ip \ map_poly inv_embed p" define iq where "iq \ map_poly inv_embed q" interpret map_poly_hom: map_poly_comm_ring_hom to_fract.. from content_ff_1_coeffs_to_fract[OF cp1] have cp: "set (coeffs p) \ range to_fract" . from content_ff_1_coeffs_to_fract[OF cq1] have cq: "set (coeffs q) \ range to_fract" . have ip: "p = map_poly to_fract ip" unfolding ip_def by (rule range_to_fract_embed_poly[OF cp]) have iq: "q = map_poly to_fract iq" unfolding iq_def by (rule range_to_fract_embed_poly[OF cq]) have cpq0: "?c (p * q) \ 0" unfolding content_ff_0_iff using cp1 cq1 content_ff_eq_dff_nonzero[of _ 1] by auto have cpq: "set (coeffs (p * q)) \ range to_fract" unfolding ip iq unfolding map_poly_hom.hom_mult[symmetric] to_fract_hom.coeffs_map_poly_hom by auto have ctnt: "?c (p * q) \ range to_fract" using content_ff_to_fract[OF cpq] . then obtain cpq where id: "?c (p * q) = to_fract cpq" by auto have dvd: "divides_ff 1 (?c (p * q))" using ctnt unfolding divides_ff_def by auto from cpq0[unfolded id] have cpq0: "cpq \ 0" unfolding to_fract_def Zero_fract_def by auto hence cpqM: "cpq \ carrier mk_monoid" by auto have "?c (p * q) =dff 1" proof (rule ccontr) assume "\ ?c (p * q) =dff 1" with dvd have "\ divides_ff (?c (p * q)) 1" unfolding eq_dff_def by auto from this[unfolded id divides_ff_def] have cpq: "\ r. cpq * r \ 1" by (auto simp: to_fract_def One_fract_def eq_fract) then have cpq1: "\ cpq dvd 1" by (auto elim:dvdE simp:ac_simps) from mset_factors_exist[OF cpq0 cpq1] obtain F where F: "mset_factors F cpq" by auto have "F \ {#}" using F by auto then obtain f where f: "f \# F" by auto with F have irrf: "irreducible f" and f0: "f \ 0" by (auto dest: mset_factorsD) from irrf have pf: "prime_elem f" by simp note * = this[unfolded prime_elem_def] from * have no_unit: "\ f dvd 1" by auto from * f0 have prime: "\ a b. f dvd a * b \ f dvd a \ f dvd b" unfolding dvd_def by force let ?f = "to_fract f" from F f have fdvd: "f dvd cpq" by (auto intro:mset_factors_imp_dvd) hence "divides_ff ?f (to_fract cpq)" by simp from divides_ff_trans[OF this, folded id, OF content_ff_divides_ff] have dvd: "\ z. z \ set (coeffs (p * q)) \ divides_ff ?f z" . { fix p :: "'a fract poly" assume cp: "?c p =dff 1" let ?P = "\ i. \ divides_ff ?f (coeff p i)" { assume "\ c \ set (coeffs p). divides_ff ?f c" hence n: "divides_ff ?f (?c p)" unfolding content_ff_iff by auto from divides_ff_trans[OF this] cp[unfolded eq_dff_def] have "divides_ff ?f 1" by auto also have "1 = to_fract 1" by simp finally have "f dvd 1" by (unfold divides_ff_to_fract) hence False using no_unit unfolding dvd_def by (auto simp: ac_simps) } then obtain cp where cp: "cp \ set (coeffs p)" and ncp: "\ divides_ff ?f cp" by auto hence "cp \ range (coeff p)" unfolding range_coeff by auto with ncp have "\ i. ?P i" by auto from LeastI_ex[OF this] not_less_Least[of _ ?P] have "\ i. ?P i \ (\ j. j < i \ divides_ff ?f (coeff p j))" by blast } note cont = this from cont[OF cp1] obtain r where r: "\ divides_ff ?f (coeff p r)" and r': "\ i. i < r \ divides_ff ?f (coeff p i)" by auto have "\ i. \ k. i < r \ coeff p i = ?f * to_fract k" using divides_ff_coeff[OF cp r'] by blast from choice[OF this] obtain rr where r': "\ i. i < r \ coeff p i = ?f * to_fract (rr i)" by blast let ?r = "coeff p r" from cont[OF cq1] obtain s where s: "\ divides_ff ?f (coeff q s)" and s': "\ i. i < s \ divides_ff ?f (coeff q i)" by auto have "\ i. \ k. i < s \ coeff q i = ?f * to_fract k" using divides_ff_coeff[OF cq s'] by blast from choice[OF this] obtain ss where s': "\ i. i < s \ coeff q i = ?f * to_fract (ss i)" by blast from range_coeffs_to_fract[OF cp] have "\ i. \ m. coeff p i = to_fract m" .. from choice[OF this] obtain pi where pi: "\ i. coeff p i = to_fract (pi i)" by blast from range_coeffs_to_fract[OF cq] have "\ i. \ m. coeff q i = to_fract m" .. from choice[OF this] obtain qi where qi: "\ i. coeff q i = to_fract (qi i)" by blast let ?s = "coeff q s" let ?g = "\ i. coeff p i * coeff q (r + s - i)" define a where "a = (\i\{.. i \ {Suc r..r + s}. pi i * (ss (r + s - i)))" have "coeff (p * q) (r + s) = (\i\r + s. ?g i)" unfolding coeff_mult .. also have "{..r+s} = {..< r} \ {r .. r+s}" by auto also have "(\i\{.. {r..r + s}. ?g i) = (\i\{.. i \ {r..r + s}. ?g i)" by (rule sum.union_disjoint, auto) also have "(\i\{..i\{.. = to_fract (f * a)" by (simp add: a_def sum_distrib_left) also have "(\ i \ {r..r + s}. ?g i) = ?g r + (\ i \ {Suc r..r + s}. ?g i)" by (subst sum.remove[of _ r], auto intro: sum.cong) also have "(\ i \ {Suc r..r + s}. ?g i) = (\ i \ {Suc r..r + s}. ?f * (to_fract (pi i) * to_fract (ss (r + s - i))))" by (rule sum.cong[OF refl], insert s' pi, auto) also have "\ = to_fract (f * b)" by (simp add: sum_distrib_left b_def) finally have cpq: "coeff (p * q) (r + s) = to_fract (f * (a + b)) + ?r * ?s" by (simp add: field_simps) { fix i from dvd[of "coeff (p * q) i"] have "divides_ff ?f (coeff (p * q) i)" using range_coeff[of "p * q"] by (cases "coeff (p * q) i = 0", auto simp: divides_ff_def) } from this[of "r + s", unfolded cpq] have "divides_ff ?f (to_fract (f * (a + b) + pi r * qi s))" unfolding pi qi by simp from this[unfolded divides_ff_to_fract] have "f dvd pi r * qi s" by (metis dvd_add_times_triv_left_iff mult.commute) from prime[OF this] have "f dvd pi r \ f dvd qi s" by auto with r s show False unfolding pi qi by auto qed } note main = this define n where "n \ normalize_content_ff :: 'a fract poly \ 'a fract poly" let ?s = "\ p. smult (content_ff p) (n p)" have "?c (p * q) = ?c (?s p * ?s q)" unfolding smult_normalize_content_ff n_def by simp also have "?s p * ?s q = smult (?c p * ?c q) (n p * n q)" by (simp add: mult.commute) also have "?c (\) =dff (?c p * ?c q) * ?c (n p * n q)" by (rule content_ff_smult) also have "?c (n p * n q) =dff 1" unfolding n_def by (rule main, insert p q, auto simp: content_ff_normalize_content_ff_1) finally show ?thesis by simp qed auto abbreviation (input) "content_ff_ff p \ content_ff (map_poly to_fract p)" lemma factorization_to_fract: assumes q: "q \ 0" and factor: "map_poly to_fract (p :: 'a :: ufd poly) = q * r" shows "\ q' r' c. c \ 0 \ q = smult c (map_poly to_fract q') \ r = smult (inverse c) (map_poly to_fract r') \ content_ff_ff q' =dff 1 \ p = q' * r'" proof - let ?c = content_ff let ?p = "map_poly to_fract p" interpret map_poly_inj_comm_ring_hom "to_fract :: 'a \ _".. define cq where "cq \ normalize_content_ff q" define cr where "cr \ smult (content_ff q) r" define q' where "q' \ map_poly inv_embed cq" define r' where "r' \ map_poly inv_embed cr" from content_ff_map_poly_to_fract have cp_ff: "?c ?p \ range to_fract" by auto from smult_normalize_content_ff[of q] have cqs: "q = smult (content_ff q) cq" unfolding cq_def .. from content_ff_normalize_content_ff_1[OF q] have c_cq: "content_ff cq =dff 1" unfolding cq_def . from content_ff_1_coeffs_to_fract[OF this] have cq_ff: "set (coeffs cq) \ range to_fract" . have factor: "?p = cq * cr" unfolding factor cr_def using cqs by (metis mult_smult_left mult_smult_right) from gauss_lemma[of cq cr] have cp: "?c ?p =dff ?c cq * ?c cr" unfolding factor . with c_cq have "?c ?p =dff ?c cr" by (metis eq_dff_mult_right_trans mult.commute mult.right_neutral) with cp_ff have "?c cr \ range to_fract" by (metis divides_ff_def to_fract_hom.hom_mult eq_dff_def image_iff range_eqI) from content_ff_to_fract_coeffs_to_fract[OF this] have cr_ff: "set (coeffs cr) \ range to_fract" by auto have cq: "cq = map_poly to_fract q'" unfolding q'_def by (rule range_to_fract_embed_poly[OF cq_ff]) have cr: "cr = map_poly to_fract r'" unfolding r'_def by (rule range_to_fract_embed_poly[OF cr_ff]) from factor[unfolded cq cr] have p: "p = q' * r'" by (simp add: injectivity) from c_cq have ctnt: "content_ff_ff q' =dff 1" using cq q'_def by force from cqs have idq: "q = smult (?c q) (map_poly to_fract q')" unfolding cq . with q have cq: "?c q \ 0" by auto have "r = smult (inverse (?c q)) cr" unfolding cr_def using cq by auto also have "cr = map_poly to_fract r'" by (rule cr) finally have idr: "r = smult (inverse (?c q)) (map_poly to_fract r')" by auto from cq p ctnt idq idr show ?thesis by blast qed lemma irreducible_PM_M_PFM: assumes irr: "irreducible p" shows "degree p = 0 \ irreducible (coeff p 0) \ degree p \ 0 \ irreducible (map_poly to_fract p) \ content_ff_ff p =dff 1" proof- interpret map_poly_inj_idom_hom to_fract.. from irr[unfolded irreducible_altdef] have p0: "p \ 0" and irr: "\ p dvd 1" "\ b. b dvd p \ \ p dvd b \ b dvd 1" by auto show ?thesis proof (cases "degree p = 0") case True from degree0_coeffs[OF True] obtain a where p: "p = [:a:]" by auto note irr = irr[unfolded p] from p p0 have a0: "a \ 0" by auto moreover have "\ a dvd 1" using irr(1) by simp moreover { fix b assume "b dvd a" "\ a dvd b" hence "[:b:] dvd [:a:]" "\ [:a:] dvd [:b:]" unfolding const_poly_dvd . from irr(2)[OF this] have "b dvd 1" unfolding const_poly_dvd_1 . } ultimately have "irreducible a" unfolding irreducible_altdef by auto with True show ?thesis unfolding p by auto next case False let ?E = "map_poly to_fract" let ?p = "?E p" have dp: "degree ?p \ 0" using False by simp from p0 have p': "?p \ 0" by simp moreover have "\ ?p dvd 1" proof assume "?p dvd 1" then obtain q where id: "?p * q = 1" unfolding dvd_def by auto have deg: "degree (?p * q) = degree ?p + degree q" by (rule degree_mult_eq, insert id, auto) from arg_cong[OF id, of degree, unfolded deg] dp show False by auto qed moreover { fix q assume "q dvd ?p" and ndvd: "\ ?p dvd q" then obtain r where fact: "?p = q * r" unfolding dvd_def by auto with p' have q0: "q \ 0" by auto from factorization_to_fract[OF this fact] obtain q' r' c where *: "c \ 0" "q = smult c (?E q')" "r = smult (inverse c) (?E r')" "content_ff_ff q' =dff 1" "p = q' * r'" by auto hence "q' dvd p" unfolding dvd_def by auto note irr = irr(2)[OF this] have "\ p dvd q'" proof assume "p dvd q'" then obtain u where q': "q' = p * u" unfolding dvd_def by auto from arg_cong[OF this, of "\ x. smult c (?E x)", unfolded *(2)[symmetric]] have "q = ?p * smult c (?E u)" by simp hence "?p dvd q" unfolding dvd_def by blast with ndvd show False .. qed from irr[OF this] have "q' dvd 1" . from divides_degree[OF this] have "degree q' = 0" by auto from degree0_coeffs[OF this] obtain a' where "q' = [:a':]" by auto from *(2)[unfolded this] obtain a where q: "q = [:a:]" by (simp add: to_fract_hom.map_poly_pCons_hom) with q0 have a: "a \ 0" by auto have "q dvd 1" unfolding q const_poly_dvd_1 using a unfolding dvd_def by (intro exI[of _ "inverse a"], auto) } ultimately have irr_p': "irreducible ?p" unfolding irreducible_altdef by auto let ?c = "content_ff" have "?c ?p \ range to_fract" by (rule content_ff_to_fract, unfold to_fract_hom.coeffs_map_poly_hom, auto) then obtain c where cp: "?c ?p = to_fract c" by auto from p' cp have c: "c \ 0" by auto have "?c ?p =dff 1" unfolding cp proof (rule ccontr) define cp where "cp = normalize_content_ff ?p" from smult_normalize_content_ff[of ?p] have cps: "?p = smult (to_fract c) cp" unfolding cp_def cp .. from content_ff_normalize_content_ff_1[OF p'] have c_cp: "content_ff cp =dff 1" unfolding cp_def . from range_to_fract_embed_poly[OF content_ff_1_coeffs_to_fract[OF c_cp]] obtain cp' where "cp = ?E cp'" by auto from cps[unfolded this] have "p = smult c cp'" by (simp add: injectivity) hence dvd: "[: c :] dvd p" unfolding dvd_def by auto have "\ p dvd [: c :]" using divides_degree[of p "[: c :]"] c False by auto from irr(2)[OF dvd this] have "c dvd 1" by simp assume "\ to_fract c =dff 1" from this[unfolded eq_dff_def One_fract_def to_fract_def[symmetric] divides_ff_def to_fract_mult] have c1: "\ r. 1 \ c * r" by (auto simp: ac_simps simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) with \c dvd 1\ show False unfolding dvd_def by blast qed with False irr_p' show ?thesis by auto qed qed lemma irreducible_M_PM: fixes p :: "'a :: ufd poly" assumes 0: "degree p = 0" and irr: "irreducible (coeff p 0)" shows "irreducible p" proof (cases "p = 0") case True thus ?thesis using assms by auto next case False from degree0_coeffs[OF 0] obtain a where p: "p = [:a:]" by auto with False have a0: "a \ 0" by auto from p irr have "irreducible a" by auto from this[unfolded irreducible_altdef] have a1: "\ a dvd 1" and irr: "\ b. b dvd a \ \ a dvd b \ b dvd 1" by auto { fix b assume *: "b dvd [:a:]" "\ [:a:] dvd b" from divides_degree[OF this(1)] a0 have "degree b = 0" by auto from degree0_coeffs[OF this] obtain bb where b: "b = [: bb :]" by auto from * irr[of bb] have "b dvd 1" unfolding b const_poly_dvd by auto } with a0 a1 show ?thesis by (auto simp: irreducible_altdef p) qed lemma primitive_irreducible_imp_degree: "primitive (p::'a::{semiring_gcd,idom} poly) \ irreducible p \ degree p > 0" by (unfold irreducible_primitive_connect[symmetric], auto) lemma irreducible_degree_field: fixes p :: "'a :: field poly" assumes "irreducible p" shows "degree p > 0" proof- { assume "degree p = 0" from degree0_coeffs[OF this] assms obtain a where p: "p = [:a:]" and a: "a \ 0" by auto hence "1 = p * [:inverse a:]" by auto hence "p dvd 1" .. hence "p \ Units mk_monoid" by simp with assms have False unfolding irreducible_def by auto } then show ?thesis by auto qed lemma irreducible_PFM_PM: assumes irr: "irreducible (map_poly to_fract p)" and ct: "content_ff_ff p =dff 1" shows "irreducible p" proof - let ?E = "map_poly to_fract" let ?p = "?E p" from ct have p0: "p \ 0" by (auto simp: eq_dff_def divides_ff_def) moreover from irreducible_degree_field[OF irr] have deg: "degree p \ 0" by simp from irr[unfolded irreducible_altdef] have irr: "\ b. b dvd ?p \ \ ?p dvd b \ b dvd 1" by auto have "\ p dvd 1" using deg divides_degree[of p 1] by auto moreover { fix q :: "'a poly" assume dvd: "q dvd p" and ndvd: "\ p dvd q" from dvd obtain r where pqr: "p = q * r" .. from arg_cong[OF this, of ?E] have pqr': "?p = ?E q * ?E r" by simp from p0 pqr have q: "q \ 0" and r: "r \ 0" by auto have dp: "degree p = degree q + degree r" unfolding pqr by (subst degree_mult_eq, insert q r, auto) from eq_dff_trans[OF eq_dff_sym[OF gauss_lemma[of "?E q" "?E r", folded pqr']] ct] have ct: "content_ff (?E q) * content_ff (?E r) =dff 1" . from content_ff_map_poly_to_fract obtain cq where cq: "content_ff (?E q) = to_fract cq" by auto from content_ff_map_poly_to_fract obtain cr where cr: "content_ff (?E r) = to_fract cr" by auto note ct[unfolded cq cr to_fract_mult eq_dff_def divides_ff_def] from this[folded hom_distribs] obtain c where c: "cq * cr * c = 1" by (auto simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) hence one: "1 = cq * (c * cr)" "1 = cr * (c * cq)" by (auto simp: ac_simps) { assume *: "degree q \ 0 \ degree r \ 0" with dp have "degree q < degree p" by auto hence "degree (?E q) < degree (?E p)" by simp hence ndvd: "\ ?p dvd ?E q" using divides_degree[of ?p "?E q"] q by auto have "?E q dvd ?p" unfolding pqr' by auto from irr[OF this ndvd] have "?E q dvd 1" . from divides_degree[OF this] * have False by auto } hence "degree q = 0 \ degree r = 0" by blast then have "q dvd 1" proof assume "degree q = 0" from degree0_coeffs[OF this] q obtain a where q: "q = [:a:]" and a: "a \ 0" by auto hence id: "set (coeffs (?E q)) = {to_fract a}" by auto have "divides_ff (to_fract a) (content_ff (?E q))" unfolding content_ff_iff id by auto from this[unfolded cq divides_ff_def, folded hom_distribs] obtain rr where cq: "cq = a * rr" by (auto simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) with one(1) have "1 = a * (rr * c * cr)" by (auto simp: ac_simps) hence "a dvd 1" .. thus ?thesis by (simp add: q) next assume "degree r = 0" from degree0_coeffs[OF this] r obtain a where r: "r = [:a:]" and a: "a \ 0" by auto hence id: "set (coeffs (?E r)) = {to_fract a}" by auto have "divides_ff (to_fract a) (content_ff (?E r))" unfolding content_ff_iff id by auto note this[unfolded cr divides_ff_def to_fract_mult] note this[folded hom_distribs] then obtain rr where cr: "cr = a * rr" by (auto simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) with one(2) have one: "1 = a * (rr * c * cq)" by (auto simp: ac_simps) from arg_cong[OF pqr[unfolded r], of "\ p. p * [:rr * c * cq:]"] have "p * [:rr * c * cq:] = q * [:a * (rr * c * cq):]" by (simp add: ac_simps) also have "\ = q" unfolding one[symmetric] by auto finally obtain r where "q = p * r" by blast hence "p dvd q" .. with ndvd show ?thesis by auto qed } ultimately show ?thesis by (auto simp:irreducible_altdef) qed lemma irreducible_cases: "irreducible p \ degree p = 0 \ irreducible (coeff p 0) \ degree p \ 0 \ irreducible (map_poly to_fract p) \ content_ff_ff p =dff 1" using irreducible_PM_M_PFM irreducible_M_PM irreducible_PFM_PM by blast lemma dvd_PM_iff: "p dvd q \ divides_ff (content_ff_ff p) (content_ff_ff q) \ map_poly to_fract p dvd map_poly to_fract q" proof - interpret map_poly_inj_idom_hom to_fract.. let ?E = "map_poly to_fract" show ?thesis (is "?l = ?r") proof assume "p dvd q" then obtain r where qpr: "q = p * r" .. from arg_cong[OF this, of ?E] have dvd: "?E p dvd ?E q" by auto from content_ff_map_poly_to_fract obtain cq where cq: "content_ff_ff q = to_fract cq" by auto from content_ff_map_poly_to_fract obtain cp where cp: "content_ff_ff p = to_fract cp" by auto from content_ff_map_poly_to_fract obtain cr where cr: "content_ff_ff r = to_fract cr" by auto from gauss_lemma[of "?E p" "?E r", folded hom_distribs qpr, unfolded cq cp cr] have "divides_ff (content_ff_ff p) (content_ff_ff q)" unfolding cq cp eq_dff_def by (metis divides_ff_def divides_ff_trans) with dvd show ?r by blast next assume ?r show ?l proof (cases "q = 0") case True with \?r\ show ?l by auto next case False note q = this hence q': "?E q \ 0" by auto from \?r\ obtain rr where qpr: "?E q = ?E p * rr" unfolding dvd_def by auto with q have p: "p \ 0" and Ep: "?E p \ 0" and rr: "rr \ 0" by auto from gauss_lemma[of "?E p" rr, folded qpr] have ct: "content_ff_ff q =dff content_ff_ff p * content_ff rr" by auto from content_ff_map_poly_to_fract[of p] obtain cp where cp: "content_ff_ff p = to_fract cp" by auto from content_ff_map_poly_to_fract[of q] obtain cq where cq: "content_ff_ff q = to_fract cq" by auto from \?r\[unfolded cp cq] have "divides_ff (to_fract cp) (to_fract cq)" .. with ct[unfolded cp cq eq_dff_def] have "content_ff rr \ range to_fract" by (metis (no_types, lifting) Ep content_ff_0_iff cp divides_ff_def divides_ff_trans mult.commute mult_right_cancel range_eqI) from range_to_fract_embed_poly[OF content_ff_to_fract_coeffs_to_fract[OF this]] obtain r where rr: "rr = ?E r" by auto from qpr[unfolded rr, folded hom_distribs] have "q = p * r" by (rule injectivity) thus "p dvd q" .. qed qed qed lemma factorial_monoid_poly: "factorial_monoid (mk_monoid :: 'a :: ufd poly monoid)" proof (fold factorial_condition_one, intro conjI) interpret M: factorial_monoid "mk_monoid :: 'a monoid" by (fact factorial_monoid) interpret PFM: factorial_monoid "mk_monoid :: 'a fract poly monoid" by (rule as_ufd.factorial_monoid) interpret PM: comm_monoid_cancel "mk_monoid :: 'a poly monoid" by (unfold_locales, auto) let ?E = "map_poly to_fract" show "divisor_chain_condition_monoid (mk_monoid::'a poly monoid)" proof (unfold_locales, unfold mk_monoid_simps) let ?rel' = "{(x::'a poly, y). x \ 0 \ y \ 0 \ properfactor x y}" let ?rel'' = "{(x::'a, y). x \ 0 \ y \ 0 \ properfactor x y}" let ?relPM = "{(x, y). x \ 0 \ y \ 0 \ x dvd y \ \ y dvd (x :: 'a poly)}" let ?relM = "{(x, y). x \ 0 \ y \ 0 \ x dvd y \ \ y dvd (x :: 'a)}" have id: "?rel' = ?relPM" using properfactor_nz by auto have id': "?rel'' = ?relM" using properfactor_nz by auto have "wf ?rel''" using M.division_wellfounded by auto hence wfM: "wf ?relM" using id' by auto let ?c = "\ p. inv_embed (content_ff_ff p)" let ?f = "\ p. (degree p, ?c p)" note wf = wf_inv_image[OF wf_lex_prod[OF wf_less wfM], of ?f] show "wf ?rel'" unfolding id proof (rule wf_subset[OF wf], clarify) fix p q :: "'a poly" assume p: "p \ 0" and q: "q \ 0" and dvd: "p dvd q" and ndvd: "\ q dvd p" from dvd obtain r where qpr: "q = p * r" .. from degree_mult_eq[of p r, folded qpr] q qpr have r: "r \ 0" and deg: "degree q = degree p + degree r" by auto show "(p,q) \ inv_image ({(x, y). x < y} <*lex*> ?relM) ?f" proof (cases "degree p = degree q") case False with deg have "degree p < degree q" by auto thus ?thesis by auto next case True with deg have "degree r = 0" by simp from degree0_coeffs[OF this] r obtain a where ra: "r = [:a:]" and a: "a \ 0" by auto from arg_cong[OF qpr, of "\ p. ?E p * [:inverse (to_fract a):]"] a have "?E p = ?E q * [:inverse (to_fract a):]" by (auto simp: ac_simps ra) hence "?E q dvd ?E p" .. with ndvd dvd_PM_iff have ndvd: "\ divides_ff (content_ff_ff q) (content_ff_ff p)" by auto from content_ff_map_poly_to_fract obtain cq where cq: "content_ff_ff q = to_fract cq" by auto from content_ff_map_poly_to_fract obtain cp where cp: "content_ff_ff p = to_fract cp" by auto from ndvd[unfolded cp cq] have ndvd: "\ cq dvd cp" by simp from iffD1[OF dvd_PM_iff,OF dvd,unfolded cq cp] have dvd: "cp dvd cq" by simp have c_p: "?c p = cp" unfolding cp by simp have c_q: "?c q = cq" unfolding cq by simp from q cq have cq0: "cq \ 0" by auto from p cp have cp0: "cp \ 0" by auto from ndvd cq0 cp0 dvd have "(?c p, ?c q) \ ?relM" unfolding c_p c_q by auto with True show ?thesis by auto qed qed qed show "primeness_condition_monoid (mk_monoid::'a poly monoid)" proof (unfold_locales, unfold mk_monoid_simps) fix p :: "'a poly" assume p: "p \ 0" and "irred p" then have irr: "irreducible p" by auto from p have p': "?E p \ 0" by auto from irreducible_PM_M_PFM[OF irr] have choice: "degree p = 0 \ irred (coeff p 0) \ degree p \ 0 \ irred (?E p) \ content_ff_ff p =dff 1" by auto show "Divisibility.prime mk_monoid p" proof (rule Divisibility.primeI, unfold mk_monoid_simps mem_Units) show "\ p dvd 1" proof assume "p dvd 1" from divides_degree[OF this] have dp: "degree p = 0" by auto from degree0_coeffs[OF this] p obtain a where p: "p = [:a:]" and a: "a \ 0" by auto with choice have irr: "irreducible a" by auto from \p dvd 1\[unfolded p] have "a dvd 1" by auto with irr show False unfolding irreducible_def by auto qed fix q r :: "'a poly" assume q: "q \ 0" and r: "r \ 0" and "factor p (q * r)" from this[unfolded factor_idom] have "p dvd q * r" by auto from iffD1[OF dvd_PM_iff this] have dvd_ct: "divides_ff (content_ff_ff p) (content_ff (?E (q * r)))" and dvd_E: "?E p dvd ?E q * ?E r" by auto from gauss_lemma[of "?E q" "?E r"] divides_ff_trans[OF dvd_ct, of "content_ff_ff q * content_ff_ff r"] have dvd_ct: "divides_ff (content_ff_ff p) (content_ff_ff q * content_ff_ff r)" unfolding eq_dff_def by auto from choice have "p dvd q \ p dvd r" proof assume "degree p \ 0 \ irred (?E p) \ content_ff_ff p =dff 1" hence deg: "degree p \ 0" and irr: "irred (?E p)" and ct: "content_ff_ff p =dff 1" by auto from PFM.irreducible_prime[OF irr] p have prime: "Divisibility.prime mk_monoid (?E p)" by auto from q r have Eq: "?E q \ carrier mk_monoid" and Er: "?E r \ carrier mk_monoid" and q': "?E q \ 0" and r': "?E r \ 0" and qr': "?E q * ?E r \ 0" by auto from PFM.prime_divides[OF Eq Er prime] q' r' qr' dvd_E have dvd_E: "?E p dvd ?E q \ ?E p dvd ?E r" by simp from ct have ct: "divides_ff (content_ff_ff p) 1" unfolding eq_dff_def by auto moreover have "\ q. divides_ff 1 (content_ff_ff q)" using content_ff_map_poly_to_fract unfolding divides_ff_def by auto from divides_ff_trans[OF ct this] have ct: "\ q. divides_ff (content_ff_ff p) (content_ff_ff q)" . with dvd_E show ?thesis using dvd_PM_iff by blast next assume "degree p = 0 \ irred (coeff p 0)" hence deg: "degree p = 0" and irr: "irred (coeff p 0)" by auto from degree0_coeffs[OF deg] p obtain a where p: "p = [:a:]" and a: "a \ 0" by auto with irr have irr: "irred a" and aM: "a \ carrier mk_monoid" by auto from M.irreducible_prime[OF irr aM] have prime: "Divisibility.prime mk_monoid a" . from content_ff_map_poly_to_fract obtain cq where cq: "content_ff_ff q = to_fract cq" by auto from content_ff_map_poly_to_fract obtain cp where cp: "content_ff_ff p = to_fract cp" by auto from content_ff_map_poly_to_fract obtain cr where cr: "content_ff_ff r = to_fract cr" by auto have "divides_ff (to_fract a) (content_ff_ff p)" unfolding p content_ff_iff using a by auto from divides_ff_trans[OF this[unfolded cp] dvd_ct[unfolded cp cq cr]] have "divides_ff (to_fract a) (to_fract (cq * cr))" by simp hence dvd: "a dvd cq * cr" by (auto simp add: divides_ff_def simp del: to_fract_hom.hom_mult simp: to_fract_hom.hom_mult[symmetric]) from content_ff_divides_ff[of "to_fract a" "?E p"] have "divides_ff (to_fract cp) (to_fract a)" using cp a p by auto hence cpa: "cp dvd a" by simp from a q r cq cr have aM: "a \ carrier mk_monoid" and qM: "cq \ carrier mk_monoid" and rM: "cr \ carrier mk_monoid" and q': "cq \ 0" and r': "cr \ 0" and qr': "cq * cr \ 0" by auto from M.prime_divides[OF qM rM prime] q' r' qr' dvd have "a dvd cq \ a dvd cr" by simp with dvd_trans[OF cpa] have dvd: "cp dvd cq \ cp dvd cr" by auto have "\ q. ?E p * (smult (inverse (to_fract a)) q) = q" unfolding p using a by (auto simp: one_poly_def) hence Edvd: "\ q. ?E p dvd q" unfolding dvd_def by metis from dvd Edvd show ?thesis apply (subst(1 2) dvd_PM_iff) unfolding cp cq cr by auto qed thus "factor p q \ factor p r" unfolding factor_idom using p q r by auto qed qed qed instance poly :: (ufd) ufd by (intro class.ufd.of_class.intro factorial_monoid_imp_ufd factorial_monoid_poly) lemma primitive_iff_some_content_dvd_1: fixes f :: "'a :: ufd poly" (* gcd_condition suffices... *) shows "primitive f \ some_gcd.listgcd (coeffs f) dvd 1" (is "_ \ ?c dvd 1") proof(intro iffI primitiveI) fix x assume "(\y. y \ set (coeffs f) \ x dvd y)" from some_gcd.listgcd_greatest[of "coeffs f", OF this] have "x dvd ?c" by simp also assume "?c dvd 1" finally show "x dvd 1". next assume "primitive f" from primitiveD[OF this some_gcd.listgcd[of _ "coeffs f"]] show "?c dvd 1" by auto qed end diff --git a/thys/Budan_Fourier/BF_Misc.thy b/thys/Budan_Fourier/BF_Misc.thy --- a/thys/Budan_Fourier/BF_Misc.thy +++ b/thys/Budan_Fourier/BF_Misc.thy @@ -1,1294 +1,1294 @@ (* Author: Wenda Li *) section "Misc results for polynomials and sign variations" theory BF_Misc imports "HOL-Computational_Algebra.Polynomial_Factorial" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Sturm_Tarski.Sturm_Tarski begin subsection \Induction on polynomial roots\ (*adapted from the poly_root_induct in Polynomial.thy.*) lemma poly_root_induct_alt [case_names 0 no_proots root]: fixes p :: "'a :: idom poly" assumes "Q 0" assumes "\p. (\a. poly p a \ 0) \ Q p" assumes "\a p. Q p \ Q ([:-a, 1:] * p)" shows "Q p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) have ?case when "p=0" using \Q 0\ that by auto moreover have ?case when "\a. poly p a = 0" using assms(2) that by blast moreover have ?case when "\a. poly p a = 0" "p\0" proof - obtain a where "poly p a =0" using \\a. poly p a = 0\ by auto then obtain q where pq:"p= [:-a,1:] * q" by (meson dvdE poly_eq_0_iff_dvd) then have "q\0" using \p\0\ by auto then have "degree qMisc\ lemma lead_coeff_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" shows "lead_coeff (pderiv p) = of_nat (degree p) * lead_coeff p" apply (auto simp:degree_pderiv coeff_pderiv) apply (cases "degree p") by (auto simp add: coeff_eq_0) lemma gcd_degree_le_min: assumes "p\0" "q\0" shows "degree (gcd p q) \ min (degree p) (degree q)" by (simp add: assms(1) assms(2) dvd_imp_degree_le) lemma lead_coeff_normalize_field: fixes p::"'a::{field,semidom_divide_unit_factor} poly" assumes "p\0" shows "lead_coeff (normalize p) = 1" by (metis (no_types, lifting) assms coeff_normalize divide_self_if dvd_field_iff is_unit_unit_factor leading_coeff_0_iff normalize_eq_0_iff normalize_idem) lemma smult_normalize_field_eq: fixes p::"'a::{field,semidom_divide_unit_factor} poly" shows "p = smult (lead_coeff p) (normalize p)" proof (rule poly_eqI) fix n have "unit_factor (lead_coeff p) = lead_coeff p" by (metis dvd_field_iff is_unit_unit_factor unit_factor_0) then show "coeff p n = coeff (smult (lead_coeff p) (normalize p)) n" by simp qed lemma lead_coeff_gcd_field: - fixes p q::"'a::{field,semidom_divide_unit_factor,factorial_ring_gcd} poly" + fixes p q::"'a::field_gcd poly" assumes "p\0 \ q\0" shows "lead_coeff (gcd p q) = 1" using assms by (metis gcd.normalize_idem gcd_eq_0_iff lead_coeff_normalize_field) lemma poly_gcd_0_iff: "poly (gcd p q) x = 0 \ poly p x=0 \ poly q x=0" by (simp add:poly_eq_0_iff_dvd) lemma degree_eq_oneE: fixes p :: "'a::zero poly" assumes "degree p = 1" obtains a b where "p = [:a,b:]" "b\0" proof - obtain a b q where p:"p=pCons a (pCons b q)" by (metis pCons_cases) with assms have "q=0" by (cases "q = 0") simp_all with p have "p=[:a,b:]" by auto moreover then have "b\0" using assms by auto ultimately show ?thesis .. qed subsection \More results about sign variations (i.e. @{term changes}\ lemma changes_0[simp]:"changes (0#xs) = changes xs" by (cases xs) auto lemma changes_Cons:"changes (x#xs) = (if filter (\x. x\0) xs = [] then 0 else if x* hd (filter (\x. x\0) xs) < 0 then 1 + changes xs else changes xs)" apply (induct xs) by auto lemma changes_filter_eq: "changes (filter (\x. x\0) xs) = changes xs" apply (induct xs) by (auto simp add:changes_Cons) lemma changes_filter_empty: assumes "filter (\x. x\0) xs = []" shows "changes xs = 0" "changes (a#xs) = 0" using assms apply (induct xs) apply auto by (metis changes_0 neq_Nil_conv) lemma changes_append: assumes "xs\ [] \ ys\ [] \ (last xs = hd ys \ last xs\0)" shows "changes (xs@ys) = changes xs + changes ys" using assms proof (induct xs) case Nil then show ?case by simp next case (Cons a xs) have ?case when "xs=[]" using that Cons apply (cases ys) by auto moreover have ?case when "ys=[]" using that Cons by auto moreover have ?case when "xs\[]" "ys\[]" proof - have "filter (\x. x \ 0) xs \[]" using that Cons apply auto by (metis (mono_tags, lifting) filter.simps(1) filter.simps(2) filter_append snoc_eq_iff_butlast) then have "changes (a # xs @ ys) = changes (a # xs) + changes ys" apply (subst (1 2) changes_Cons) using that Cons by auto then show ?thesis by auto qed ultimately show ?case by blast qed lemma changes_drop_dup: assumes "xs\ []" "ys\ [] \ last xs=hd ys" shows "changes (xs@ys) = changes (xs@ tl ys)" using assms proof (induct xs) case Nil then show ?case by simp next case (Cons a xs) have ?case when "ys=[]" using that by simp moreover have ?case when "ys\[]" "xs=[]" using that Cons apply auto by (metis changes.simps(3) list.exhaust_sel not_square_less_zero) moreover have ?case when "ys\[]" "xs\[]" proof - define ts ts' where "ts = filter (\x. x \ 0) (xs @ ys)" and "ts' = filter (\x. x \ 0) (xs @ tl ys)" have "(ts = [] \ ts' = []) \ hd ts = hd ts'" proof (cases "filter (\x. x \ 0) xs = []") case True then have "last xs = 0" using \xs\[]\ by (metis (mono_tags, lifting) append_butlast_last_id append_is_Nil_conv filter.simps(2) filter_append list.simps(3)) then have "hd ys=0" using Cons(3)[rule_format, OF \ys\[]\] \xs\[]\ by auto then have "filter (\x. x \ 0) ys = filter (\x. x \ 0) (tl ys)" by (metis (mono_tags, lifting) filter.simps(2) list.exhaust_sel that(1)) then show ?thesis unfolding ts_def ts'_def by auto next case False then show ?thesis unfolding ts_def ts'_def by auto qed moreover have "changes (xs @ ys) = changes (xs @ tl ys)" apply (rule Cons(1)) using that Cons(3) by auto moreover have "changes (a # xs @ ys) = (if ts = [] then 0 else if a * hd ts < 0 then 1 + changes (xs @ ys) else changes (xs @ ys))" using changes_Cons[of a "xs @ ys",folded ts_def] . moreover have "changes (a # xs @ tl ys) = (if ts' = [] then 0 else if a * hd ts' < 0 then 1 + changes (xs @ tl ys) else changes (xs @ tl ys))" using changes_Cons[of a "xs @ tl ys",folded ts'_def] . ultimately show ?thesis by auto qed ultimately show ?case by blast qed (* TODO: the following lemmas contain duplicates of some lemmas in Winding_Number_Eval/Missing_Algebraic.thy Will resolve later. *) lemma Im_poly_of_real: "Im (poly p (of_real x)) = poly (map_poly Im p) x" apply (induct p) by (auto simp add:map_poly_pCons) lemma Re_poly_of_real: "Re (poly p (of_real x)) = poly (map_poly Re p) x" apply (induct p) by (auto simp add:map_poly_pCons) subsection \More about @{term map_poly} and @{term of_real}\ lemma of_real_poly_map_pCons[simp]:"map_poly of_real (pCons a p) = pCons (of_real a) (map_poly of_real p)" by (simp add: map_poly_pCons) lemma of_real_poly_map_plus[simp]: "map_poly of_real (p + q) = map_poly of_real p + map_poly of_real q" apply (rule poly_eqI) by (auto simp add: coeff_map_poly) lemma of_real_poly_map_smult[simp]:"map_poly of_real (smult s p) = smult (of_real s) (map_poly of_real p)" apply (rule poly_eqI) by (auto simp add: coeff_map_poly) lemma of_real_poly_map_mult[simp]:"map_poly of_real (p*q) = map_poly of_real p * map_poly of_real q" by (induct p,intro poly_eqI,auto) lemma of_real_poly_map_poly: "of_real (poly p x) = poly (map_poly of_real p) (of_real x)" by (induct p,auto) lemma of_real_poly_map_power:"map_poly of_real (p^n) = (map_poly of_real p) ^ n" by (induct n,auto) (*FIXME: not duplicate*) lemma of_real_poly_eq_iff [simp]: "map_poly of_real p = map_poly of_real q \ p = q" by (auto simp: poly_eq_iff coeff_map_poly) (*FIXME: not duplicate*) lemma of_real_poly_eq_0_iff [simp]: "map_poly of_real p = 0 \ p = 0" by (auto simp: poly_eq_iff coeff_map_poly) subsection \More about @{term order}\ lemma order_multiplicity_eq: assumes "p\0" shows "order a p = multiplicity [:-a,1:] p" by (metis assms multiplicity_eqI order_1 order_2) lemma order_gcd: assumes "p\0" "q\0" shows "order x (gcd p q) = min (order x p) (order x q)" proof - have "prime [:- x, 1:]" apply (auto simp add: prime_elem_linear_poly normalize_poly_def intro!:primeI) by (simp add: pCons_one) then show ?thesis using assms by (auto simp add:order_multiplicity_eq intro:multiplicity_gcd) qed lemma order_linear[simp]: "order x [:-a,1:] = (if x=a then 1 else 0)" by (auto simp add:order_power_n_n[where n=1,simplified] order_0I) lemma map_poly_order_of_real: assumes "p\0" shows "order (of_real t) (map_poly of_real p) = order t p" using assms proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) then have "order t p = 0" using order_root by blast moreover have "poly (map_poly of_real p) (of_real x) \0" for x apply (subst of_real_poly_map_poly[symmetric]) using no_proots order_root by simp then have "order (of_real t) (map_poly of_real p) = 0" using order_root by blast ultimately show ?case by auto next case (root a p) define a1 where "a1=[:-a,1:]" have [simp]:"a1\0" "p\0" unfolding a1_def using root(2) by auto have "order (of_real t) (map_poly of_real a1) = order t a1" unfolding a1_def by simp then show ?case apply (fold a1_def) by (simp add:order_mult root) qed lemma order_pcompose: assumes "pcompose p q\0" shows "order x (pcompose p q) = order x (q-[:poly q x:]) * order (poly q x) p" using \pcompose p q\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have "order x (p \\<^sub>p q) = 0" apply (rule order_0I) using no_proots by (auto simp:poly_pcompose) moreover have "order (poly q x) p = 0" apply (rule order_0I) using no_proots by (auto simp:poly_pcompose) ultimately show ?case by auto next case (root a p) define a1 where "a1=[:-a,1:]" have [simp]: "a1\0" "p\0" "a1 \\<^sub>p q \0" "p \\<^sub>p q \ 0" subgoal using root(2) unfolding a1_def by simp subgoal using root(2) by auto using root(2) by (fold a1_def,auto simp:pcompose_mult) have "order x ((a1 * p) \\<^sub>p q) = order x (a1 \\<^sub>p q) + order x (p \\<^sub>p q)" unfolding pcompose_mult by (auto simp: order_mult) also have "... = order x (q-[:poly q x:]) * (order (poly q x) a1 + order (poly q x) p)" proof - have "order x (a1 \\<^sub>p q) = order x (q-[:poly q x:]) * order (poly q x) a1" unfolding a1_def apply (auto simp: pcompose_pCons algebra_simps diff_conv_add_uminus ) by (simp add: order_0I) moreover have "order x (p \\<^sub>p q) = order x (q - [:poly q x:]) * order (poly q x) p" apply (rule root.hyps) by auto ultimately show ?thesis by (auto simp:algebra_simps) qed also have "... = order x (q - [:poly q x:]) * order (poly q x) (a1 * p)" by (auto simp:order_mult) finally show ?case unfolding a1_def . qed subsection \Polynomial roots / zeros\ definition proots_within::"'a::comm_semiring_0 poly \ 'a set \ 'a set" where "proots_within p s = {x\s. poly p x=0}" abbreviation proots::"'a::comm_semiring_0 poly \ 'a set" where "proots p \ proots_within p UNIV" lemma proots_def: "proots p = {x. poly p x=0}" unfolding proots_within_def by auto lemma proots_within_empty[simp]: "proots_within p {} = {}" unfolding proots_within_def by auto lemma proots_within_0[simp]: "proots_within 0 s = s" unfolding proots_within_def by auto lemma proots_withinI[intro,simp]: "poly p x=0 \ x\s \ x\proots_within p s" unfolding proots_within_def by auto lemma proots_within_iff[simp]: "x\proots_within p s \ poly p x=0 \ x\s" unfolding proots_within_def by auto lemma proots_within_union: "proots_within p A \ proots_within p B = proots_within p (A \ B)" unfolding proots_within_def by auto lemma proots_within_times: fixes s::"'a::{semiring_no_zero_divisors,comm_semiring_0} set" shows "proots_within (p*q) s = proots_within p s \ proots_within q s" unfolding proots_within_def by auto lemma proots_within_gcd: - fixes s::"'a::factorial_ring_gcd set" + fixes s::"'a::{factorial_ring_gcd,semiring_gcd_mult_normalize} set" shows "proots_within (gcd p q) s= proots_within p s \ proots_within q s" unfolding proots_within_def by (auto simp add: poly_eq_0_iff_dvd) lemma proots_within_inter: "NO_MATCH UNIV s \ proots_within p s = proots p \ s" unfolding proots_within_def by auto lemma proots_within_proots[simp]: "proots_within p s \ proots p" unfolding proots_within_def by auto lemma finite_proots[simp]: fixes p :: "'a::idom poly" shows "p\0 \ finite (proots_within p s)" unfolding proots_within_def using poly_roots_finite by fast lemma proots_within_pCons_1_iff: fixes a::"'a::idom" shows "proots_within [:-a,1:] s = (if a\s then {a} else {})" "proots_within [:a,-1:] s = (if a\s then {a} else {})" by (cases "a\s",auto) lemma proots_within_uminus[simp]: fixes p :: "'a::comm_ring poly" shows "proots_within (- p) s = proots_within p s" by auto lemma proots_within_smult: fixes a::"'a::{semiring_no_zero_divisors,comm_semiring_0}" assumes "a\0" shows "proots_within (smult a p) s = proots_within p s" unfolding proots_within_def using assms by auto subsection \Polynomial roots counting multiplicities.\ (*counting the number of proots WITH MULTIPLICITIES within a set*) definition proots_count::"'a::idom poly \ 'a set \ nat" where "proots_count p s = (\r\proots_within p s. order r p)" lemma proots_count_emtpy[simp]:"proots_count p {} = 0" unfolding proots_count_def by auto lemma proots_count_times: fixes s :: "'a::idom set" assumes "p*q\0" shows "proots_count (p*q) s = proots_count p s + proots_count q s" proof - define pts where "pts=proots_within p s" define qts where "qts=proots_within q s" have [simp]: "finite pts" "finite qts" using \p*q\0\ unfolding pts_def qts_def by auto have "(\r\pts \ qts. order r p) = (\r\pts. order r p)" proof (rule comm_monoid_add_class.sum.mono_neutral_cong_right,simp_all) show "\i\pts \ qts - pts. order i p = 0" unfolding pts_def qts_def proots_within_def using order_root by fastforce qed moreover have "(\r\pts \ qts. order r q) = (\r\qts. order r q)" proof (rule comm_monoid_add_class.sum.mono_neutral_cong_right,simp_all) show "\i\pts \ qts - qts. order i q = 0" unfolding pts_def qts_def proots_within_def using order_root by fastforce qed ultimately show ?thesis unfolding proots_count_def apply (simp add:proots_within_times order_mult[OF \p*q\0\] sum.distrib) apply (fold pts_def qts_def) by auto qed lemma proots_count_power_n_n: shows "proots_count ([:- a, 1:]^n) s = (if a\s \ n>0 then n else 0)" proof - have "proots_within ([:- a, 1:] ^ n) s= (if a\s \ n>0 then {a} else {})" unfolding proots_within_def by auto thus ?thesis unfolding proots_count_def using order_power_n_n by auto qed lemma degree_proots_count: fixes p::"complex poly" shows "degree p = proots_count p UNIV" proof (induct "degree p" arbitrary:p ) case 0 then obtain c where c_def:"p=[:c:]" using degree_eq_zeroE by auto then show ?case unfolding proots_count_def apply (cases "c=0") by (auto intro!:sum.infinite simp add:infinite_UNIV_char_0 order_0I) next case (Suc n) then have "degree p\0" and "p\0" by auto obtain z where "poly p z = 0" using Fundamental_Theorem_Algebra.fundamental_theorem_of_algebra \degree p\0\ constant_degree[of p] by auto define onez where "onez=[:-z,1:]" have [simp]: "onez\0" "degree onez = 1" unfolding onez_def by auto obtain q where q_def:"p= onez * q" using poly_eq_0_iff_dvd \poly p z = 0\ dvdE unfolding onez_def by blast hence "q\0" using \p\0\ by auto hence "n=degree q" using degree_mult_eq[of onez q] \Suc n = degree p\ apply (fold q_def) by auto hence "degree q = proots_count q UNIV" using Suc.hyps(1) by simp moreover have " Suc 0 = proots_count onez UNIV" unfolding onez_def using proots_count_power_n_n[of z 1 UNIV] by auto ultimately show ?case unfolding q_def using degree_mult_eq[of onez q] proots_count_times[of onez q UNIV] \q\0\ by auto qed lemma proots_count_smult: fixes a::"'a::{semiring_no_zero_divisors,idom}" assumes "a\0" shows "proots_count (smult a p) s= proots_count p s" proof (cases "p=0") case True then show ?thesis by auto next case False then show ?thesis unfolding proots_count_def using order_smult[OF assms] proots_within_smult[OF assms] by auto qed lemma proots_count_pCons_1_iff: fixes a::"'a::idom" shows "proots_count [:-a,1:] s = (if a\s then 1 else 0)" unfolding proots_count_def by (cases "a\s",auto simp add:proots_within_pCons_1_iff order_power_n_n[of _ 1,simplified]) lemma proots_count_uminus[simp]: "proots_count (- p) s = proots_count p s" unfolding proots_count_def by simp lemma card_proots_within_leq: assumes "p\0" shows "proots_count p s \ card (proots_within p s)" using assms proof (induct rule:poly_root_induct[of _ "\x. x\s"]) case 0 then show ?case unfolding proots_within_def proots_count_def by auto next case (no_roots p) then have "proots_within p s = {}" by auto then show ?case unfolding proots_count_def by auto next case (root a p) have "card (proots_within ([:- a, 1:] * p) s) \ card (proots_within [:- a, 1:] s)+card (proots_within p s)" unfolding proots_within_times by (auto simp add:card_Un_le) also have "... \ 1+ proots_count p s" proof - have "card (proots_within [:- a, 1:] s) \ 1" proof (cases "a\s") case True then have "proots_within [:- a, 1:] s = {a}" by auto then show ?thesis by auto next case False then have "proots_within [:- a, 1:] s = {}" by auto then show ?thesis by auto qed moreover have "card (proots_within p s) \ proots_count p s" apply (rule root.hyps) using root by auto ultimately show ?thesis by auto qed also have "... = proots_count ([:- a,1:] * p) s" apply (subst proots_count_times) subgoal by (metis mult_eq_0_iff pCons_eq_0_iff root.prems zero_neq_one) using root by (auto simp add:proots_count_pCons_1_iff) finally have "card (proots_within ([:- a, 1:] * p) s) \ proots_count ([:- a, 1:] * p) s" . then show ?case by (metis (no_types, hide_lams) add.inverse_inverse add.inverse_neutral minus_pCons mult_minus_left proots_count_uminus proots_within_uminus) qed (*FIXME: not duplicate*) lemma proots_count_0_imp_empty: assumes "proots_count p s=0" "p\0" shows "proots_within p s = {}" proof - have "card (proots_within p s) = 0" using card_proots_within_leq[OF \p\0\,of s] \proots_count p s=0\ by auto moreover have "finite (proots_within p s)" using \p\0\ by auto ultimately show ?thesis by auto qed lemma proots_count_leq_degree: assumes "p\0" shows "proots_count p s\ degree p" using assms proof (induct rule:poly_root_induct[of _ "\x. x\s"]) case 0 then show ?case by auto next case (no_roots p) then have "proots_within p s = {}" by auto then show ?case unfolding proots_count_def by auto next case (root a p) have "proots_count ([:a, - 1:] * p) s = proots_count [:a, - 1:] s + proots_count p s" apply (subst proots_count_times) using root by auto also have "... = 1 + proots_count p s" proof - have "proots_count [:a, - 1:] s =1" by (metis (no_types, lifting) add.inverse_inverse add.inverse_neutral minus_pCons proots_count_pCons_1_iff proots_count_uminus root.hyps(1)) then show ?thesis by auto qed also have "... \ degree ([:a,-1:] * p)" apply (subst degree_mult_eq) subgoal by auto subgoal using root by auto subgoal using root by (simp add: \p \ 0\) done finally show ?case . qed (*TODO: not a duplicate*) lemma proots_count_union_disjoint: assumes "A \ B = {}" "p\0" shows "proots_count p (A \ B) = proots_count p A + proots_count p B" unfolding proots_count_def apply (subst proots_within_union[symmetric]) apply (subst sum.union_disjoint) using assms by auto lemma proots_count_cong: assumes order_eq:"\x\s. order x p = order x q" and "p\0" and "q\0" shows "proots_count p s = proots_count q s" unfolding proots_count_def proof (rule sum.cong) have "poly p x = 0 \ poly q x = 0" when "x\s" for x using order_eq that by (simp add: assms(2) assms(3) order_root) then show "proots_within p s = proots_within q s" by auto show "\x. x \ proots_within q s \ order x p = order x q" using order_eq by auto qed lemma proots_count_of_real: assumes "p\0" shows "proots_count (map_poly of_real p) ((of_real::_\'a::{real_algebra_1,idom}) ` s) = proots_count p s" proof - define k where "k=(of_real::_\'a)" have "proots_within (map_poly of_real p) (k ` s) =k ` (proots_within p s)" unfolding proots_within_def k_def by (auto simp add:of_real_poly_map_poly[symmetric]) then have "proots_count (map_poly of_real p) (k ` s) = (\r\k ` (proots_within p s). order r (map_poly of_real p))" unfolding proots_count_def by simp also have "... = sum ((\r. order r (map_poly of_real p)) \ k) (proots_within p s)" apply (subst sum.reindex) unfolding k_def by (auto simp add: inj_on_def) also have "... = proots_count p s" unfolding proots_count_def apply (rule sum.cong) unfolding k_def comp_def using \p\0\ by (auto simp add:map_poly_order_of_real) finally show ?thesis unfolding k_def . qed (*Is field really necessary here?*) lemma proots_pcompose: fixes p q::"'a::field poly" assumes "p\0" "degree q=1" shows "proots_count (pcompose p q) s = proots_count p (poly q ` s)" proof - obtain a b where ab:"q=[:a,b:]" "b\0" using \degree q=1\ degree_eq_oneE by metis define f where "f=(\y. (y-a)/b)" have f_eq:"f (poly q x) = x" "poly q (f x) = x" for x unfolding f_def using ab by auto have "proots_count (p \\<^sub>p q) s = (\r\ f ` proots_within p (poly q ` s). order r (p \\<^sub>p q))" unfolding proots_count_def apply (rule arg_cong2[where f =sum]) apply (auto simp:poly_pcompose proots_within_def f_eq) by (metis (mono_tags, lifting) f_eq(1) image_eqI mem_Collect_eq) also have "... = (\x\proots_within p (poly q ` s). order (f x) (p \\<^sub>p q))" apply (subst sum.reindex) subgoal unfolding f_def inj_on_def using \b\0\ by auto by simp also have "... = (\x\proots_within p (poly q ` s). order x p)" proof - have "p \\<^sub>p q \ 0" using assms(1) assms(2) pcompose_eq_0 by force moreover have "order (f x) (q - [:x:]) = 1" for x proof - have "order (f x) (q - [:x:]) = order (f x) (smult b [:-((x - a) / b),1:])" unfolding f_def using ab by auto also have "... = 1" apply (subst order_smult) using \b\0\ unfolding f_def by auto finally show ?thesis . qed ultimately have "order (f x) (p \\<^sub>p q) = order x p" for x apply (subst order_pcompose) using f_eq by auto then show ?thesis by auto qed also have "... = proots_count p (poly q ` s)" unfolding proots_count_def by auto finally show ?thesis . qed subsection \Composition of a polynomial and a rational function\ (*composition of a polynomial and a rational function. Maybe a more general version in the future?*) definition fcompose::"'a ::field poly \ 'a poly \ 'a poly \ 'a poly" where "fcompose p q r = fst (fold_coeffs (\a (c,d). (d*[:a:] + q * c,r*d)) p (0,1))" lemma fcompose_0 [simp]: "fcompose 0 q r = 0" by (simp add: fcompose_def) lemma fcompose_const[simp]:"fcompose [:a:] q r = [:a:]" unfolding fcompose_def by (cases "a=0") auto lemma fcompose_pCons: "fcompose (pCons a p) q1 q2 = smult a (q2^(degree (pCons a p))) + q1 * fcompose p q1 q2" proof (cases "p=0") case False define ff where "ff=(\a (c, d). (d * [:a:] + q1 * c, q2 * d))" define fc where "fc=fold_coeffs ff p (0, 1)" have snd_ff:"snd fc = (if p=0 then 1 else q2^(degree p + 1))" unfolding fc_def apply (induct p) subgoal by simp subgoal for a p by (auto simp add:ff_def split:if_splits prod.splits) done have "fcompose (pCons a p) q1 q2 = fst (fold_coeffs ff (pCons a p) (0, 1))" unfolding fcompose_def ff_def by simp also have "... = fst (ff a fc)" using False unfolding fc_def by auto also have "... = snd fc * [:a:] + q1 * fst fc" unfolding ff_def by (auto split:prod.splits) also have "... = smult a (q2^(degree (pCons a p))) + q1 * fst fc" using snd_ff False by auto also have "... = smult a (q2^(degree (pCons a p))) + q1 * fcompose p q1 q2" unfolding fc_def ff_def fcompose_def by simp finally show ?thesis . qed simp lemma fcompose_uminus: "fcompose (-p) q r = - fcompose p q r" by (induct p) (auto simp:fcompose_pCons) lemma fcompose_add_less: assumes "degree p1 > degree p2" shows "fcompose (p1+p2) q1 q2 = fcompose p1 q1 q2 + q2^(degree p1-degree p2) * fcompose p2 q1 q2" using assms proof (induction p1 p2 rule: poly_induct2) case (pCons a1 p1 a2 p2) have ?case when "p2=0" using that by (simp add:fcompose_pCons smult_add_left) moreover have ?case when "p2\0" "\ degree p2 < degree p1" using that pCons(2) by auto moreover have ?case when "p2\0" "degree p2< degree p1" proof - define d1 d2 where "d1=degree (pCons a1 p1)" and "d2=degree (pCons a2 p2) " define fp1 fp2 where "fp1= fcompose p1 q1 q2" and "fp2=fcompose p2 q1 q2" have "fcompose (pCons a1 p1 + pCons a2 p2) q1 q2 = fcompose (pCons (a1+a2) (p1+p2)) q1 q2" by simp also have "... = smult (a1 + a2) (q2 ^ d1) + q1 * fcompose (p1 + p2) q1 q2" proof - have "degree (pCons (a1 + a2) (p1 + p2)) = d1" unfolding d1_def using that degree_add_eq_left by fastforce then show ?thesis unfolding fcompose_pCons by simp qed also have "... = smult (a1 + a2) (q2 ^ d1) + q1 * (fp1 + q2 ^ (d1 - d2) * fp2)" proof - have "degree p1 - degree p2 = d1 - d2" unfolding d1_def d2_def using that by simp then show ?thesis unfolding pCons(1)[OF that(2),folded fp1_def fp2_def] by simp qed also have "... = fcompose (pCons a1 p1) q1 q2 + q2 ^ (d1 - d2) * fcompose (pCons a2 p2) q1 q2" proof - have "d1 > d2" unfolding d1_def d2_def using that by auto then show ?thesis unfolding fcompose_pCons apply (fold d1_def d2_def fp1_def fp2_def) by (simp add:algebra_simps smult_add_left power_add[symmetric]) qed finally show ?thesis unfolding d1_def d2_def . qed ultimately show ?case by blast qed simp lemma fcompose_add_eq: assumes "degree p1 = degree p2" shows "q2^(degree p1 - degree (p1+p2)) * fcompose (p1+p2) q1 q2 = fcompose p1 q1 q2 + fcompose p2 q1 q2" using assms proof (induction p1 p2 rule: poly_induct2) case (pCons a1 p1 a2 p2) have ?case when "p1+p2=0" proof - have "p2=-p1" using that by algebra then show ?thesis by (simp add:fcompose_pCons fcompose_uminus smult_add_left) qed moreover have ?case when "p1=0" proof - have "p2=0" using pCons(2) that by (auto split:if_splits) then show ?thesis using that by simp qed moreover have ?case when "p1\0" "p1+p2\0" proof - define d1 d2 dp where "d1=degree (pCons a1 p1)" and "d2=degree (pCons a2 p2)" and "dp = degree p1 - degree (p1+p2)" define fp1 fp2 where "fp1= fcompose p1 q1 q2" and "fp2=fcompose p2 q1 q2" have "q2 ^ (degree (pCons a1 p1) - degree (pCons a1 p1 + pCons a2 p2)) * fcompose (pCons a1 p1 + pCons a2 p2) q1 q2 = q2 ^ dp * fcompose (pCons (a1+a2) (p1 +p2)) q1 q2" unfolding dp_def using that by auto also have "... = smult (a1 + a2) (q2 ^ d1) + q1 * (q2 ^ dp * fcompose (p1 + p2) q1 q2)" proof - have "degree p1 \ degree (p1 + p2)" by (metis degree_add_le degree_pCons_eq_if not_less_eq_eq order_refl pCons.prems zero_le) then show ?thesis unfolding fcompose_pCons dp_def d1_def using that by (simp add:algebra_simps power_add[symmetric]) qed also have "... = smult (a1 + a2) (q2 ^ d1) + q1 * (fp1 + fp2)" apply (subst pCons(1)[folded dp_def fp1_def fp2_def]) subgoal by (metis degree_pCons_eq_if diff_Suc_Suc diff_zero not_less_eq_eq pCons.prems zero_le) subgoal by simp done also have "... = fcompose (pCons a1 p1) q1 q2 + fcompose (pCons a2 p2) q1 q2" proof - have *:"d1 = degree (pCons a2 p2)" unfolding d1_def using pCons(2) by simp show ?thesis unfolding fcompose_pCons apply (fold d1_def fp1_def fp2_def *) by (simp add:smult_add_left algebra_simps) qed finally show ?thesis . qed ultimately show ?case by blast qed simp lemma fcompose_add_const: "fcompose ([:a:] + p) q1 q2 = smult a (q2 ^ degree p) + fcompose p q1 q2" apply (cases p) by (auto simp add:fcompose_pCons smult_add_left) lemma fcompose_smult: "fcompose (smult a p) q1 q2 = smult a (fcompose p q1 q2)" by (induct p) (simp_all add:fcompose_pCons smult_add_right) lemma fcompose_mult: "fcompose (p1*p2) q1 q2 = fcompose p1 q1 q2 * fcompose p2 q1 q2" proof (induct p1) case 0 then show ?case by simp next case (pCons a p1) have ?case when "p1=0 \ p2=0" using that by (auto simp add:fcompose_smult) moreover have ?case when "p1\0" "p2\0" "a=0" using that by (simp add:fcompose_pCons pCons) moreover have ?case when "p1\0" "p2\0" "a\0" proof - have "fcompose (pCons a p1 * p2) q1 q2 = fcompose (pCons 0 (p1 * p2) + smult a p2) q1 q2" by (simp add:algebra_simps) also have "... = fcompose (pCons 0 (p1 * p2)) q1 q2 + q2 ^ (degree p1 +1) * fcompose (smult a p2) q1 q2" proof - have "degree (pCons 0 (p1 * p2)) > degree (smult a p2)" using that by (simp add: degree_mult_eq) from fcompose_add_less[OF this,of q1 q2] that show ?thesis by (simp add:degree_mult_eq) qed also have "... = fcompose (pCons a p1) q1 q2 * fcompose p2 q1 q2" using that by (simp add:fcompose_pCons fcompose_smult pCons algebra_simps) finally show ?thesis . qed ultimately show ?case by blast qed lemma fcompose_poly: assumes "poly q2 x\0" shows "poly p (poly q1 x/poly q2 x) = poly (fcompose p q1 q2) x / poly (q2^(degree p)) x" apply (induct p) using assms by (simp_all add:fcompose_pCons field_simps) lemma poly_fcompose: assumes "poly q2 x\0" shows "poly (fcompose p q1 q2) x = poly p (poly q1 x/poly q2 x) * (poly q2 x)^(degree p)" using fcompose_poly[OF assms] assms by (auto simp add:field_simps) lemma poly_fcompose_0_denominator: assumes "poly q2 x=0" shows "poly (fcompose p q1 q2) x = poly q1 x ^ degree p * lead_coeff p" apply (induct p) using assms by (auto simp add:fcompose_pCons) lemma fcompose_0_denominator:"fcompose p q1 0 = smult (lead_coeff p) (q1^degree p)" apply (induct p) by (auto simp:fcompose_pCons) lemma fcompose_nzero: fixes p::"'a::field poly" assumes "p\0" and "q2\0" and nconst:"\c. q1 \ smult c q2" and infi:"infinite (UNIV::'a set)" shows "fcompose p q1 q2 \ 0" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have False when "fcompose p q1 q2 = 0" proof - obtain x where "poly q2 x\0" proof - have "finite (proots q2)" using \q2\0\ by auto then have "\x. poly q2 x\0" by (meson UNIV_I ex_new_if_finite infi proots_withinI) then show ?thesis using that by auto qed define y where "y = poly q1 x / poly q2 x" have "poly p y = 0" using \fcompose p q1 q2 = 0\ fcompose_poly[OF \poly q2 x\0\,of p q1,folded y_def] by simp then show False using no_proots(1) by auto qed then show ?case by auto next case (root a p) have "fcompose [:- a, 1:] q1 q2 \ 0" unfolding fcompose_def using nconst[rule_format,of a] by simp moreover have "fcompose p q1 q2 \ 0" using root by fastforce ultimately show ?case unfolding fcompose_mult by auto qed subsection \Bijection (@{term bij_betw}) and the number of polynomial roots\ lemma proots_fcompose_bij_eq: fixes p::"'a::field poly" assumes bij:"bij_betw (\x. poly q1 x/poly q2 x) A B" and "p\0" and nzero:"\x\A. poly q2 x\0" and max_deg: "max (degree q1) (degree q2) \ 1" and nconst:"\c. q1 \ smult c q2" and infi:"infinite (UNIV::'a set)" shows "proots_count p B = proots_count (fcompose p q1 q2) A" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have "proots_count p B = 0" proof - have "proots_within p B = {}" using no_proots by auto then show ?thesis unfolding proots_count_def by auto qed moreover have "proots_count (fcompose p q1 q2) A = 0" proof - have "proots_within (fcompose p q1 q2) A = {}" using no_proots unfolding proots_within_def by (smt div_0 empty_Collect_eq fcompose_poly nzero) then show ?thesis unfolding proots_count_def by auto qed ultimately show ?case by auto next case (root b p) have "proots_count ([:- b, 1:] * p) B = proots_count [:- b, 1:] B + proots_count p B" using proots_count_times[OF \[:- b, 1:] * p \ 0\] by simp also have "... = proots_count (fcompose [:- b, 1:] q1 q2) A + proots_count (fcompose p q1 q2) A" proof - define g where "g=(\x. poly q1 x/poly q2 x)" have "proots_count [:- b, 1:] B = proots_count (fcompose [:- b, 1:] q1 q2) A" proof (cases "b\B") case True then have "proots_count [:- b, 1:] B = 1" unfolding proots_count_pCons_1_iff by simp moreover have "proots_count (fcompose [:- b, 1:] q1 q2) A = 1" proof - obtain a where "b=g a" "a\A" using bij[folded g_def] True by (metis bij_betwE bij_betw_the_inv_into f_the_inv_into_f_bij_betw) define qq where "qq=q1 - smult b q2" have qq_0:"poly qq a=0" and qq_deg: "degree qq\1" and \qq\0\ unfolding qq_def subgoal using \b=g a\ nzero[rule_format,OF \a\A\] unfolding g_def by auto subgoal using max_deg by (simp add: degree_diff_le) subgoal using nconst[rule_format,of b] by auto done have "proots_within qq A = {a}" proof - have "a\proots_within qq A" using qq_0 \a\A\ by auto moreover have "card (proots_within qq A) = 1" proof - have "finite (proots_within qq A)" using \qq\0\ by simp moreover have "proots_within qq A \ {}" using \a\proots_within qq A\ by auto ultimately have "card (proots_within qq A) \0" by auto moreover have "card (proots_within qq A) \ 1" by (meson \qq \ 0\ card_proots_within_leq le_trans proots_count_leq_degree qq_deg) ultimately show ?thesis by auto qed ultimately show ?thesis by (metis card_1_singletonE singletonD) qed moreover have "order a qq=1" by (metis One_nat_def \qq \ 0\ le_antisym le_zero_eq not_less_eq_eq order_degree order_root qq_0 qq_deg) ultimately show ?thesis unfolding fcompose_def proots_count_def qq_def by auto qed ultimately show ?thesis by auto next case False then have "proots_count [:- b, 1:] B = 0" unfolding proots_count_pCons_1_iff by simp moreover have "proots_count (fcompose [:- b, 1:] q1 q2) A = 0" proof - have "proots_within (fcompose [:- b, 1:] q1 q2) A = {}" proof (rule ccontr) assume "proots_within (fcompose [:- b, 1:] q1 q2) A \ {}" then obtain a where "a\A" "poly q1 a = b * poly q2 a" unfolding fcompose_def proots_within_def by auto then have "b = g a" unfolding g_def using nzero[rule_format,OF \a\A\] by auto then have "b\B" using \a\A\ bij[folded g_def] using bij_betwE by blast then show False using False by auto qed then show ?thesis unfolding proots_count_def by auto qed ultimately show ?thesis by simp qed moreover have "proots_count p B = proots_count (fcompose p q1 q2) A" apply (rule root.hyps) using mult_eq_0_iff root.prems by blast ultimately show ?thesis by auto qed also have "... = proots_count (fcompose ([:- b, 1:] * p) q1 q2) A" proof (cases "A={}") case False have "fcompose [:- b, 1:] q1 q2 \0" using nconst[rule_format,of b] unfolding fcompose_def by auto moreover have "fcompose p q1 q2 \ 0" apply (rule fcompose_nzero[OF _ _ nconst infi]) subgoal using \[:- b, 1:] * p \ 0\ by auto subgoal using nzero False by auto done ultimately show ?thesis unfolding fcompose_mult apply (subst proots_count_times) by auto qed auto finally show ?case . qed lemma proots_card_fcompose_bij_eq: fixes p::"'a::field poly" assumes bij:"bij_betw (\x. poly q1 x/poly q2 x) A B" and "p\0" and nzero:"\x\A. poly q2 x\0" and max_deg: "max (degree q1) (degree q2) \ 1" and nconst:"\c. q1 \ smult c q2" and infi:"infinite (UNIV::'a set)" shows "card (proots_within p B) = card (proots_within (fcompose p q1 q2) A)" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have "proots_within p B = {}" using no_proots by auto moreover have "proots_within (fcompose p q1 q2) A = {}" using no_proots fcompose_poly by (smt Collect_empty_eq divide_eq_0_iff nzero proots_within_def) ultimately show ?case by auto next case (root b p) then have [simp]:"p\0" by auto have ?case when "b\B \ poly p b=0" proof - have "proots_within ([:- b, 1:] * p) B = proots_within p B" using that by auto moreover have "proots_within (fcompose ([:- b, 1:] * p) q1 q2) A = proots_within (fcompose p q1 q2) A" using that nzero unfolding fcompose_mult proots_within_times apply (auto simp add: poly_fcompose) using bij bij_betwE by blast ultimately show ?thesis using root by auto qed moreover have ?case when "b\B" "poly p b\0" proof - define bb where "bb=[:- b, 1:]" have "card (proots_within (bb * p) B) = card {b} + card (proots_within p B)" proof - have "proots_within bb B = {b}" using that unfolding bb_def by auto then show ?thesis unfolding proots_within_times apply (subst card_Un_disjoint) by (use that in auto) qed also have "... = 1 + card (proots_within (fcompose p q1 q2) A)" using root.hyps by simp also have "... = card (proots_within (fcompose (bb * p) q1 q2) A)" unfolding proots_within_times fcompose_mult proof (subst card_Un_disjoint) obtain a where b_poly:"b=poly q1 a / poly q2 a" and "a\A" by (metis (no_types, lifting) \b \ B\ bij bij_betwE bij_betw_the_inv_into f_the_inv_into_f_bij_betw) define bbq pq where "bbq=fcompose bb q1 q2" and "pq=fcompose p q1 q2" have bbq_0:"poly bbq a=0" and bbq_deg: "degree bbq\1" and "bbq\0" unfolding bbq_def bb_def subgoal using \a \ A\ b_poly nzero poly_fcompose by fastforce subgoal by (metis (no_types, lifting) degree_add_le degree_pCons_eq_if degree_smult_le dual_order.trans fcompose_const fcompose_pCons max.boundedE max_deg mult_cancel_left2 one_neq_zero one_poly_eq_simps(1) power.simps) subgoal by (metis \a \ A\ \poly (fcompose [:- b, 1:] q1 q2) a = 0\ fcompose_nzero infi nconst nzero one_neq_zero pCons_eq_0_iff) done show "finite (proots_within bbq A)" using \bbq\0\ by simp show "finite (proots_within pq A)" unfolding pq_def by (metis \a \ A\ \p \ 0\ fcompose_nzero finite_proots infi nconst nzero poly_0 pq_def) have bbq_a:"proots_within bbq A = {a}" proof - have "a\proots_within bbq A" by (simp add: \a \ A\ bbq_0) moreover have "card (proots_within bbq A) = 1" proof - have "card (proots_within bbq A) \0" using \a\proots_within bbq A\ \finite (proots_within bbq A)\ by auto moreover have "card (proots_within bbq A) \ 1" by (meson \bbq \ 0\ card_proots_within_leq le_trans proots_count_leq_degree bbq_deg) ultimately show ?thesis by auto qed ultimately show ?thesis by (metis card_1_singletonE singletonD) qed show "proots_within (bbq) A \ proots_within (pq) A = {}" using b_poly bbq_a fcompose_poly nzero pq_def that(2) by fastforce show "1 + card (proots_within pq A) = card (proots_within bbq A) + card (proots_within pq A)" using bbq_a by simp qed finally show ?thesis unfolding bb_def . qed ultimately show ?case by auto qed lemma proots_pcompose_bij_eq: fixes p::"'a::idom poly" assumes bij:"bij_betw (\x. poly q x) A B" and "p\0" and q_deg: "degree q = 1" shows "proots_count p B = proots_count (p \\<^sub>p q) A" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have "proots_count p B = 0" proof - have "proots_within p B = {}" using no_proots by auto then show ?thesis unfolding proots_count_def by auto qed moreover have "proots_count (p \\<^sub>p q) A = 0" proof - have "proots_within (p \\<^sub>p q) A = {}" using no_proots unfolding proots_within_def by (auto simp:poly_pcompose) then show ?thesis unfolding proots_count_def by auto qed ultimately show ?case by auto next case (root b p) have "proots_count ([:- b, 1:] * p) B = proots_count [:- b, 1:] B + proots_count p B" using proots_count_times[OF \[:- b, 1:] * p \ 0\] by simp also have "... = proots_count ([:- b, 1:] \\<^sub>p q) A + proots_count (p \\<^sub>p q) A" proof - have "proots_count [:- b, 1:] B = proots_count ([:- b, 1:] \\<^sub>p q) A" proof (cases "b\B") case True then have "proots_count [:- b, 1:] B = 1" unfolding proots_count_pCons_1_iff by simp moreover have "proots_count ([:- b, 1:] \\<^sub>p q) A = 1" proof - obtain a where "b=poly q a" "a\A" using True bij by (metis bij_betwE bij_betw_the_inv_into f_the_inv_into_f_bij_betw) define qq where "qq=[:- b:] + q" have qq_0:"poly qq a=0" and qq_deg: "degree qq\1" and \qq\0\ unfolding qq_def subgoal using \b=poly q a\ by auto subgoal using q_deg by (simp add: degree_add_le) subgoal using q_deg add.inverse_unique by force done have "proots_within qq A = {a}" proof - have "a\proots_within qq A" using qq_0 \a\A\ by auto moreover have "card (proots_within qq A) = 1" proof - have "finite (proots_within qq A)" using \qq\0\ by simp moreover have "proots_within qq A \ {}" using \a\proots_within qq A\ by auto ultimately have "card (proots_within qq A) \0" by auto moreover have "card (proots_within qq A) \ 1" by (meson \qq \ 0\ card_proots_within_leq le_trans proots_count_leq_degree qq_deg) ultimately show ?thesis by auto qed ultimately show ?thesis by (metis card_1_singletonE singletonD) qed moreover have "order a qq=1" by (metis One_nat_def \qq \ 0\ le_antisym le_zero_eq not_less_eq_eq order_degree order_root qq_0 qq_deg) ultimately show ?thesis unfolding pcompose_def proots_count_def qq_def by auto qed ultimately show ?thesis by auto next case False then have "proots_count [:- b, 1:] B = 0" unfolding proots_count_pCons_1_iff by simp moreover have "proots_count ([:- b, 1:] \\<^sub>p q) A = 0" proof - have "proots_within ([:- b, 1:] \\<^sub>p q) A = {}" unfolding pcompose_def apply auto using False bij bij_betwE by blast then show ?thesis unfolding proots_count_def by auto qed ultimately show ?thesis by simp qed moreover have "proots_count p B = proots_count (p \\<^sub>p q) A" apply (rule root.hyps) using \[:- b, 1:] * p \ 0\ by auto ultimately show ?thesis by auto qed also have "... = proots_count (([:- b, 1:] * p) \\<^sub>p q) A" unfolding pcompose_mult apply (subst proots_count_times) subgoal by (metis (no_types, lifting) One_nat_def add.right_neutral degree_0 degree_mult_eq degree_pCons_eq_if degree_pcompose mult_eq_0_iff one_neq_zero one_pCons pcompose_mult q_deg root.prems) by simp finally show ?case . qed lemma proots_card_pcompose_bij_eq: fixes p::"'a::idom poly" assumes bij:"bij_betw (\x. poly q x) A B" and "p\0" and q_deg: "degree q = 1" shows "card (proots_within p B) = card (proots_within (p \\<^sub>p q) A)" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by auto next case (no_proots p) have "proots_within p B = {}" using no_proots by auto moreover have "proots_within (p \\<^sub>p q) A = {}" using no_proots by (simp add: poly_pcompose proots_within_def) ultimately show ?case by auto next case (root b p) then have [simp]:"p\0" by auto have ?case when "b\B \ poly p b=0" proof - have "proots_within ([:- b, 1:] * p) B = proots_within p B" using that by auto moreover have "proots_within (([:- b, 1:] * p) \\<^sub>p q) A = proots_within (p \\<^sub>p q) A" using that unfolding pcompose_mult proots_within_times apply (auto simp add: poly_pcompose) using bij bij_betwE by blast ultimately show ?thesis using root.hyps[OF \p\0\] by auto qed moreover have ?case when "b\B" "poly p b\0" proof - define bb where "bb=[:- b, 1:]" have "card (proots_within (bb * p) B) = card {b} + card (proots_within p B)" proof - have "proots_within bb B = {b}" using that unfolding bb_def by auto then show ?thesis unfolding proots_within_times apply (subst card_Un_disjoint) by (use that in auto) qed also have "... = 1 + card (proots_within (p \\<^sub>p q) A)" using root.hyps by simp also have "... = card (proots_within ((bb * p) \\<^sub>p q) A)" unfolding proots_within_times pcompose_mult proof (subst card_Un_disjoint) obtain a where "b=poly q a" "a\A" by (metis \b \ B\ bij bij_betwE bij_betw_the_inv_into f_the_inv_into_f_bij_betw) define bbq pq where "bbq=bb \\<^sub>p q" and "pq=p \\<^sub>p q" have bbq_0:"poly bbq a=0" and bbq_deg: "degree bbq\1" and "bbq\0" unfolding bbq_def bb_def poly_pcompose subgoal using \b=poly q a\ by auto subgoal using q_deg by (simp add: degree_add_le degree_pcompose) subgoal using pcompose_eq_0 q_deg by fastforce done show "finite (proots_within bbq A)" using \bbq\0\ by simp show "finite (proots_within pq A)" unfolding pq_def by (metis \p \ 0\ finite_proots pcompose_eq_0 q_deg zero_less_one) have bbq_a:"proots_within bbq A = {a}" proof - have "a\proots_within bbq A" unfolding bb_def proots_within_def poly_pcompose bbq_def using \b=poly q a\ \a\A\ by simp moreover have "card (proots_within bbq A) = 1" proof - have "card (proots_within bbq A) \0" using \a\proots_within bbq A\ \finite (proots_within bbq A)\ by auto moreover have "card (proots_within bbq A) \ 1" by (meson \bbq \ 0\ card_proots_within_leq le_trans proots_count_leq_degree bbq_deg) ultimately show ?thesis by auto qed ultimately show ?thesis by (metis card_1_singletonE singletonD) qed show "proots_within (bbq) A \ proots_within (pq) A = {}" using bbq_a \b = poly q a\ that(2) unfolding pq_def by (simp add:poly_pcompose) show "1 + card (proots_within pq A) = card (proots_within bbq A) + card (proots_within pq A)" using bbq_a by simp qed finally show ?thesis unfolding bb_def . qed ultimately show ?case by auto qed end \ No newline at end of file diff --git a/thys/Count_Complex_Roots/More_Polynomials.thy b/thys/Count_Complex_Roots/More_Polynomials.thy --- a/thys/Count_Complex_Roots/More_Polynomials.thy +++ b/thys/Count_Complex_Roots/More_Polynomials.thy @@ -1,392 +1,392 @@ (* Author: Wenda Li *) section \More useful lemmas related polynomials\ theory More_Polynomials imports Winding_Number_Eval.Missing_Algebraic Winding_Number_Eval.Missing_Transcendental Sturm_Tarski.PolyMisc Budan_Fourier.BF_Misc begin subsection \More about @{term order}\ lemma order_normalize[simp]:"order x (normalize p) = order x p" by (metis dvd_normalize_iff normalize_eq_0_iff order_1 order_2 order_unique_lemma) lemma order_gcd: assumes "p\0" "q\0" shows "order x (gcd p q) = min (order x p) (order x q)" proof - define xx op oq where "xx=[:- x, 1:]" and "op = order x p" and "oq = order x q" obtain pp where pp:"p = xx ^ op * pp" "\ xx dvd pp" using order_decomp[OF \p\0\,of x,folded xx_def op_def] by auto obtain qq where qq:"q = xx ^ oq * qq" "\ xx dvd qq" using order_decomp[OF \q\0\,of x,folded xx_def oq_def] by auto define pq where "pq = gcd pp qq" have p_unfold:"p = (pq * xx ^ (min op oq)) * ((pp div pq) * xx ^ (op - min op oq))" and [simp]:"coprime xx (pp div pq)" and "pp\0" proof - have "xx ^ op = xx ^ (min op oq) * xx ^ (op - min op oq)" by (simp flip:power_add) moreover have "pp = pq * (pp div pq)" unfolding pq_def by simp ultimately show "p = (pq * xx ^ (min op oq)) * ((pp div pq) * xx ^ (op - min op oq))" unfolding pq_def pp by(auto simp:algebra_simps) show "coprime xx (pp div pq)" apply (rule prime_elem_imp_coprime[OF prime_elem_linear_poly[of 1 "-x",simplified],folded xx_def]) using \pp = pq * (pp div pq)\ pp(2) by auto qed (use pp \p\0\ in auto) have q_unfold:"q = (pq * xx ^ (min op oq)) * ((qq div pq) * xx ^ (oq - min op oq))" and [simp]:"coprime xx (qq div pq)" proof - have "xx ^ oq = xx ^ (min op oq) * xx ^ (oq - min op oq)" by (simp flip:power_add) moreover have "qq = pq * (qq div pq)" unfolding pq_def by simp ultimately show "q = (pq * xx ^ (min op oq)) * ((qq div pq) * xx ^ (oq - min op oq))" unfolding pq_def qq by(auto simp:algebra_simps) show "coprime xx (qq div pq)" apply (rule prime_elem_imp_coprime[OF prime_elem_linear_poly[of 1 "-x",simplified],folded xx_def]) using \qq = pq * (qq div pq)\ qq(2) by auto qed have "gcd p q=normalize (pq * xx ^ (min op oq))" proof - have "coprime (pp div pq * xx ^ (op - min op oq)) (qq div pq * xx ^ (oq - min op oq))" proof (cases "op>oq") case True then have "oq - min op oq = 0" by auto moreover have "coprime (xx ^ (op - min op oq)) (qq div pq)" by auto moreover have "coprime (pp div pq) (qq div pq)" apply (rule div_gcd_coprime[of pp qq,folded pq_def]) using \pp\0\ by auto ultimately show ?thesis by auto next case False then have "op - min op oq = 0" by auto moreover have "coprime (pp div pq) (xx ^ (oq - min op oq))" by (auto simp:coprime_commute) moreover have "coprime (pp div pq) (qq div pq)" apply (rule div_gcd_coprime[of pp qq,folded pq_def]) using \pp\0\ by auto ultimately show ?thesis by auto qed then show ?thesis unfolding p_unfold q_unfold apply (subst gcd_mult_left) by auto qed then have "order x (gcd p q) = order x pq + order x (xx ^ (min op oq))" apply simp apply (subst order_mult) using assms(1) p_unfold by auto also have "... = order x (xx ^ (min op oq))" using pp(2) qq(2) unfolding pq_def xx_def by (auto simp add: order_0I poly_eq_0_iff_dvd) also have "... = min op oq" unfolding xx_def by (rule order_power_n_n) also have "... = min (order x p) (order x q)" unfolding op_def oq_def by simp finally show ?thesis . qed lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n-1)) * pderiv p" apply (cases n) using pderiv_power_Suc by auto (*TODO: to replace the one (with the same name) in the library, as this version does not require 'a to be a field?*) lemma order_pderiv: fixes p::"'a::{idom,semiring_char_0} poly" assumes "p\0" "poly p x=0" shows "order x p = Suc (order x (pderiv p))" using assms proof - define xx op where "xx=[:- x, 1:]" and "op = order x p" have "op \0" unfolding op_def using assms order_root by blast obtain pp where pp:"p = xx ^ op * pp" "\ xx dvd pp" using order_decomp[OF \p\0\,of x,folded xx_def op_def] by auto have p_der:"pderiv p = smult (of_nat op) (xx^(op -1)) * pp + xx^op*pderiv pp" unfolding pp(1) by (auto simp:pderiv_mult pderiv_power xx_def algebra_simps pderiv_pCons) have "xx^(op -1) dvd (pderiv p)" unfolding p_der by (metis One_nat_def Suc_pred assms(1) assms(2) dvd_add dvd_mult_right dvd_triv_left neq0_conv op_def order_root power_Suc smult_dvd_cancel) moreover have "\ xx^op dvd (pderiv p)" proof assume "xx ^ op dvd pderiv p" then have "xx ^ op dvd smult (of_nat op) (xx^(op -1) * pp)" unfolding p_der by (simp add: dvd_add_left_iff) then have "xx ^ op dvd (xx^(op -1)) * pp" apply (elim dvd_monic[rotated]) using \op\0\ by (auto simp:lead_coeff_power xx_def) then have "xx ^ (op-1) * xx dvd (xx^(op -1))" using \\ xx dvd pp\ by (simp add: \op \ 0\ mult.commute power_eq_if) then have "xx dvd 1" using assms(1) pp(1) by auto then show False unfolding xx_def by (meson assms(1) dvd_trans one_dvd order_decomp) qed ultimately have "op - 1 = order x (pderiv p)" using order_unique_lemma[of x "op-1" "pderiv p",folded xx_def] \op\0\ by auto then show ?thesis using \op\0\ unfolding op_def by auto qed subsection \More about @{term rsquarefree}\ lemma rsquarefree_0[simp]: "\ rsquarefree 0" unfolding rsquarefree_def by simp lemma rsquarefree_times: assumes "rsquarefree (p*q)" shows "rsquarefree q" using assms proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) then have [simp]:"p\0" "q\0" "\a. order a p = 0" using order_0I by auto have "order a (p * q) = 0 \ order a q = 0" "order a (p * q) = 1 \ order a q = 1" for a subgoal by (subst order_mult) auto subgoal by (subst order_mult) auto done then show ?case using \rsquarefree (p * q)\ unfolding rsquarefree_def by simp next case (root a p) define pq aa where "pq = p * q" and "aa = [:- a, 1:]" have [simp]:"pq\0" "aa\0" "order a aa=1" subgoal using pq_def root.prems by auto subgoal by (simp add: aa_def) subgoal by (metis aa_def order_power_n_n power_one_right) done have "rsquarefree (aa * pq)" unfolding aa_def pq_def using root(2) by (simp add:algebra_simps) then have "rsquarefree pq" unfolding rsquarefree_def by (auto simp add:order_mult) from root(1)[OF this[unfolded pq_def]] show ?case . qed lemma rsquarefree_smult_iff: assumes "s\0" shows "rsquarefree (smult s p) \ rsquarefree p" unfolding rsquarefree_def using assms by (auto simp add:order_smult) lemma card_proots_within_rsquarefree: assumes "rsquarefree p" shows "proots_count p s = card (proots_within p s)" using assms proof (induct rule:poly_root_induct[of _ "\x. x\s"]) case 0 then have False by simp then show ?case by simp next case (no_roots p) then show ?case by (metis all_not_in_conv card_empty proots_count_def proots_within_iff sum.empty) next case (root a p) have "proots_count ([:a, - 1:] * p) s = 1 + proots_count p s" apply (subst proots_count_times) subgoal using root.prems rsquarefree_def by blast subgoal by (metis (no_types, hide_lams) add.inverse_inverse add.inverse_neutral minus_pCons proots_count_pCons_1_iff proots_count_uminus root.hyps(1)) done also have "... = 1 + card (proots_within p s)" proof - have "rsquarefree p" using \rsquarefree ([:a, - 1:] * p)\ by (elim rsquarefree_times) from root(2)[OF this] show ?thesis by simp qed also have "... = card (proots_within ([:a, - 1:] * p) s)" unfolding proots_within_times proof (subst card_Un_disjoint) have [simp]:"p\0" using root.prems by auto show "finite (proots_within [:a, - 1:] s)" "finite (proots_within p s)" by auto show " 1 + card (proots_within p s) = card (proots_within [:a, - 1:] s) + card (proots_within p s)" using \a \ s\ apply (subst proots_within_pCons_1_iff) by simp have "poly p a\0" proof (rule ccontr) assume "\ poly p a \ 0" then have "order a p >0" by (simp add: order_root) moreover have "order a [:a,-1:] = 1" by (metis (no_types, hide_lams) add.inverse_inverse add.inverse_neutral minus_pCons order_power_n_n order_uminus power_one_right) ultimately have "order a ([:a, - 1:] * p) > 1" apply (subst order_mult) subgoal using root.prems by auto subgoal by auto done then show False using \rsquarefree ([:a, - 1:] * p)\ unfolding rsquarefree_def using gr_implies_not0 less_not_refl2 by blast qed then show " proots_within [:a, - 1:] s \ proots_within p s = {}" using proots_within_pCons_1_iff(2) by auto qed finally show ?case . qed lemma rsquarefree_gcd_pderiv: - fixes p::"'a::{factorial_ring_gcd,semiring_char_0} poly" + fixes p::"'a::{factorial_ring_gcd,semiring_gcd_mult_normalize,semiring_char_0} poly" assumes "p\0" shows "rsquarefree (p div (gcd p (pderiv p)))" proof (cases "pderiv p = 0") case True have "poly (unit_factor p) x \0" for x using unit_factor_is_unit[OF \p\0\] by (meson assms dvd_trans order_decomp poly_eq_0_iff_dvd unit_factor_dvd) then have "order x (unit_factor p) = 0" for x using order_0I by blast then show ?thesis using True \p\0\ unfolding rsquarefree_def by simp next case False define q where "q = p div (gcd p (pderiv p))" have "q\0" unfolding q_def by (simp add: assms dvd_div_eq_0_iff) have order_pq:"order x p = order x q + min (order x p) (order x (pderiv p))" for x proof - have *:"p = q * gcd p (pderiv p)" unfolding q_def by simp show ?thesis apply (subst *) using \q\0\ \p\0\ \pderiv p\0\ by (simp add:order_mult order_gcd) qed have "order x q = 0 \ order x q=1" for x proof (cases "poly p x=0") case True from order_pderiv[OF \p\0\ this] have "order x p = order x (pderiv p) + 1" by simp then show ?thesis using order_pq[of x] by auto next case False then have "order x p = 0" by (simp add: order_0I) then have "order x q = 0" using order_pq[of x] by simp then show ?thesis by simp qed then show ?thesis using \q\0\ unfolding rsquarefree_def q_def by auto qed lemma poly_gcd_pderiv_iff: - fixes p::"'a::{semiring_char_0,factorial_ring_gcd} poly" + fixes p::"'a::{semiring_char_0,factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "poly (p div (gcd p (pderiv p))) x =0 \ poly p x=0" proof (cases "pderiv p=0") case True then obtain a where "p=[:a:]" using pderiv_iszero by auto then show ?thesis by (auto simp add: unit_factor_poly_def) next case False then have "p\0" using pderiv_0 by blast define q where "q = p div (gcd p (pderiv p))" have "q\0" unfolding q_def by (simp add: \p\0\ dvd_div_eq_0_iff) have order_pq:"order x p = order x q + min (order x p) (order x (pderiv p))" for x proof - have *:"p = q * gcd p (pderiv p)" unfolding q_def by simp show ?thesis apply (subst *) using \q\0\ \p\0\ \pderiv p\0\ by (simp add:order_mult order_gcd) qed have "order x q =0 \ order x p = 0" proof (cases "poly p x=0") case True from order_pderiv[OF \p\0\ this] have "order x p = order x (pderiv p) + 1" by simp then show ?thesis using order_pq[of x] by auto next case False then have "order x p = 0" by (simp add: order_0I) then have "order x q = 0" using order_pq[of x] by simp then show ?thesis using \order x p = 0\ by simp qed then show ?thesis apply (fold q_def) unfolding order_root using \p\0\ \q\0\ by auto qed subsection \Composition of a polynomial and a circular path\ lemma poly_circlepath_tan_eq: fixes z0::complex and r::real and p::"complex poly" defines "q1\ fcompose p [:(z0+r)*\,z0-r:] [:\,1:]" and "q2 \ [:\,1:] ^ degree p" assumes "0\t" "t\1" "t\1/2" shows "poly p (circlepath z0 r t) = poly q1 (tan (pi*t)) / poly q2 (tan (pi*t))" (is "?L = ?R") proof - have "?L = poly p (z0+ r*exp (2 * of_real pi * \ * t))" unfolding circlepath by simp also have "... = ?R" proof - define f where "f = (poly p \ (\x::real. z0 + r * exp (\ * x)))" have f_eq:"f t = ((\x::real. poly q1 x / poly q2 x) o (\x. tan (x/2)) ) t" when "cos (t / 2) \ 0" for t proof - have "f t = poly p (z0 + r * (cos t + \ * sin t)) " unfolding f_def exp_Euler by (auto simp add:cos_of_real sin_of_real) also have "... = poly p ((\x. ((z0-r)*x+(z0+r)*\) / (\+x)) (tan (t/2)))" proof - define tt where "tt=complex_of_real (tan (t / 2))" define rr where "rr = complex_of_real r" have "cos t = (1-tt*tt) / (1 + tt * tt)" "sin t = 2*tt / (1 + tt * tt)" unfolding sin_tan_half[of "t/2",simplified] cos_tan_half[of "t/2",OF that, simplified] tt_def by (auto simp add:power2_eq_square) moreover have "1 + tt * tt \ 0" unfolding tt_def apply (fold of_real_mult) by (metis (no_types, hide_lams) mult_numeral_1 numeral_One of_real_add of_real_eq_0_iff of_real_numeral sum_squares_eq_zero_iff zero_neq_one) ultimately have "z0 + r * ( (cos t) + \ * (sin t)) =(z0*(1+tt*tt)+rr*(1-tt*tt)+\*rr*2*tt ) / (1 + tt * tt) " apply (fold rr_def,simp add:add_divide_distrib) by (simp add:algebra_simps) also have "... = ((z0-rr)*tt+z0*\+rr*\) / (tt + \)" proof - have "tt + \ \ 0" using \1 + tt * tt \ 0\ by (metis i_squared neg_eq_iff_add_eq_0 square_eq_iff) then show ?thesis using \1 + tt * tt \ 0\ by (auto simp add:divide_simps algebra_simps) qed finally have "z0 + r * ( (cos t) + \ * (sin t)) = ((z0-rr)*tt+z0*\+rr*\) / (tt + \)" . then show ?thesis unfolding tt_def rr_def by (auto simp add:algebra_simps power2_eq_square) qed also have "... = (poly p o ((\x. ((z0-r)*x+(z0+r)*\) / (\+x)) o (\x. tan (x/2)) )) t" unfolding comp_def by (auto simp:tan_of_real) also have "... = ((\x::real. poly q1 x / poly q2 x) o (\x. tan (x/2)) ) t" unfolding q2_def q1_def apply (subst fcompose_poly[symmetric]) subgoal for x apply simp by (metis Re_complex_of_real add_cancel_right_left complex_i_not_zero imaginary_unit.sel(1) plus_complex.sel(1) rcis_zero_arg rcis_zero_mod) subgoal by (auto simp:tan_of_real algebra_simps) done finally show ?thesis . qed have "cos (pi * t) \0" unfolding cos_zero_iff_int2 proof assume "\i. pi * t = real_of_int i * pi + pi / 2" then obtain i where "pi * t = real_of_int i * pi + pi / 2" by auto then have "pi * t=pi * (real_of_int i + 1 / 2)" by (simp add:algebra_simps) then have "t=real_of_int i + 1 / 2" by auto then show False using \0\t\ \t\1\ \t\1/2\ by auto qed from f_eq[of "2*pi*t",simplified,OF this] show "?thesis" unfolding f_def comp_def by (auto simp add:algebra_simps) qed finally show ?thesis . qed end \ No newline at end of file diff --git a/thys/Hermite/Hermite.thy b/thys/Hermite/Hermite.thy --- a/thys/Hermite/Hermite.thy +++ b/thys/Hermite/Hermite.thy @@ -1,2128 +1,2133 @@ (* Title: Hermite.thy Author: Jose DivasĂ³n Author: JesĂºs Aransay *) section\Hermite Normal Form\ theory Hermite imports Echelon_Form.Echelon_Form_Inverse Echelon_Form.Examples_Echelon_Form_Abstract "HOL-Computational_Algebra.Euclidean_Algorithm" begin subsection\Some previous properties\ subsubsection\Rings\ subclass (in bezout_ring_div) euclidean_ring proof qed subsubsection\Polynomials\ lemma coeff_dvd_poly: "[:coeff a (degree a):] dvd (a::'a::{field} poly)" proof (cases "a=0") case True thus ?thesis unfolding dvd_def by simp next case False thus ?thesis by (intro dvd_trans[OF is_unit_triv one_dvd]) simp qed lemma poly_dvd_antisym2: fixes p q :: "'a::field poly" assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p div [:coeff p (degree p):] = q div [:coeff q (degree q):]" proof (cases "p = 0") case True thus ?thesis by (metis dvd1 dvd_0_left_iff) next case False have q: "q \ 0" by (metis False dvd2 dvd_0_left_iff) 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 a show ?thesis proof (cases a, auto split: if_splits, metis \a \ 0\) fix aa :: 'a assume a1: "aa \ 0" assume "q = smult aa p" have "[:aa * coeff p (degree p):] dvd smult aa p" using a1 by (metis (no_types) coeff_dvd_poly coeff_smult degree_smult_eq) thus "p div [:coeff p (degree p):] = smult aa p div [:aa * coeff p (degree p):]" using a1 by (simp add: False coeff_dvd_poly dvd_div_div_eq_mult) qed qed subsubsection\Units\ lemma unit_prod: assumes "finite S" shows "is_unit (prod (\i. U $ i $ i) S) = (\i\S. is_unit (U $ i $ i))" using assms proof (induct) case empty thus ?case by auto next case (insert a S) have "prod (\i. U $ i $ i) (insert a S) = U $ a $ a * prod (\i. U $ i $ i) S" by (simp add: insert.hyps(2)) thus ?case using is_unit_mult_iff insert.hyps by auto qed subsubsection\Upper triangular matrices\ lemma is_unit_diagonal: fixes U::"'a::{comm_ring_1, algebraic_semidom}^'n::{finite, wellorder}^'n::{finite, wellorder}" assumes U: "upper_triangular U" and det_U: "is_unit (det U)" shows "\i. is_unit (U $ i $ i)" proof - have "is_unit (prod (\i. U $ i $ i) UNIV)" using det_U det_upperdiagonal[of U] U unfolding upper_triangular_def by auto hence "\i\UNIV. is_unit (U $ i $ i)" using unit_prod[of UNIV U] by simp thus ?thesis by simp qed lemma upper_triangular_mult: fixes A::"'a::{semiring_1}^'n::{mod_type}^'n::{mod_type}" assumes A: "upper_triangular A" and B: "upper_triangular B" shows "upper_triangular (A**B)" proof (unfold upper_triangular_def matrix_matrix_mult_def, vector, auto) fix i j::'n assume ji: "jk\UNIV. A $ i $ k * B $ k $ j) = 0" proof (rule sum.neutral, clarify) fix x show "A $ i $ x * B $ x $ j = 0" proof (cases "xj" using ji by auto thus ?thesis using A B ji unfolding upper_triangular_def by auto qed qed qed lemma upper_triangular_adjugate: fixes A::"(('a::comm_ring_1,'n::{wellorder, finite}) vec, 'n) vec" assumes A: "upper_triangular A" shows "upper_triangular (adjugate A)" proof (auto simp add: cofactor_def upper_triangular_def adjugate_def transpose_def cofactorM_def) fix i j::'n assume ji: "j < i" with A show "det (minorM A j i) = 0" unfolding minorM_eq det_sq_matrix_eq[symmetric] from_vec_to_vec det_minor_row by (subst Square_Matrix.det_upperdiagonal) (auto simp: upd_row.rep_eq from_vec.rep_eq row_def axis_def upper_triangular_def intro!: prod_zero) qed lemma upper_triangular_inverse: fixes A::"(('a::{euclidean_semiring,comm_ring_1},'n::{wellorder, finite}) vec, 'n) vec" assumes A: "upper_triangular A" and inv_A: "invertible A" shows "upper_triangular (matrix_inv A)" using upper_triangular_adjugate[OF A] unfolding invertible_imp_matrix_inv[OF inv_A] unfolding matrix_scalar_mult_def upper_triangular_def by auto lemma upper_triangular_mult_diagonal: fixes A::"(('a::{semiring_1},'n::{wellorder, finite}) vec, 'n) vec" assumes A: "upper_triangular A" and B: "upper_triangular B" shows "(A**B) $ i $ i = A $ i $ i * B $ i $ i" proof - have UNIV_rw: "UNIV = (insert i (UNIV-{i}))" by auto have sum_0: "(\k\UNIV-{i}. A $ i $ k * B $ k $ i) = 0" proof (rule sum.neutral, rule) fix x assume x: "x \ UNIV - {i}" show "A $ i $ x * B $ x $ i = 0" proof (cases "xi" using x by auto thus ?thesis using B unfolding upper_triangular_def by auto qed qed have "(A**B) $ i $ i = (\k\UNIV. A $ i $ k * B $ k $ i)" unfolding matrix_matrix_mult_def by simp also have "... = (\k\(insert i (UNIV-{i})). A $ i $ k * B $ k $ i)" using UNIV_rw by simp also have "... = (A $ i $ i * B $ i $ i) + (\k\UNIV-{i}. A $ i $ k * B $ k $ i)" by (rule sum.insert, simp_all) finally show ?thesis unfolding sum_0 by simp qed subsubsection\More properties of mod type\ lemma add_left_neutral: fixes a::"'n::mod_type" shows "(a + b = a) = (b = 0)" by (auto, metis add_left_cancel monoid_add_class.add.right_neutral) lemma from_nat_1: "from_nat 1 = 1" unfolding from_nat_def o_def Abs'_def by (metis Rep_1 Rep_mod of_nat_1 one_def) subsubsection\Div and Mod\ lemma dvd_minus_eq_mod: fixes c::"'a::unique_euclidean_ring" assumes "c \ 0" and "c dvd a - b" shows "a mod c = b mod c" using assms dvd_div_mult_self[of c] by (metis add.commute diff_add_cancel mod_mult_self1) lemma eq_mod_dvd_minus: fixes c::"'a::unique_euclidean_ring" assumes "c \ 0" and "a mod c = b mod c" shows "c dvd a - b" using assms by (simp add: mod_eq_dvd_iff) lemma dvd_cong_not_eq_mod: fixes c::"'a::unique_euclidean_ring" assumes "xa mod c \ xb" and "c dvd xa mod c - xb" and "c \ 0" shows "xb mod c \ xb" using assms by (metis (no_types, lifting) diff_add_cancel dvdE mod_mod_trivial mod_mult_self4) lemma diff_mod_cong_0: fixes c::"'a::unique_euclidean_ring" assumes "xa mod c \ xb mod c" and" c dvd xa mod c - xb mod c" shows "c = 0" using assms dvd_cong_not_eq_mod mod_mod_trivial by blast lemma cong_diff_mod: fixes c::"'a::unique_euclidean_ring" assumes "xa \ xb" and "c dvd xa - xb" and "xa = xa mod c" shows "xb \ xb mod c" by (metis assms diff_eq_diff_eq diff_numeral_special(12) dvd_0_left dvd_minus_eq_mod) lemma exists_k_mod: fixes c::"'a::unique_euclidean_ring" shows "\k. a mod c = a + k * c" by (metis add.commute diff_add_cancel diff_minus_eq_add mult_div_mod_eq mult.commute mult_minus_left) subsection\Units, associated and congruent relations\ context semiring_1 begin definition "Units = {x::'a. (\k. 1 = x * k)}" end context ring_1 begin definition cong::"'a\'a\'a\bool" where "cong a c b = (\k. (a - c) = b * k)" lemma cong_eq: "cong a c b = (b dvd (a - c))" unfolding ring_1_class.cong_def dvd_def by simp end context normalization_semidom begin lemma Units_eq: "Units = {x. x dvd 1}" unfolding Units_def dvd_def .. lemma normalize_Units: "x \ Units \ normalize x = 1" by (intro is_unit_normalize) (simp_all add: Units_eq) lemma associated_eq: "(normalize a = normalize b) \ (\u\Units. a = u*b)" proof assume A: "normalize a = normalize b" show "\u\Units. a = u*b" proof (cases "a = 0 \ b = 0") case False from A have "a = (unit_factor a div unit_factor b) * b" by (metis mult_not_zero normalize_0 normalize_mult_unit_factor mult.left_commute unit_div_mult_self unit_factor_is_unit) moreover from False have "unit_factor a div unit_factor b \ Units" by (simp add: Units_eq unit_div) ultimately show ?thesis by blast next case True with A normalize_eq_0_iff[of a] normalize_eq_0_iff[of b] have "a = 0" "b = 0" by auto thus ?thesis by (auto intro!: exI[of _ 1] simp: Units_def) qed -qed (auto simp: normalize_Units normalize_mult) +qed (auto simp: normalize_Units Units_def) end context unique_euclidean_ring begin definition "associated_rel = {(a,b). normalize a = normalize b}" lemma equiv_associated: shows "equiv UNIV associated_rel" unfolding associated_rel_def equiv_def refl_on_def sym_def trans_def by simp definition "congruent_rel b = {(a,c). cong a c b}" lemma relf_congruent_rel: "refl (congruent_rel b)" unfolding refl_on_def congruent_rel_def unfolding cong_def by (auto, metis mult_zero_right) lemma sym_congruent_rel: "sym (congruent_rel b)" unfolding sym_def congruent_rel_def unfolding cong_def by (auto, metis add_commute add_minus_cancel diff_conv_add_uminus minus_mult_commute mult.left_commute mult_1_left) lemma trans_congruent_rel: "trans (congruent_rel b)" unfolding trans_def congruent_rel_def unfolding cong_def by (auto, metis add_assoc diff_add_cancel diff_conv_add_uminus distrib_left) lemma equiv_congruent: "equiv UNIV (congruent_rel b)" unfolding equiv_def using relf_congruent_rel sym_congruent_rel trans_congruent_rel by auto end subsection\Associates and residues functions\ context normalization_semidom begin definition ass_function :: "('a \ 'a) \ bool" where "ass_function f = ((\a. normalize a = normalize (f a)) \ pairwise (\a b. normalize a \ normalize b) (range f))" definition "Complete_set_non_associates S = (\f. ass_function f \ f`UNIV = S \ (pairwise (\a b. normalize a \ normalize b) S))" end context ring_1 begin definition res_function :: "('a \ 'a \ 'a) \ bool" where "res_function f = (\c. (\a b. cong a b c \ f c a = f c b) \ pairwise (\a b. \ cong a b c) (range (f c)) \ (\a. \k. f c a = a + k*c))" definition "Complete_set_residues g = (\f. res_function f \ (\c. (pairwise (\a b. \ cong a b c) (f c`UNIV)) \ g c = f c`UNIV))" end lemma ass_function_Complete_set_non_associates: assumes f: "ass_function f" shows "Complete_set_non_associates (f`UNIV)" unfolding Complete_set_non_associates_def ass_function_def apply (rule exI[of _ f]) using f unfolding ass_function_def unfolding pairwise_def by fast lemma in_Ass_not_associated: assumes Ass_S: "Complete_set_non_associates S" and x: "x\S" and y: "y\S" and x_not_y: "x\y" shows "normalize x \ normalize y" using assms unfolding Complete_set_non_associates_def pairwise_def by auto lemma ass_function_0: assumes r: "ass_function ass" shows "(ass x = 0) = (x = 0)" using assms unfolding ass_function_def pairwise_def by (metis normalize_eq_0_iff) lemma ass_function_0': assumes r: "ass_function ass" shows "(ass x div x = 0) = (x=0)" - using assms unfolding ass_function_def pairwise_def - by (metis ass_function_0 associatedD2 div_self div_by_0 dvd_normalize_div - normalize_0 normalize_1 one_neq_zero r) - +proof safe + assume *: "ass x div x = 0" + from r have **: "normalize (ass x) = normalize x" + by (simp add: ass_function_def) + from associatedD2[OF this] have "x dvd ass x" + by simp + with * ** show "x = 0" + by (auto simp: dvd_div_eq_0_iff) +qed auto lemma res_function_Complete_set_residues: assumes f: "res_function f" shows "Complete_set_residues (\c. (f c)`UNIV)" unfolding Complete_set_residues_def apply (rule exI[of _ f]) using f unfolding res_function_def by blast lemma in_Res_not_congruent: assumes res_g: "Complete_set_residues g" and x: "x \ g b" and y: "y \ g b" and x_not_y: "x\y" shows "\ cong x y b" using assms unfolding Complete_set_residues_def unfolding pairwise_def by auto subsubsection\Concrete instances in Euclidean rings\ definition "ass_function_euclidean (p::'a::{normalization_euclidean_semiring, euclidean_ring}) = normalize p" definition "res_function_euclidean b (n::'a::{euclidean_ring}) = (if b = 0 then n else (n mod b))" lemma ass_function_euclidean: "ass_function ass_function_euclidean" unfolding ass_function_def image_def ass_function_euclidean_def pairwise_def by auto lemma res_function_euclidean: "res_function (res_function_euclidean :: 'a :: unique_euclidean_ring \ _)" by (auto simp add: pairwise_def res_function_def cong_eq image_def res_function_euclidean_def dvd_minus_eq_mod) (auto simp add: dvd_cong_not_eq_mod eq_mod_dvd_minus diff_mod_cong_0 cong_diff_mod exists_k_mod) subsubsection\Concrete case of the integer ring\ definition "ass_function_int (n::int) = abs n" lemma ass_function_int: "ass_function_int = ass_function_euclidean" by (unfold ass_function_int_def ass_function_euclidean_def) simp lemma ass_function_int_UNIV: "(ass_function_int` UNIV) = {x. x\0}" unfolding ass_function_int_def image_def by (auto, metis abs_of_nonneg) subsection\Definition of Hermite Normal Form\ text\ It is worth noting that there is not a single definition of Hermite Normal Form in the literature. For instance, some authors restrict their definitions to the case of square nonsingular matrices. Other authors just work with integer matrices. Furthermore, given a matrix $A$ its Hermite Normal Form $H$ can be defined to be upper triangular or lower triangular. In addition, the transformation from $A$ to $H$ can be made by means of elementary row operations or elementary column operations. In our case, we will work as general as possible, so our input will be any matrix (including nonsquare ones). The output will be an upper triangular matrix obtained by means of elementary row operations. Hence, given a complete set of nonassociates and a complete set of residues, $H$ is said to be in Hermite Normal Form if: \begin{enumerate} \item H is in Echelon Form \item The first nonzero element of a nonzero row belongs to the complete set of nonassociates \item Let $h$ be the first nonzero element of a nonzero row. Then each element above $h$ belongs to the corresponding complete set of residues of $h$ \end{enumerate} A matrix $H$ is the Hermite Normal Form of a matrix $A$ if: \begin{enumerate} \item There exists an invertible matrix $P$ such that $A = PH$ \item H is in Hermite Normal Form \end{enumerate} The Hermite Normal Form is usually applied to integer matrices. As we have already said, there is no one single definition of it, so some authors impose different conditions. In the particular case of integer matrices, leading coefficients (the first nonzero element of a nonzero row) are usually required to be positive, but it is also possible to impose them to be negative since we would only have to multiply by $-1$. In the case of the elements $h_{ik}$ above a leading coefficient $h_{ij}$, some authors demand $0 \leq h_{ik} < h_{ij}$, other ones impose the conditions $h_{ik} \leq 0$ and \mbox{$\mid h_{ik} \mid < h_{ij}$}, and other ones $- \frac{h_{ij}}{2} < h_{ik} \leq \frac{h_{ij}}{2}$. More different options are also possible. All the possibilities can be represented selecting a complete set of nonassociates and a complete set of residues. The algorithm to compute the Hermite Normal Form will be parameterised by functions which obtain the appropriate leading coefficient and the suitable elements above them. We can execute the algorithm with different functions to get exactly which Hermite Normal Form we want. Once we fix such a complete set of nonassociates and the corresponding complete set of residues, the Hermite Normal Form is unique. \ subsubsection\Echelon form up to row k\ text\We present the definition of echelon form up to a row k (not included).\ definition "echelon_form_upt_row A k = ( (\i. to_nat i < k \ is_zero_row i A \ \ (\j. j>i \ to_nat j < k \ \ is_zero_row j A)) \ (\i j. i < j \ to_nat j < k \ \ is_zero_row i A \ \ is_zero_row j A \ (LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)) )" lemma echelon_form_upt_row_condition1_explicit: assumes "echelon_form_upt_row A k" and "to_nat i < k" and "is_zero_row i A" shows "\ (\j. j>i \ to_nat j < k \ \ is_zero_row j A)" using assms unfolding echelon_form_upt_row_def by blast lemma echelon_form_upt_row_condition1_explicit': assumes "echelon_form_upt_row A k" and "to_nat i < k" and "is_zero_row i A" and "i\j" and "to_nat j < k" shows "is_zero_row j A" proof (cases "i=j") case True thus ?thesis using assms by auto next case False thus ?thesis using assms unfolding echelon_form_upt_row_def by simp qed lemma echelon_form_upt_row_condition1_explicit_neg: assumes "echelon_form_upt_row A k" and iA: "\ is_zero_row i A" and ia_i: "ia < i" and i: "to_nat i < k" shows "\ is_zero_row ia A" proof - have "to_nat ia < k" by (metis ia_i i less_trans to_nat_mono) thus ?thesis using assms unfolding echelon_form_upt_row_def by blast qed lemma echelon_form_upt_row_condition2_explicit: assumes "echelon_form_upt_row A k" and "ia < j" and "to_nat j < k" and "\ is_zero_row ia A" and "\ is_zero_row j A" shows "(LEAST n. A $ ia $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using assms unfolding echelon_form_upt_row_def by auto lemma echelon_form_upt_row_intro: assumes"(\i. to_nat i < k \ is_zero_row i A \ \ (\j. i to_nat j < k \ \ is_zero_row j A))" and "(\i j. i < j \ to_nat j < k \ \ is_zero_row i A \ \ is_zero_row j A \ (LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0))" shows "echelon_form_upt_row A k" using assms unfolding echelon_form_upt_row_def by simp lemma echelon_form_echelon_form_upt_row: "echelon_form A = echelon_form_upt_row A (nrows A)" by (simp add: to_nat_less_card echelon_form_def echelon_form_upt_row_def ncols_def nrows_def echelon_form_upt_k_def is_zero_row_upt_k_def is_zero_row_def) subsubsection\Hermite Normal Form up to row k\ text\Predicate to check if a matrix is in Hermite Normal form up to row k (not included).\ definition "Hermite_upt_row A k associates residues = ( Complete_set_non_associates associates \ Complete_set_residues residues \ echelon_form_upt_row A k \ (\i. to_nat i < k \ \ is_zero_row i A \ A $ i $ (LEAST n. A $ i $ n \ 0) \ associates) \ (\i. to_nat i < k \ \ is_zero_row i A \ (\j 0) \ residues (A $ i $ (LEAST n. A $ i $ n \ 0)))) )" text\The definition of Hermite Normal Form is now introduced:\ definition Hermite::"'a::{bezout_ring_div,normalization_semidom} set \ ('a \ 'a set) \ (('a, 'b::{mod_type}) vec, 'c::{mod_type}) vec \ bool" where "Hermite associates residues A = ( Complete_set_non_associates associates \ (Complete_set_residues residues) \ echelon_form A \ (\i. \ is_zero_row i A \ A $ i $ (LEAST n. A $ i $ n \ 0) \ associates) \ (\i. \ is_zero_row i A \ (\j. j A $ j $ (LEAST n. A $ i $ n \ 0) \ residues (A $ i $ (LEAST n. A $ i $ n \ 0)))) )" lemma Hermite_Hermite_upt_row: "Hermite ass res A = Hermite_upt_row A (nrows A) ass res" by (simp add: Hermite_def Hermite_upt_row_def to_nat_less_card is_zero_row_def nrows_def ncols_def echelon_form_echelon_form_upt_row) lemma Hermite_intro: assumes "Complete_set_non_associates associates" and "Complete_set_residues residues" and "echelon_form A " and "(\i. \ is_zero_row i A \ A $ i $ (LEAST n. A $ i $ n \ 0) \ associates)" and "(\i. \ is_zero_row i A \ (\j. j A $ j $ (LEAST n. A $ i $ n \ 0) \ residues (A $ i $ (LEAST n. A $ i $ n \ 0))))" shows "Hermite associates residues A" using assms unfolding Hermite_def by simp subsection\Definition of an algorithm to compute the Hermite Normal Form\ text\ The algorithm is parameterised by three functions: \begin{itemize} \item The function that computes de B\'ezout identity (necessary to compute the echelon form). \item The function that given an element, it returns its representative element in the associated equivalent class, which will be an element in the complete set of nonassociates. \item The function that given two elements $a$ and $b$, it returns its representative element in the congruent equivalent class of $b$, which will be an element in the complete set of residues of $b$. \end{itemize} \ primrec Hermite_reduce_above :: "'a::unique_euclidean_ring^'cols::mod_type^'rows::mod_type\nat \'rows\'cols\ ('a\'a\'a) \ 'a^'cols::mod_type^'rows::mod_type" where "Hermite_reduce_above A 0 i j res = A" | "Hermite_reduce_above A (Suc n) i j res = (let i'=((from_nat n)::'rows); Aij = A $ i $ j; Ai'j = A $ i' $ j in Hermite_reduce_above (row_add A i' i (((res Aij (Ai'j)) - (Ai'j)) div Aij)) n i j res)" definition "Hermite_of_row_i ass res A i = ( if is_zero_row i A then A else let j = (LEAST n. A $ i $ n \ 0); Aij= (A $ i $ j); A' = mult_row A i ((ass Aij) div Aij) in Hermite_reduce_above A' (to_nat i) i j res)" definition "Hermite_of_upt_row_i A i ass res = foldl (Hermite_of_row_i ass res) A (map from_nat [0..Proving the correctness of the algorithm\ subsubsection\The proof\ lemma Hermite_reduce_above_preserves: assumes n: "n\to_nat a" shows "(Hermite_reduce_above A n i j res) $ a $ b = A $ a $ b" using n proof (induct n arbitrary: A) case 0 thus ?case by simp next case (Suc n) thus ?case by (auto simp add: Let_def row_add_def) (metis Suc_le_eq from_nat_mono from_nat_to_nat_id less_irrefl to_nat_less_card) qed lemma Hermite_reduce_above_works: assumes n: "n \ to_nat i" and a: "to_nat a < n" shows "(Hermite_reduce_above A n i j res) $ a $ b = row_add A a i ((res (A$i$j) (A$a$j) - (A$a$j)) div (A$i$j)) $ a $ b" using n a proof (induct n arbitrary: A) case 0 thus ?case by (simp add: row_add_def) next case (Suc n) define A' where "A' = row_add A (from_nat n) i ((res (A $ i $ j) (A $ from_nat n $ j) - A $ from_nat n $ j) div A $ i $ j)" have n: "n < nrows A" unfolding nrows_def by (metis Suc.prems(1) Suc_le_eq less_trans to_nat_less_card) show ?case proof (cases "to_nat a = n") case False have a_less_n: "to_nat a < n" by (metis False Suc.prems(2) less_antisym) have "Hermite_reduce_above A (Suc n) i j res $ a $ b = Hermite_reduce_above A' n i j res $ a $ b " by (simp add: Let_def A'_def) also have "... = row_add A' a i ((res (A' $ i $ j) (A' $ a $ j) - A' $ a $ j) div A' $ i $ j) $ a $ b" by (rule Suc.hyps[OF _ a_less_n], simp add: Suc.prems(1) Suc_leD) also have "... = row_add A a i ((res (A $ i $ j) (A $ a $ j) - A $ a $ j) div A $ i $ j) $ a $ b" unfolding row_add_def A'_def using a_less_n Suc.prems n to_nat_from_nat_id[OF n[unfolded nrows_def]] by auto finally show ?thesis . next case True hence a_eq_fn_n: "a = from_nat n" by auto have "Hermite_reduce_above A (Suc n) i j res $ a $ b = Hermite_reduce_above A' n i j res $ a $ b " by (simp add: Let_def A'_def) also have "... = A' $ a $ b" by (rule Hermite_reduce_above_preserves, simp add: True) finally show ?thesis unfolding A'_def a_eq_fn_n . qed qed lemma Hermite_of_row_preserves_below: assumes i_a: "i 0)" and not_zero_i_A: "\ is_zero_row i A" and e: "echelon_form A" shows "(Hermite_of_row_i ass res A i) $ a $ b = A $ a $ b" proof (auto simp add: Hermite_of_row_i_def Let_def) let ?n="(LEAST n. A $ i $ n \ 0)" let ?M="(mult_row A i (ass (A $ i $ ?n) div A $ i $ ?n))" let ?H="Hermite_reduce_above ?M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" have Aib: " A $ i $ b = 0" by (metis (mono_tags) b not_less_Least) show "?H $ a $ b = A $ a $ b" proof (cases "a\i") case True have "?H $ a $ b = ?M $ a $ b" by (rule Hermite_reduce_above_preserves) (metis True to_nat_mono') also have "... = A $ a $ b" using Aib unfolding mult_row_def by auto finally show ?thesis . next let ?R="row_add ?M a i ((res (?M $ i $ ?n) (?M $ a $ ?n) - ?M $ a $ ?n) div ?M $ i $ ?n)" case False hence ia: "i>a" by simp have "?H $ a $ b = ?R $ a $ b" by (rule Hermite_reduce_above_works, auto simp add: ia to_nat_mono) also have "... = A $ a $ b" using ia Aib unfolding row_add_def mult_row_def by auto finally show ?thesis . qed qed lemma echelon_form_Hermite_of_condition1: fixes res ass i A defines M: "M \ mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes e: "echelon_form A" and a: "ass_function ass" and not_zero_iA: "\ is_zero_row i A" and zero_ia_H: "is_zero_row ia H" and ia_j: "ia < j" shows "is_zero_row j H" proof (cases "is_zero_row ia A") case True have zero_jA: "is_zero_row j A" by (metis True e echelon_form_condition1 ia_j) have ij: "i 0" by (metis (mono_tags, lifting) LeastI is_zero_row_def is_zero_row_upt_k_def not_zero_ia_A) show ?thesis proof (cases "i \ ia") case True have "H $ ia $ ?n = M $ ia $ ?n" unfolding H by (rule Hermite_reduce_above_preserves, simp add: True to_nat_mono') also have "... \ 0" unfolding M mult_row_def using A_ia_n ass_function_0'[OF a] by auto finally have "H $ ia $ ?n \ 0" . hence not_zero_ia_H: "\ is_zero_row ia H" unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto thus ?thesis using zero_ia_H by contradiction next case False let ?m="(LEAST m. A $ i $ m \ 0)" let ?R="row_add M ia i ((res (M $ i $ ?m) (M $ ia $ ?m) - M $ ia $ ?m) div M $ i $ ?m)" have ia_less_i: "ia 0" using ia_less_i A_im A_ia_n unfolding row_add_def M mult_row_def by auto finally have "H $ ia $ ?n \ 0" . hence not_zero_ia_H: "\ is_zero_row ia H" unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto thus ?thesis using zero_ia_H by contradiction qed qed lemma row_zero_A_imp_row_zero_H: fixes res ass i A defines M: "M \ mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes e: "echelon_form A" and not_zero_iA: "\ is_zero_row i A" and zero_j_A: "is_zero_row j A" shows "is_zero_row j H" proof (auto simp add: is_zero_row_def is_zero_row_upt_k_def ncols_def) fix a have A_ja: "A $ j $ a = 0" using zero_j_A by (simp add: is_zero_row_def is_zero_row_upt_k_def ncols_def) show "H $ j $ a = 0" proof (cases "i\j") case True have "H $ j $ a = M $ j $ a" unfolding H by (rule Hermite_reduce_above_preserves, simp add: True to_nat_mono') also have "... = 0" unfolding M mult_row_def using True A_ja by auto finally show ?thesis . next let ?n="(LEAST n. A $ i $ n \ 0)" let ?R="row_add M j i ((res (M $ i $ ?n) (M $ j $ ?n) - M $ j $ ?n) div M $ i $ ?n)" case False hence ji: "j mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes i_ia: "i is_zero_row ia H" shows "(LEAST n. A $ ia $ n \ 0) = (LEAST n. H $ ia $ n \ 0)" proof (rule Least_equality) let ?n="(LEAST n. H $ ia $ n \ 0)" have "A $ ia $ ?n = M $ ia $ ?n" unfolding M mult_row_def using i_ia by auto also have "... = H $ ia $ ?n " unfolding H by (rule Hermite_reduce_above_preserves[symmetric]) (metis i_ia dual_order.strict_iff_order to_nat_mono') also have "... \ 0" by (metis (mono_tags) LeastI is_zero_row_def' not_zero_ia_H) finally show "A $ ia $ (LEAST n. H $ ia $ n \ 0) \ 0" . next fix y assume A_ia_y: "A $ ia $ y \ 0" have "H $ ia $ y = M $ ia $ y" unfolding H by (rule Hermite_reduce_above_preserves) (metis i_ia dual_order.strict_iff_order to_nat_mono') also have "... \ 0" unfolding M mult_row_def using i_ia A_ia_y by auto finally show "(LEAST n. H $ ia $ n \ 0) \ y" by (rule Least_le) qed lemma Hermite_reduce_above_Least_eq: fixes res ass i A defines M: "M \ mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes a: "ass_function ass" and not_zero_iA: "\ is_zero_row i A" shows "(LEAST n. A $ i $ n \ 0) = (LEAST n. H $ i $ n \ 0)" proof (rule Least_equality[symmetric]) let ?n="(LEAST n. A $ i $ n \ 0)" have Ain: "A $ i $ ?n \ 0" by (metis (mono_tags, lifting) LeastI is_zero_row_def' not_zero_iA) have "H $ i $ ?n = M $ i $ ?n" unfolding H by (rule Hermite_reduce_above_preserves, simp) also have "... \ 0" unfolding M mult_row_def by (auto simp add: Ain ass_function_0'[OF a]) finally show "H $ i $ ?n \ 0" . fix y assume H_iy: "H $ i $ y \ 0" show "(LEAST n. A $ i $ n \ 0) \ y" proof (rule Least_le, rule ccontr, simp) assume Aiy: "A $ i $ y = 0" have "H $ i $ y = M $ i $ y" unfolding H by (rule Hermite_reduce_above_preserves, simp) also have "... = (ass (A $ i $ ?n) div A $ i $ ?n) * A $ i $ y" unfolding M mult_row_def by auto also have "... = 0" unfolding Aiy by simp finally show False using H_iy by contradiction qed qed lemma Hermite_reduce_above_Least_eq_ge: fixes res ass i A defines M: "M \ mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes e: "echelon_form A" and not_zero_iA: "\ is_zero_row i A" and not_zero_ia_A: "\ is_zero_row ia A" and not_zero_ia_H: "\ is_zero_row ia H" and ia_less_i: "ia < i" shows "(LEAST n. H $ ia $ n \ 0) = (LEAST n. A $ ia $ n \ 0)" proof - let ?least_H = "(LEAST n. H $ ia $ n \ 0)" let ?least_A = "(LEAST n. A $ ia $ n \ 0)" let ?n= "(LEAST n. A $ i $ n \ 0)" let ?Ain ="A $ i $ ?n" let ?R="row_add M ia i ((res (M $ i $ ?n) (M $ ia $ ?n) - M $ ia $ ?n) div M $ i $ ?n)" have A_ia_least_A: "A $ ia $ ?least_A \ 0" by (metis (mono_tags, lifting) LeastI is_zero_row_def' not_zero_ia_A) have H_ia_least_H: "H $ ia $ ?least_H \ 0" by (metis (mono_tags, lifting) LeastI is_zero_row_def' not_zero_ia_H) have A_i_least_ia_0: "A $ i $ (LEAST n. A $ ia $ n \ 0) = 0" proof - have "(LEAST n. A $ ia $ n \ 0) < (LEAST n. A $ i $ n \ 0)" using e echelon_form_condition1 echelon_form_condition2_explicit ia_less_i not_zero_iA by blast thus ?thesis using not_less_Least by blast qed have H_ia_least_A: "H $ ia $ ?least_A \ 0" proof - have "H $ ia $ ?least_A = ?R $ ia $ ?least_A" unfolding H by (rule Hermite_reduce_above_works, simp_all add: ia_less_i to_nat_mono) also have "... \ 0" using ia_less_i unfolding row_add_def M mult_row_def by (auto simp add: A_i_least_ia_0 A_ia_least_A) finally show ?thesis . qed hence least_H_le_least_A: "?least_H \ ?least_A" by (metis (mono_tags) not_less not_less_Least) have A_i_least_H: "A $ i $ ?least_H = 0" proof - have "(LEAST n. A $ ia $ n \ 0) < (LEAST n. A $ i $ n \ 0)" using e echelon_form_condition1 echelon_form_condition2_explicit ia_less_i not_zero_iA by blast thus ?thesis using not_less_Least least_H_le_least_A by (metis (mono_tags) dual_order.strict_trans2) qed have "A $ ia $ ?least_H \ 0" proof - have ia_not_i: "ia \ i" using ia_less_i by simp have "?R $ ia $ ?least_H = H $ ia $ ?least_H" unfolding H by (rule Hermite_reduce_above_works[symmetric], simp_all add: ia_less_i to_nat_mono) also have "... \ 0" by (rule H_ia_least_H) finally have R_ia_least_H: "?R $ ia $ ?least_H \ 0" . hence "A $ ia $ ?least_H + (res (ass (?Ain) div ?Ain * ?Ain) (A $ ia $ (LEAST n. A $ i $ n \ 0)) - A $ ia $ (LEAST n. A $ i $ n \ 0)) div (ass (?Ain) div ?Ain * ?Ain) * (ass (?Ain) div ?Ain * A $ i $ ?least_H) \ 0" using ia_not_i unfolding row_add_def M mult_row_def by auto thus ?thesis using ia_less_i A_i_least_H unfolding row_add_def M mult_row_def by auto qed hence least_A_le_least_H: "?least_A \ ?least_H" by (metis (poly_guards_query) Least_le) show ?thesis using least_A_le_least_H least_H_le_least_A by simp qed lemma Hermite_reduce_above_Least: fixes res ass i A defines M: "M \ mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes e: "echelon_form A" and a: "ass_function ass" and not_zero_iA: "\ is_zero_row i A" and not_zero_ia_A: "\ is_zero_row ia A" and not_zero_ia_H: "\ is_zero_row ia H" shows "(LEAST n. H $ ia $ n \ 0) = (LEAST n. A $ ia $ n \ 0)" proof (cases "iaia" by simp show ?thesis proof (cases "ia=i") case True show ?thesis unfolding True H M by (rule Hermite_reduce_above_Least_eq[symmetric, OF a not_zero_iA]) next case False hence i_ia: "i mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" defines H: "H \ Hermite_reduce_above M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" assumes e: "echelon_form A" and a: "ass_function ass" and not_zero_iA: "\ is_zero_row i A" and ia_less_j: "ia < j" and not_zero_ia_H: "\ is_zero_row ia H" and not_zero_j_H: "\ is_zero_row j H" shows "(LEAST n. H $ ia $ n \ 0) < (LEAST n. H $ j $ n \ 0)" proof - let ?n= "(LEAST n. A $ i $ n \ 0)" have Ain: "A $ i $ ?n \ 0" by (metis (mono_tags) LeastI is_zero_row_def' not_zero_iA) have not_zero_j_A: "\ is_zero_row j A" using row_zero_A_imp_row_zero_H[OF e not_zero_iA] not_zero_j_H unfolding H M by blast have not_zero_ia_A: "\ is_zero_row ia A" using row_zero_A_imp_row_zero_H[OF e not_zero_iA] not_zero_ia_H unfolding H M by blast have Least_le_A: "(LEAST n. A $ ia $ n \ 0) < (LEAST n. A $ j $ n \ 0)" by (rule echelon_form_condition2_explicit[OF e ia_less_j not_zero_ia_A not_zero_j_A]) show ?thesis proof (cases "i 0) = (LEAST n. H $ ia $ n \ 0)" unfolding H M by (rule Hermite_reduce_above_Least_eq_le[OF True], metis H M not_zero_ia_H) moreover have Least_A_ia_H_ia: "(LEAST n. A $ j $ n \ 0) = (LEAST n. H $ j $ n \ 0)" unfolding H M by (rule Hermite_reduce_above_Least_eq_le[OF ij], metis H M not_zero_j_H) ultimately show ?thesis using Least_le_A by simp next case False hence ia_le_i: "ia\i" by simp show ?thesis proof (cases "i=ia") case True thus ?thesis using Hermite_reduce_above_Least_eq[OF a not_zero_iA] Least_le_A using Hermite_reduce_above_Least_eq_le[OF ia_less_j] using not_zero_j_H unfolding H M by fastforce next case False hence ia_less_i: "ia 0) = (LEAST n. A $ ia $ n \ 0)" unfolding H M by (rule Hermite_reduce_above_Least_eq_ge[OF e not_zero_iA not_zero_ia_A _ ia_less_i]) (metis H M not_zero_ia_H) show ?thesis proof (cases "j 0) = (LEAST n. A $ j $ n \ 0)" unfolding H M by (rule Hermite_reduce_above_Least_eq_ge[OF e not_zero_iA not_zero_j_A _ True]) (metis H M not_zero_j_H) thus ?thesis by (simp add: Least_H_ia_A_ia Least_le_A) next case False hence j_ge_i: "j\i" by auto show ?thesis proof (cases "j=i") case True have "(LEAST n. H $ j $ n \ 0) = (LEAST n. A $ j $ n \ 0)" unfolding H M using Hermite_reduce_above_Least_eq True a not_zero_iA by fastforce thus ?thesis by (simp add: Least_H_ia_A_ia Least_le_A) next case False hence j_sg_i: "j>i" using j_ge_i by simp have "(LEAST n. H $ j $ n \ 0) = (LEAST n. A $ j $ n \ 0)" unfolding H M by (rule Hermite_reduce_above_Least_eq_le[symmetric, OF j_sg_i]) (metis H M not_zero_j_H) thus ?thesis by (simp add: Least_H_ia_A_ia Least_le_A) qed qed qed qed qed lemma echelon_form_Hermite_of_row: assumes a: "ass_function ass" and "res_function res" and e: "echelon_form A" shows "echelon_form (Hermite_of_row_i ass res A i)" proof (rule echelon_form_intro, auto simp add: Hermite_of_row_i_def Let_def) fix ia j assume "is_zero_row i A" and "is_zero_row ia A" and "ia < j" thus "is_zero_row j A" using echelon_form_condition1[OF e] by blast next fix ia j assume "is_zero_row i A" and "ia < j" and "\ is_zero_row ia A" and "\ is_zero_row j A" thus "(LEAST n. A $ ia $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using echelon_form_condition2[OF e] by blast next fix ia j assume "\ is_zero_row i A" and "is_zero_row ia (Hermite_reduce_above (mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))) (to_nat i) i (LEAST n. A $ i $ n \ 0) res)" and "ia < j" thus "is_zero_row j (Hermite_reduce_above (mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))) (to_nat i) i (LEAST n. A $ i $ n \ 0) res)" using echelon_form_Hermite_of_condition1[OF e a] by blast next fix ia j let ?H="(Hermite_reduce_above (mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))) (to_nat i) i (LEAST n. A $ i $ n \ 0) res)" assume "\ is_zero_row i A" and "ia < j" and "\ is_zero_row ia ?H" and "\ is_zero_row j ?H" thus "(LEAST n. ?H $ ia $ n \ 0) < (LEAST n. ?H $ j $ n \ 0)" using echelon_form_Hermite_of_condition2[OF e a] by blast qed lemma echelon_form_fold_Hermite_of_row_i: assumes e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" shows "echelon_form (foldl (Hermite_of_row_i ass res) A (map from_nat [0.. is_zero_row i A" shows "(Hermite_of_row_i ass res A i) $ i $ (LEAST n. (Hermite_of_row_i ass res A i) $ i $ n \ 0) \ range ass" unfolding Hermite_of_row_i_def using not_zero_i_A proof (auto simp add: Let_def) let ?M="(mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))) " let ?H="Hermite_reduce_above ?M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" let ?Ain="A $ i $ (LEAST n. A $ i $ n \ 0)" have Ain: "?Ain \ 0" by (metis (mono_tags) LeastI is_zero_row_def' not_zero_i_A) have least_eq: "(LEAST n. ?H $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" by (rule Hermite_reduce_above_Least_eq[OF a not_zero_i_A, symmetric]) have "?H $ i $ (LEAST n. ?H $ i $ n \ 0) = ?M $ i $ (LEAST n. ?H $ i $ n \ 0)" by (rule Hermite_reduce_above_preserves, simp) also have "... = ass (?Ain) div ?Ain * ?Ain" unfolding mult_row_def least_eq[unfolded mult_row_def] by simp also have "... = ass ?Ain" proof (rule dvd_div_mult_self) show "?Ain dvd ass ?Ain" using a unfolding ass_function_def by (simp add: associatedD2) qed also have "... \ range ass" by simp finally show "?H $ i $ (LEAST n. ?H $ i $ n \ 0) \ range ass" . qed lemma Hermite_of_upt_row_preserves_below: assumes i: "to_nat a\k" shows "Hermite_of_upt_row_i A k ass res $ a $ b = A $ a $ b" using i proof (induct k) case 0 thus ?case unfolding Hermite_of_upt_row_i_def by auto next case (Suc k) show ?case by (simp add: Hermite_of_upt_row_i_def, metis Hermite_of_upt_row_i_def Hermite_of_row_preserves_below Suc.hyps Suc.prems Suc_leD Suc_le_eq from_nat_mono from_nat_to_nat_id to_nat_less_card) qed lemma not_zero_Hermite_reduce_above: fixes ass i A defines M: "M\(mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0)))" assumes not_zero_a_A: "\ is_zero_row a A" and not_zero_i_A: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and n: "n \ to_nat i" shows "\ is_zero_row a (Hermite_reduce_above M n i (LEAST n. A $ i $ n \ 0) res)" proof - let ?H = "(Hermite_reduce_above M n i (LEAST n. A $ i $ n \ 0) res)" let ?n="LEAST n. A $ a $ n \ 0" let ?m="LEAST n. A $ i $ n \ 0" have Aan: "A $ a $ ?n \ 0" by (metis (mono_tags) LeastI not_zero_a_A is_zero_row_def') have Aim: "A $ i $ ?m \ 0" by (metis (mono_tags) LeastI not_zero_i_A is_zero_row_def') show ?thesis proof (cases "n\to_nat a") case True have "?H $ a $ ?n = M $ a $ ?n" by (metis Hermite_reduce_above_preserves True) also have "... \ 0" unfolding M mult_row_def using ass_function_0'[OF a] Aan by auto finally show ?thesis unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto next let ?R="row_add M a i ((res (M $ i $ ?m) (M $ a $ ?m) - M $ a $ ?m) div M $ i $ ?m)" case False hence a_n: "to_nat a < n" by simp have ai: "a < i" by (metis False dual_order.trans n nat_less_le not_less_iff_gr_or_eq to_nat_mono) have "(LEAST n. A $ a $ n \ 0) < ?m" by (rule echelon_form_condition2_explicit[OF e ai not_zero_a_A not_zero_i_A]) hence Ain: "A $ i $ (LEAST n. A $ a $ n \ 0) = 0" by (metis (full_types) not_less_Least) have a_not_i: "a \ i" by (metis False n) have "?H $ a $ ?n = ?R $ a $ ?n" by (rule Hermite_reduce_above_works[OF n a_n]) also have "... \ 0" using a_not_i Aan Ain unfolding row_add_def M mult_row_def by auto finally show ?thesis unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto qed qed lemma Least_Hermite_of_row_i: assumes i: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" shows "(LEAST n. Hermite_of_row_i ass res A i $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" proof - let ?M="mult_row A i (ass (A $ i $ (LEAST n. A $ i $ n \ 0)) div A $ i $ (LEAST n. A $ i $ n \ 0))" let ?H="Hermite_reduce_above ?M (to_nat i) i (LEAST n. A $ i $ n \ 0) res" have "(LEAST n. Hermite_of_row_i ass res A i $ i $ n \ 0) = (LEAST n. ?H $ i $ n \ 0)" using i unfolding Hermite_of_row_i_def unfolding Let_def by auto also have "... = (LEAST n. A $ i $ n \ 0)" by (rule Hermite_reduce_above_Least[OF e a i i]) (rule not_zero_Hermite_reduce_above[OF i i e a], simp) finally show ?thesis . qed lemma Least_Hermite_of_row_i2: assumes i: "\ is_zero_row i A" and k: "\ is_zero_row k A" and e: "echelon_form A" and a: "ass_function ass" shows "(LEAST n. Hermite_of_row_i ass res A k $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" proof - let ?M="mult_row A k (ass (A $ k $ (LEAST n. A $ k $ n \ 0)) div A $ k $ (LEAST n. A $ k $ n \ 0))" let ?H="Hermite_reduce_above ?M (to_nat k) k (LEAST n. A $ k $ n \ 0) res" have "(LEAST n. Hermite_of_row_i ass res A k $ i $ n \ 0) = (LEAST n. ?H $ i $ n \ 0)" using k unfolding Hermite_of_row_i_def unfolding Let_def by auto also have "... = (LEAST n. A $ i $ n \ 0)" by (rule Hermite_reduce_above_Least[OF e a k i]) (simp add: a e i k not_zero_Hermite_reduce_above) finally show ?thesis . qed lemma Hermite_of_row_i_works: fixes i A ass defines n:"n \(LEAST n. A $ i $ n \ 0)" defines M:"M \ (mult_row A i (ass (A $ i $ n) div A $ i $ n))" assumes ai: "a < i" and i: "\ is_zero_row i A" shows "Hermite_of_row_i ass res A i $ a $ b = row_add M a i ((res (M $ i $ n) (M $ a $ n) - M $ a $ n) div M $ i $ n) $ a $ b" proof - let ?H="Hermite_reduce_above M (to_nat i) i n res" have "Hermite_of_row_i ass res A i $ a $ b = ?H $ a $ b" unfolding Hermite_of_row_i_def Let_def M n using i by simp also have "... = row_add M a i ((res (M $ i $ n) (M $ a $ n) - M $ a $ n) div M $ i $ n) $ a $ b" by (rule Hermite_reduce_above_works, auto simp add: ai to_nat_mono) finally show ?thesis . qed lemma Hermite_of_row_i_works2: fixes i A ass defines n:"n \(LEAST n. A $ i $ n \ 0)" defines M:"M \ (mult_row A i (ass (A $ i $ n) div A $ i $ n))" assumes i: "\ is_zero_row i A" shows "Hermite_of_row_i ass res A i $ i $ b = M $ i $ b" proof - let ?H="Hermite_reduce_above M (to_nat i) i n res" have "Hermite_of_row_i ass res A i $ i $ b = ?H $ i $ b" unfolding Hermite_of_row_i_def Let_def M n using i by simp also have "... = M $ i $ b" by (rule Hermite_reduce_above_preserves, simp) finally show ?thesis . qed lemma Hermite_of_upt_row_preserves_nonzero_rows_ge: assumes i: "\ is_zero_row i A" and i2: "to_nat i\k" shows "\ is_zero_row i (Hermite_of_upt_row_i A k ass res)" proof - let ?n="LEAST n. A $ i $ n \ 0" have Ain: "A $ i $ ?n \ 0" by (metis (mono_tags) LeastI i is_zero_row_def') have "Hermite_of_upt_row_i A k ass res $ i $ ?n = A $ i $ ?n" by (rule Hermite_of_upt_row_preserves_below[OF i2]) also have "... \ 0" by (metis (mono_tags) LeastI i is_zero_row_def') finally have "Hermite_of_upt_row_i A k ass res $ i $ (LEAST n. A $ i $ n \ 0) \ 0" . thus ?thesis unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto qed lemma Hermite_of_upt_row_i_Least_ge: assumes i: "\ is_zero_row i A" and i2: "to_nat i\k" shows "(LEAST n. Hermite_of_upt_row_i A k ass res $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" proof (rule Least_equality) let ?n="LEAST n. A $ i $ n \ 0" have Ain: "A $ i $ ?n \ 0" by (metis (mono_tags) LeastI i is_zero_row_def') have "Hermite_of_upt_row_i A k ass res $ i $ ?n = A $ i $ ?n" by (rule Hermite_of_upt_row_preserves_below[OF i2]) also have "... \ 0" by (metis (mono_tags) LeastI i is_zero_row_def') finally show "Hermite_of_upt_row_i A k ass res $ i $ (LEAST n. A $ i $ n \ 0) \ 0" . fix y assume H: "Hermite_of_upt_row_i A k ass res $ i $ y \ 0" show "(LEAST n. A $ i $ n \ 0) \ y" by (rule Least_le, metis Hermite_of_upt_row_preserves_below H i2) qed lemma Hermite_of_upt_row_i_Least: assumes iA: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" and k: "k\nrows A" shows "(LEAST n. Hermite_of_upt_row_i A k ass res $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" proof (cases "to_nat i\k") case True thus ?thesis using Hermite_of_upt_row_i_Least_ge iA by blast next case False hence i_less_k: "to_nat inrows A" using Suc.prems unfolding nrows_def by simp have k2: "k is_zero_row i A'" unfolding A'_def2 by (rule Hermite_of_upt_row_preserves_nonzero_rows_ge, auto simp add: Suc.prems True) have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_row_i ass res A' (from_nat k)" unfolding Hermite_of_upt_row_i_def A'_def by auto also have "(LEAST n. ... $ i $ n \ 0) = (LEAST n. A' $ i $ n \ 0)" unfolding i_fn_k by (rule Least_Hermite_of_row_i[OF not_zero_i_A' e a]) also have "... = (LEAST n. A $ i $ n \ 0)" unfolding A'_def2 by (rule Hermite_of_upt_row_i_Least_ge, auto simp add: True Suc.prems) finally show ?thesis . next case False hence i_less_k: "to_nat i < k" using Suc.prems by simp hence i_less_k2: "i < from_nat k" by (metis from_nat_mono from_nat_to_nat_id k2 nrows_def) show ?thesis proof (cases "is_zero_row (from_nat k) A'") case True have H: "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_upt_row_i A k ass res" using True by (simp add: Hermite_of_upt_row_i_def Hermite_of_row_i_def A'_def Let_def ) show ?thesis unfolding H by (rule Suc.hyps[OF i_less_k], auto simp add: Suc.prems k) next case False have not_zero_i_A': "\ is_zero_row i A'" using e False i_less_k2 echelon_form_condition1 by blast have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_row_i ass res A' (from_nat k)" unfolding Hermite_of_upt_row_i_def A'_def by auto also have "(LEAST n. ... $ i $ n \ 0) = (LEAST n. A' $ i $ n \ 0)" by (rule Least_Hermite_of_row_i2[OF not_zero_i_A' False e a]) also have "... = (LEAST n. A $ i $ n \ 0)" unfolding A'_def2 by (rule Suc.hyps[OF i_less_k], auto simp add: Suc.prems k) finally show ?thesis . qed qed qed qed lemma Hermite_of_upt_row_preserves_nonzero_rows: assumes i: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" and k: "k\nrows A" shows "\ is_zero_row i (Hermite_of_upt_row_i A k ass res)" proof - let ?n="LEAST n. A $ i $ n \ 0" have Ain: "A $ i $ ?n \ 0" by (metis (mono_tags) LeastI i is_zero_row_def') show ?thesis proof (cases "to_nat i\k") case True thus ?thesis using Hermite_of_upt_row_preserves_nonzero_rows_ge i by blast next case False hence i_less_k: "to_nat i nrows A" using Suc.prems unfolding nrows_def by auto have k_nrows2: "k < nrows A" using Suc.prems unfolding nrows_def by auto define A' where "A' = foldl (Hermite_of_row_i ass res) A (map from_nat [0.. 0) = (LEAST n. A $ i $ n \ 0)" unfolding A'_def2 by (rule Hermite_of_upt_row_i_Least[OF _ e a r], auto simp add: k_nrows Suc.prems) have e: "echelon_form A'" unfolding A'_def2 by (simp add: a e echelon_form_Hermite_of_upt_row_i r) show ?case proof (cases "to_nat i = k") let ?M="mult_row A' i (ass (A' $ i $ (LEAST n. A' $ i $ n \ 0)) div A' $ i $ (LEAST n. A' $ i $ n \ 0))" case True hence fn_k_i: "from_nat k = i" by (metis from_nat_to_nat_id) have not_zero_i_A': "\ is_zero_row i A'" by (unfold A'_def2) (rule Hermite_of_upt_row_preserves_nonzero_rows_ge, auto simp add: True Suc.prems) have A'_i_l: "(A' $ i $ (LEAST n. A' $ i $ n \ 0)) \ 0" by (metis (mono_tags) LeastI is_zero_row_def' not_zero_i_A') have "Hermite_of_upt_row_i A (Suc k) ass res $ i $ ?n = Hermite_of_row_i ass res A' (from_nat k) $ i $ ?n" unfolding Hermite_of_upt_row_i_def A'_def by simp also have "... = ?M $ i $ ?n" unfolding fn_k_i by (rule Hermite_of_row_i_works2[OF not_zero_i_A']) also have "... \ 0" using A'_i_l unfolding mult_row_def by (simp add: ass_function_0'[OF a] least_A'_A) finally show ?thesis unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto next case False hence i_k: "to_nat i < k" by (metis Suc.prems(1) less_antisym) hence i_k2: "i< from_nat k" using i_k Suc.prems by (metis from_nat_mono from_nat_to_nat_id k_nrows2 nrows_def) have not_zero_i_A': "\ is_zero_row i A'" unfolding A'_def2 by (rule Suc.hyps[OF i_k Suc.prems(2) k_nrows]) show ?thesis proof (cases "is_zero_row (from_nat k) A'") case False have Ain: "A' $ i $ (LEAST n. A $ i $ n \ 0) \ 0" unfolding least_A'_A[symmetric] by (metis (mono_tags) LeastI is_zero_row_def' not_zero_i_A') have Akn: "A' $ (from_nat k) $ (LEAST n. A $ i $ n \ 0) = 0" proof - have "(LEAST n. A' $ i $ n \ 0) < (LEAST n. A' $ (from_nat k) $ n \ 0)" by (rule echelon_form_condition2_explicit[OF e i_k2 not_zero_i_A' False]) thus ?thesis by (metis (mono_tags) least_A'_A not_less_Least) qed let ?m="(LEAST n. A' $ from_nat k $ n \ 0)" let ?M="mult_row A' (from_nat k) (ass (A' $ from_nat k $ ?m) div A' $ from_nat k $ ?m)" have "Hermite_of_upt_row_i A (Suc k) ass res $ i $ ?n = Hermite_of_row_i ass res A' (from_nat k) $ i $ ?n" unfolding Hermite_of_upt_row_i_def A'_def by simp also have "... = row_add (mult_row A' (from_nat k) (ass (A' $ from_nat k $ ?m) div A' $ from_nat k $ ?m)) i (from_nat k) ((res (?M $ from_nat k $ ?m) (?M $ i $ ?m) - ?M $ i $ ?m) div ?M $ from_nat k $ ?m) $ i $ ?n" by (rule Hermite_of_row_i_works[OF i_k2 False]) also have "... \ 0" using i_k2 Ain Akn unfolding row_add_def mult_row_def by auto finally show ?thesis unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto next case True have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_upt_row_i A k ass res" using True by (simp add: Hermite_of_upt_row_i_def Hermite_of_row_i_def A'_def Let_def) thus ?thesis using not_zero_i_A' unfolding A'_def2 by simp qed qed qed qed qed lemma Hermite_of_upt_row_i_in_range: fixes k ass res assumes not_zero_i_A: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" and k: "to_nat inrows A" shows "Hermite_of_upt_row_i A k ass res $ i $ (LEAST n. A $ i $ n \ 0) \ range ass" using k not_zero_i_A k2 proof (induct k) case 0 thus ?case by auto next case (Suc k) have k: "k\nrows A" using Suc.prems unfolding nrows_def by simp have k2: "k 0)) div A' $ i $ (LEAST n. A' $ i $ n \ 0))" have not_zero_A': "\ is_zero_row i A'" using Hermite_of_upt_row_preserves_nonzero_rows[OF not_zero_i_A e a r k] unfolding A'_def Hermite_of_upt_row_i_def by simp have e_A': "echelon_form A'" by (metis A'_def a e echelon_form_fold_Hermite_of_row_i r) have least_eq: "(LEAST n. (Hermite_of_row_i ass res A' i) $ i $ n \ 0) = (LEAST n. A' $ i $ n \ 0)" by (rule Least_Hermite_of_row_i[OF not_zero_A' e_A' a]) have least_eq2: "(LEAST n. A' $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" unfolding A'_def2 by (rule Hermite_of_upt_row_i_Least[OF not_zero_i_A e a r k]) show ?case proof (cases "to_nat i = k") case True have fn_k_i: "from_nat k = i" by (metis True from_nat_to_nat_id) have "(Hermite_of_upt_row_i A (Suc k) ass res) $ i $ (LEAST n. A $ i $ n \ 0) = (Hermite_of_row_i ass res A' i) $ i $ (LEAST n. A $ i $ n \ 0)" unfolding Hermite_of_upt_row_i_def by (simp add: A'_def fn_k_i) also have "... = (Hermite_of_row_i ass res A' i) $ i $ (LEAST n. (Hermite_of_row_i ass res A' i) $ i $ n \ 0)" unfolding least_eq least_eq2 .. also have "... \ range ass" by (rule in_ass_Hermite_of_row[OF a r not_zero_A']) finally show ?thesis . next case False hence i_less_k: "to_nat i < k" using Suc.prems by auto hence i_less_k2: "i < from_nat k" using Suc.prems by (metis from_nat_mono from_nat_to_nat_id k2 nrows_def) show ?thesis proof (cases "is_zero_row (from_nat k) A'") case True have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_upt_row_i A k ass res" using True by (simp add: Hermite_of_upt_row_i_def Hermite_of_row_i_def A'_def Let_def ) thus ?thesis using Suc.hyps not_zero_i_A k i_less_k by auto next case False have "(Hermite_of_upt_row_i A (Suc k) ass res) $ i $ (LEAST n. A $ i $ n \ 0) = (Hermite_of_row_i ass res A' (from_nat k)) $ i $ (LEAST n. A $ i $ n \ 0)" unfolding Hermite_of_upt_row_i_def A'_def by auto also have "... = A' $ i $ (LEAST n. A $ i $ n \ 0)" proof (rule Hermite_of_row_preserves_previous_cols[OF _ False e_A']) show "(LEAST n. A $ i $ n \ 0) < (LEAST n. A' $ mod_type_class.from_nat k $ n \ 0)" unfolding least_eq2[symmetric] by (rule echelon_form_condition2_explicit[OF e_A' i_less_k2 not_zero_A' False]) qed also have "... \ range ass" unfolding A'_def using Suc.prems Suc.hyps unfolding Hermite_of_upt_row_i_def using i_less_k by auto finally show ?thesis . qed qed qed lemma Hermite_of_upt_row_preserves_zero_rows_ge: assumes i: "is_zero_row i A" and k: "k \ nrows A" and ik: "to_nat i\k" shows "is_zero_row i (Hermite_of_upt_row_i A k ass res)" proof (unfold is_zero_row_def', clarify) fix j have "Hermite_of_upt_row_i A k ass res $ i $ j = A $ i $ j" by (metis Hermite_of_upt_row_preserves_below ik) also have "... = 0" using i unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by simp finally show "Hermite_of_upt_row_i A k ass res $ i $ j = 0" . qed lemma Hermite_of_upt_row_preserves_zero_rows: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes i: "is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" and k: "k \ nrows A" shows "is_zero_row i (Hermite_of_upt_row_i A k ass res)" proof (cases "to_nat i\k") case True show ?thesis by (rule Hermite_of_upt_row_preserves_zero_rows_ge[OF i k True]) next case False hence i_k: "to_nat i < k" by simp show ?thesis using k i_k proof (induct k) case 0 thus ?case unfolding Hermite_of_upt_row_i_def by (simp add: i) next case (Suc k) have k: "k\nrows A" using Suc.prems unfolding nrows_def by auto have k2: "k is_zero_row i (Hermite_of_upt_row_i A k ass res)" using echelon_form_condition1 by (metis A'_def2 False e a r echelon_form_Hermite_of_upt_row_i i_less_k2) ultimately show ?thesis by contradiction qed qed qed qed lemma Hermite_of_preserves_zero_rows: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes i: "is_zero_row i (echelon_form_of A bezout)" and a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" shows "is_zero_row i (Hermite_of A ass res bezout)" unfolding Hermite_of_def Let_def by (rule Hermite_of_upt_row_preserves_zero_rows[OF i echelon_form_echelon_form_of[OF b] a r]) (auto simp add: nrows_def) lemma Hermite_of_Least: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes i: "\ is_zero_row i (Hermite_of A ass res bezout)" and a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" shows "(LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) = (LEAST n. (echelon_form_of A bezout) $ i $ n \ 0)" proof - have non_zero_i_eA: "\ is_zero_row i (echelon_form_of A bezout)" using Hermite_of_preserves_zero_rows[OF _ a r b] i by auto have e: "echelon_form (echelon_form_of A bezout)" by (rule echelon_form_echelon_form_of[OF b]) show ?thesis unfolding Hermite_of_def Let_def by (rule Hermite_of_upt_row_i_Least[OF non_zero_i_eA e a r], simp add: nrows_def) qed lemma in_associates_Hermite_of: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" and i: "\ is_zero_row i (Hermite_of A ass res bezout)" shows "Hermite_of A ass res bezout $ i $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) \ range ass" proof - have non_zero_i_eA: "\ is_zero_row i (echelon_form_of A bezout)" using Hermite_of_preserves_zero_rows[OF _ a r b] i by auto have e: "echelon_form (echelon_form_of A bezout)" by (rule echelon_form_echelon_form_of[OF b]) have least: "(LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) = (LEAST n. (echelon_form_of A bezout) $ i $ n \ 0)" by (rule Hermite_of_Least[OF i a r b]) show ?thesis unfolding least unfolding Hermite_of_def Let_def by (rule Hermite_of_upt_row_i_in_range[OF non_zero_i_eA e a r]) (auto simp add: to_nat_less_card nrows_def) qed lemma Hermite_of_row_i_range_res: assumes ji: "j is_zero_row i A" and r: "res_function res" shows "Hermite_of_row_i ass res A i $ j $ (LEAST n. A $ i $ n \ 0) \ range (res (Hermite_of_row_i ass res A i $ i $ (LEAST n. A $ i $ n \ 0)))" proof - let ?n="(LEAST n. A $ i $ n \ 0)" define M where "M = mult_row A i (ass (A $ i $ ?n) div A $ i $ ?n)" let ?R="row_add M j i ((res (M $ i $ ?n) (M $ j $ ?n) - M $ j $ ?n) div M $ i $ ?n)" have Hii: "Hermite_of_row_i ass res A i $ i $ ?n = M $ i $ ?n" unfolding M_def by (rule Hermite_of_row_i_works2[OF not_zero_i_A]) have rw: "Hermite_of_row_i ass res A i $ j $ ?n = ?R $ j $ ?n" unfolding M_def by (rule Hermite_of_row_i_works[OF ji not_zero_i_A]) show ?thesis proof - have "\b ba. \bb. ba + bb * b = res b ba" using r unfolding res_function_def by metis thus ?thesis using rw unfolding image_def Hii row_add_def by auto (metis (lifting) add_diff_cancel_left' nonzero_mult_div_cancel_left mult.commute mult_eq_0_iff) qed qed lemma Hermite_of_upt_row_i_in_range_res: fixes k ass res assumes not_zero_i_A: "\ is_zero_row i A" and e: "echelon_form A" and a: "ass_function ass" and r: "res_function res" and k: "to_nat inrows A" and j: "j 0) \ range (res (Hermite_of_upt_row_i A k ass res $ i $ (LEAST n. A $ i $ n \ 0)))" using k not_zero_i_A k2 proof (induct k) case 0 thus ?case by auto next case (Suc k) let ?n="(LEAST n. A $ i $ n \ 0)" have k: "k\nrows A" using Suc.prems unfolding nrows_def by simp have k2: "k 0)) div A' $ i $ (LEAST n. A' $ i $ n \ 0))" have not_zero_A': "\ is_zero_row i A'" using Hermite_of_upt_row_preserves_nonzero_rows[OF not_zero_i_A e a r k] unfolding A'_def Hermite_of_upt_row_i_def by simp have e_A': "echelon_form A'" by (metis A'_def a e echelon_form_fold_Hermite_of_row_i r) have least_eq: "(LEAST n. (Hermite_of_row_i ass res A' i) $ i $ n \ 0) = (LEAST n. A' $ i $ n \ 0)" by (rule Least_Hermite_of_row_i[OF not_zero_A' e_A' a]) have least_eq2: "(LEAST n. A' $ i $ n \ 0) = (LEAST n. A $ i $ n \ 0)" unfolding A'_def2 by (rule Hermite_of_upt_row_i_Least[OF not_zero_i_A e a r k]) show ?case proof (cases "to_nat i = k") case True have fn_k_i: "from_nat k = i" by (metis True from_nat_to_nat_id) have H_rw: "(Hermite_of_upt_row_i A (Suc k) ass res $ i $ (LEAST n. A $ i $ n \ 0)) = (Hermite_of_row_i ass res A' i $ i $ (LEAST n. A' $ i $ n \ 0))" by (simp add: Hermite_of_upt_row_i_def A'_def fn_k_i least_eq2[unfolded A'_def]) have "(Hermite_of_upt_row_i A (Suc k) ass res) $ j $ (LEAST n. A $ i $ n \ 0) = (Hermite_of_row_i ass res A' i) $ j $ (LEAST n. A $ i $ n \ 0)" unfolding Hermite_of_upt_row_i_def by (simp add: A'_def fn_k_i) also have "... = (Hermite_of_row_i ass res A' i) $ j $ (LEAST n. A' $ i $ n \ 0)" unfolding least_eq2 .. also have "... \ range (res (Hermite_of_row_i ass res A' i $ i $ (LEAST n. A' $ i $ n \ 0)))" by (rule Hermite_of_row_i_range_res[OF j not_zero_A' r]) also have "... = range ((res (Hermite_of_upt_row_i A (Suc k) ass res $ i $ (LEAST n. A $ i $ n \ 0))))" unfolding H_rw .. finally show ?thesis . next case False hence i_less_k: "to_nat i < k" using Suc.prems by auto hence i_less_k2: "i < from_nat k" using Suc.prems by (metis from_nat_mono from_nat_to_nat_id k2 nrows_def) show ?thesis proof (cases "is_zero_row (from_nat k) A'") case True have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_upt_row_i A k ass res" using True by (simp add: Hermite_of_upt_row_i_def Hermite_of_row_i_def A'_def Let_def ) thus ?thesis using Suc.hyps not_zero_i_A k i_less_k by auto next case False have H_rw: "(Hermite_of_upt_row_i A (Suc k) ass res $ i $ (LEAST n. A $ i $ n \ 0)) = A' $ i $ (LEAST n. A $ i $ n \ 0)" proof (auto simp add: Hermite_of_upt_row_i_def A'_def[symmetric], rule Hermite_of_row_preserves_previous_cols[OF _ False e_A']) have "(LEAST n. A' $ i $ n \ 0) < (LEAST n. A' $ mod_type_class.from_nat k $ n \ 0)" by (rule echelon_form_condition2_explicit[OF e_A' i_less_k2 not_zero_A' False]) thus "(LEAST n. A $ i $ n \ 0) < (LEAST n. A' $ mod_type_class.from_nat k $ n \ 0)" unfolding least_eq2 . qed have "(Hermite_of_upt_row_i A (Suc k) ass res) $ j $ (LEAST n. A $ i $ n \ 0) = (Hermite_of_row_i ass res A' (from_nat k)) $ j $ (LEAST n. A $ i $ n \ 0)" unfolding Hermite_of_upt_row_i_def A'_def by auto also have "... = A' $ j $ (LEAST n. A $ i $ n \ 0)" proof (rule Hermite_of_row_preserves_previous_cols[OF _ False e_A']) show "(LEAST n. A $ i $ n \ 0) < (LEAST n. A' $ mod_type_class.from_nat k $ n \ 0)" unfolding least_eq2[symmetric] by (rule echelon_form_condition2_explicit[OF e_A' i_less_k2 not_zero_A' False]) qed also have "... \ range (res (Hermite_of_upt_row_i A k ass res $ i $ (LEAST n. A $ i $ n \ 0)))" unfolding A'_def2 by (rule Suc.hyps[OF i_less_k], auto simp add: Suc.prems k) also have "... = range (res (Hermite_of_upt_row_i A (Suc k) ass res $ i $ (LEAST n. A $ i $ n \ 0)))" unfolding H_rw A'_def2 .. finally show ?thesis . qed qed qed lemma in_residues_Hermite_of: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" and i: "\ is_zero_row i (Hermite_of A ass res bezout)" and ji: "j < i" shows "Hermite_of A ass res bezout $ j $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) \ range (res (Hermite_of A ass res bezout $ i $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0)))" proof - have non_zero_i_eA: "\ is_zero_row i (echelon_form_of A bezout)" using Hermite_of_preserves_zero_rows[OF _ a r b] i by auto have e: "echelon_form (echelon_form_of A bezout)" by (rule echelon_form_echelon_form_of[OF b]) have least: "(LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) = (LEAST n. (echelon_form_of A bezout) $ i $ n \ 0)" by (rule Hermite_of_Least[OF i a r b]) show ?thesis unfolding least unfolding Hermite_of_def Let_def by (rule Hermite_of_upt_row_i_in_range_res[OF non_zero_i_eA e a r _ _ ji]) (auto simp add: to_nat_less_card nrows_def) qed lemma Hermite_Hermite_of: assumes a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" shows "Hermite (range ass) (\c. range (res c)) (Hermite_of A ass res bezout)" proof (rule Hermite_intro, auto) show "Complete_set_non_associates (range ass)" by (simp add: ass_function_Complete_set_non_associates a) show "Complete_set_residues (\c. range (res c))" by (simp add: r res_function_Complete_set_residues) show "echelon_form (Hermite_of A ass res bezout)" by (simp add: a b echelon_form_Hermite_of r) fix i assume i: "\ is_zero_row i (Hermite_of A ass res bezout)" show "Hermite_of A ass res bezout $ i $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) \ range ass" by (rule in_associates_Hermite_of[OF a r b i]) next fix i j assume i: "\ is_zero_row i (Hermite_of A ass res bezout)" and j: "j < i" show "Hermite_of A ass res bezout $ j $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0) \ range (res (Hermite_of A ass res bezout $ i $ (LEAST n. Hermite_of A ass res bezout $ i $ n \ 0)))" by (rule in_residues_Hermite_of[OF a r b i j]) qed subsubsection\Proving that the Hermite Normal Form is computed by means of elementary operations\ lemma invertible_Hermite_reduce_above: assumes n: "n \ to_nat i" shows "\P. invertible P \ Hermite_reduce_above A n i j res = P ** A" using n proof (induct n arbitrary: A) case 0 thus ?case by (auto, metis invertible_def matrix_mul_lid) next case (Suc n) let ?R="(row_add A (from_nat n) i ((res (A $ i $ j) (A $ from_nat n $ j) - A $ from_nat n $ j) div A $ i $ j))" obtain Q where inv_Q: "invertible Q" and H_QR: "Hermite_reduce_above ?R n i j res = Q ** ?R" using Suc.hyps Suc.prems by auto let ?P="(row_add (mat 1) (from_nat n) i ((res (A $ i $ j) (A $ from_nat n $ j) - A $ from_nat n $ j) div A $ i $ j))" have inv_P: "invertible ?P" proof (rule invertible_row_add) show "mod_type_class.from_nat n \ i" by (metis Suc.prems Suc_le_eq add_to_nat_def from_nat_mono less_irrefl monoid_add_class.add.right_neutral to_nat_0 to_nat_less_card) qed have inv_QP: "invertible (Q ** ?P)" by (metis inv_P inv_Q invertible_mult) have "Hermite_reduce_above A (Suc n) i j res = Hermite_reduce_above ?R n i j res" by (auto simp add: Let_def) also have "... = Q ** ?R" unfolding H_QR .. also have "... = Q ** (?P ** A)" by (subst row_add_mat_1[symmetric], rule refl) also have "... = (Q ** ?P) ** A" by (simp add: matrix_mul_assoc) finally show ?case using inv_QP by auto qed lemma invertible_Hermite_of_row_i: assumes a: "ass_function ass" shows "\P. invertible P \ Hermite_of_row_i ass res A i = P ** A" unfolding Hermite_of_row_i_def proof (auto simp add: Let_def, metis invertible_def matrix_mul_lid) let ?n="LEAST n. A $ i $ n \ 0" let ?M="mult_row A i (ass (A $ i $ ?n) div A $ i $ ?n)" let ?P="mult_row (mat 1) i (ass (A $ i $ ?n) div A $ i $ ?n)" let ?Ain="A $ i $ ?n" have ass_dvd: "ass ?Ain dvd ?Ain" using a unfolding ass_function_def by (simp add: associatedD1) have ass_dvd': "?Ain dvd ass ?Ain" using a unfolding ass_function_def by (simp add: associatedD1) assume iA: "\ is_zero_row i A" have Ain_0: "A $ i $ ?n \ 0" by (metis (mono_tags) LeastI iA is_zero_row_def') have ass_Ain_0: "ass (A $ i $ ?n) \ 0" by (metis Ain_0 ass_dvd dvd_0_left_iff) have inv_P: "invertible ?P" proof (rule invertible_mult_row[of _ "A $ i $ ?n div ass (A $ i $ ?n)"]) have "ass ?Ain div ?Ain * (?Ain div ass ?Ain) = (ass ?Ain div ?Ain * ?Ain) div ass ?Ain" by (rule div_mult_swap[OF ass_dvd]) also have "... = (ass ?Ain) div ass ?Ain" unfolding dvd_div_mult_self[OF ass_dvd'] .. also have "... = 1" using ass_Ain_0 by auto finally show "ass ?Ain div ?Ain * (?Ain div ass ?Ain) = 1" . have "?Ain div ass (?Ain) * (ass (?Ain) div ?Ain) = (?Ain div ass (?Ain) * ass (?Ain)) div ?Ain" by (rule div_mult_swap[OF ass_dvd']) also have "... = ?Ain div ?Ain" unfolding dvd_div_mult_self[OF ass_dvd] .. also have "... = 1" using Ain_0 by simp finally show "?Ain div ass (?Ain) * (ass (?Ain) div ?Ain) = 1" . qed obtain Q where inv_Q: "invertible Q" and H_QM: "Hermite_reduce_above ?M (to_nat i) i ?n res = Q ** ?M" using invertible_Hermite_reduce_above by fast have inv_QP: "invertible (Q**?P)" by (metis inv_P inv_Q invertible_mult) have "Hermite_reduce_above ?M (to_nat i) i ?n res = Q ** ?M" by (rule H_QM) also have "... = Q ** (?P ** A)" by (subst mult_row_mat_1[symmetric], rule refl) also have "... = (Q ** ?P) ** A" by (simp add: matrix_mul_assoc) finally show "\P. invertible P \ Hermite_reduce_above ?M (to_nat i) i ?n res = P ** A" using inv_QP by auto qed lemma invertible_Hermite_of_upt_row_i: assumes a: "ass_function ass" shows "\P. invertible P \ Hermite_of_upt_row_i A k ass res = P ** A" proof (induct k arbitrary: A) case 0 thus ?case unfolding Hermite_of_upt_row_i_def by (auto, metis invertible_def matrix_mul_lid) next case (Suc k) obtain Q where inv_Q: "invertible Q" and H_QA: "Hermite_of_upt_row_i A k ass res = Q ** A" using Suc.hyps by auto obtain P where inv_P: "invertible P" and H_PH: "Hermite_of_row_i ass res (Hermite_of_upt_row_i A k ass res) (from_nat k) = P ** (Hermite_of_upt_row_i A k ass res)" using invertible_Hermite_of_row_i[OF a] by blast have inv_PQ: "invertible (P**Q)" by (simp add: inv_P inv_Q invertible_mult) have "Hermite_of_upt_row_i A (Suc k) ass res = Hermite_of_row_i ass res (Hermite_of_upt_row_i A k ass res) (from_nat k)" unfolding Hermite_of_upt_row_i_def by auto also have "... = P ** (Hermite_of_upt_row_i A k ass res)" unfolding H_PH .. also have "... = P ** (Q ** A)" unfolding H_QA .. also have "... = (P ** Q) ** A" by (simp add: matrix_mul_assoc) finally show ?case using inv_PQ by blast qed lemma invertible_Hermite_of: fixes A::"'a::{bezout_ring_div,normalization_semidom,unique_euclidean_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes a: "ass_function ass" and b: "is_bezout_ext bezout" shows "\P. invertible P \ Hermite_of A ass res bezout = P ** A" proof - obtain P where inv_P: "invertible P" and H_PH: "Hermite_of_upt_row_i (echelon_form_of A bezout) (nrows A) ass res = P ** (echelon_form_of A bezout)" using invertible_Hermite_of_upt_row_i[OF a] by blast obtain Q where inv_Q: "invertible Q" and E_QA: "(echelon_form_of A bezout) = Q ** A" using echelon_form_of_invertible[OF b, of A] by auto have inv_PQ: "invertible (P**Q)" by (simp add: inv_P inv_Q invertible_mult) have "Hermite_of A ass res bezout = Hermite_of_upt_row_i (echelon_form_of A bezout) (nrows A) ass res" unfolding Hermite_of_def Let_def .. also have "... = P ** (Q ** A)" unfolding H_PH unfolding E_QA .. also have "... = (P ** Q) ** A" by (simp add: matrix_mul_assoc) finally show ?thesis using inv_PQ by blast qed subsubsection\The final theorem\ lemma Hermite: assumes a: "ass_function ass" and r: "res_function res" and b: "is_bezout_ext bezout" shows "\P. invertible P \ (Hermite_of A ass res bezout) = P ** A \ Hermite (range ass) (\c. range (res c)) (Hermite_of A ass res bezout)" using invertible_Hermite_of[OF a b] Hermite_Hermite_of[OF a r b] by fast subsection\Proving the uniqueness of the Hermite Normal Form\ lemma diagonal_least_nonzero: fixes H :: "(('a :: {bezout_ring_div, normalization_euclidean_semiring, unique_euclidean_ring}, 'b :: mod_type) vec, 'b) vec" assumes H: "Hermite associates residues H" and inv_H: "invertible H" and up_H: "upper_triangular H" shows "(LEAST n. H $ i $ n \ 0) = i" proof (rule Least_equality) show "H $ i $ i \ 0" by (metis (full_types) inv_H invertible_iff_is_unit is_unit_diagonal not_is_unit_0 up_H) fix y assume Hiy: "H $ i $ y \ 0" show "i \ y" using up_H unfolding upper_triangular_def by (metis (poly_guards_query) Hiy not_less) qed lemma diagonal_in_associates: fixes H :: "(('a :: {bezout_ring_div, normalization_euclidean_semiring, unique_euclidean_ring}, 'b :: mod_type) vec, 'b) vec" assumes H: "Hermite associates residues H" and inv_H: "invertible H" and up_H: "upper_triangular H" shows "H $ i $ i \ associates" proof - have "H $ i $ i \ 0" by (metis (full_types) inv_H invertible_iff_is_unit is_unit_diagonal not_is_unit_0 up_H) hence "\ is_zero_row i H" unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto thus ?thesis using H unfolding Hermite_def unfolding diagonal_least_nonzero[OF H inv_H up_H] by auto qed lemma above_diagonal_in_residues: fixes H :: "(('a :: {bezout_ring_div, normalization_euclidean_semiring, unique_euclidean_ring}, 'b :: mod_type) vec, 'b) vec" assumes H: "Hermite associates residues H" and inv_H: "invertible H" and up_H: "upper_triangular H" and j_i: "j 0) \ residues (H $ i $ (LEAST n. H $ i $ n \ 0))" proof - have "H $ i $ i \ 0" by (metis (full_types) inv_H invertible_iff_is_unit is_unit_diagonal not_is_unit_0 up_H) hence "\ is_zero_row i H" unfolding is_zero_row_def is_zero_row_upt_k_def ncols_def by auto thus ?thesis using H j_i unfolding Hermite_def unfolding diagonal_least_nonzero[OF H inv_H up_H] by auto qed text\The uniqueness of the Hermite Normal Form is proven following the proof presented in the book Integral Matrices (1972) by Morris Newman.\ lemma Hermite_unique: fixes K::"'a::{bezout_ring_div,normalization_euclidean_semiring,unique_euclidean_ring}^'n::mod_type^'n::mod_type" assumes A_PH: "A = P ** H" and A_QK: "A = Q ** K" and inv_A: "invertible A" and inv_P: "invertible P" and inv_Q: "invertible Q" and H: "Hermite associates residues H" and K: "Hermite associates residues K" shows "H = K" proof - have cs_residues: "Complete_set_residues residues" using H unfolding Hermite_def by simp have inv_H: "invertible H" by (metis A_PH inv_A inv_P invertible_def invertible_mult matrix_mul_assoc matrix_mul_lid) have inv_K: "invertible K" by (metis A_QK inv_A inv_Q invertible_def invertible_mult matrix_mul_assoc matrix_mul_lid) define U where "U = (matrix_inv P)**Q" have inv_U: "invertible U" by (metis U_def inv_P inv_Q invertible_def invertible_mult matrix_inv_left matrix_inv_right) have H_UK: "H = U ** K" using A_PH A_QK inv_P by (metis U_def matrix_inv_left matrix_mul_assoc matrix_mul_lid) have "det K *k U = H ** adjugate K" unfolding H_UK matrix_mul_assoc[symmetric] mult_adjugate_det matrix_mul_mat .. have upper_triangular_H: "upper_triangular H" by (metis H Hermite_def echelon_form_imp_upper_triagular) have upper_triangular_K: "upper_triangular K" by (metis K Hermite_def echelon_form_imp_upper_triagular) have upper_triangular_U: "upper_triangular U" by (metis H_UK inv_K matrix_inv_right matrix_mul_assoc matrix_mul_rid upper_triangular_H upper_triangular_K upper_triangular_inverse upper_triangular_mult) have unit_det_U: "is_unit (det U)" by (metis inv_U invertible_iff_is_unit) have is_unit_diagonal_U: "(\i. is_unit (U $ i $ i))" by (rule is_unit_diagonal[OF upper_triangular_U unit_det_U]) have Uii_1: "(\i. (U $ i $ i) = 1)" and Hii_Kii: "(\i. (H $ i $ i) = (K $ i $ i))" proof (auto) fix i have Hii: "H $ i $ i \ associates" by (rule diagonal_in_associates[OF H inv_H upper_triangular_H]) have Kii: "K $ i $ i \ associates" by (rule diagonal_in_associates[OF K inv_K upper_triangular_K]) have ass_Hii_Kii: "normalize (H $ i $ i) = normalize (K $ i $ i)" by (meson associatedI inv_H inv_K invertible_iff_is_unit is_unit_diagonal unit_imp_dvd upper_triangular_H upper_triangular_K) show Hii_eq_Kii: "H $ i $ i = K $ i $ i" by (metis Hermite_def Hii K Kii ass_Hii_Kii in_Ass_not_associated) have "H $ i $ i = U $ i $ i * K $ i $ i" by (metis H_UK upper_triangular_K upper_triangular_U upper_triangular_mult_diagonal) thus "U $ i $ i = 1" unfolding Hii_eq_Kii mult_cancel_right1 by (metis Hii_eq_Kii inv_H invertible_iff_is_unit is_unit_diagonal not_is_unit_0 upper_triangular_H) qed have zero_above: "\j s. j\1 \ j < ncols A - to_nat s \ U $ s $ (s + from_nat j) = 0" proof (clarify) fix j s assume "1 \ j" and "j < ncols A - (to_nat (s::'n))" thus "U $ s $ (s + from_nat j) = 0" proof (induct j rule: less_induct) fix p assume induct_step: "(\y. y < p \ 1 \ y \ y < ncols A - to_nat s \ U $ s $ (s + from_nat y) = 0)" and p1: "1 \ p" and p2: "p < ncols A - to_nat s" have s_less: "s < s + from_nat p" using p1 p2 unfolding ncols_def by (metis One_nat_def add.commute add_diff_cancel_right' add_lessD1 add_to_nat_def from_nat_to_nat_id less_diff_conv neq_iff not_le to_nat_from_nat_id to_nat_le zero_less_Suc) show "U $ s $ (s + from_nat p) = 0" proof - have UNIV_rw: "UNIV = insert s (UNIV-{s})" by auto have UNIV_s_rw: "UNIV-{s} = insert (s + from_nat p) ((UNIV-{s}) - {s + from_nat p})" using p1 p2 s_less unfolding ncols_def by (auto simp: algebra_simps) have sum_rw: "(\k\UNIV-{s}. U $ s $ k * K $ k $ (s + from_nat p)) = U $ s $ (s + from_nat p) * K $ (s + from_nat p) $ (s + from_nat p) + (\k\(UNIV-{s})-{s + from_nat p}. U $ s $ k * K $ k $ (s + from_nat p))" using UNIV_s_rw sum.insert by (metis (erased, lifting) Diff_iff finite singletonI) have sum_0: "(\k\(UNIV-{s})-{s + from_nat p}. U $ s $ k * K $ k $ (s + from_nat p)) = 0" proof (rule sum.neutral, rule) fix x assume x: "x \ UNIV - {s} - {s + from_nat p}" show "U $ s $ x * K $ x $ (s + from_nat p) = 0" proof (cases "xs" using x by (metis Diff_iff neq_iff singletonI) show ?thesis proof (cases "x a" by (auto simp add: a_def p1 p2) (metis Suc_leI to_nat_mono x_g_s zero_less_diff) show "a < ncols A - to_nat s" using a_p p2 by auto qed thus ?thesis by simp next case False hence "x>s+from_nat p" using x_g_s x by auto thus ?thesis using upper_triangular_K unfolding upper_triangular_def by auto qed qed qed have "H $ s $ (s + from_nat p) = (\k\UNIV. U $ s $ k * K $ k $ (s + from_nat p))" unfolding H_UK matrix_matrix_mult_def by auto also have "... = (\k\insert s (UNIV-{s}). U $ s $ k * K $ k $ (s + from_nat p))" using UNIV_rw by simp also have "... = U $ s $ s * K $ s $ (s + from_nat p) + (\k\UNIV-{s}. U $ s $ k * K $ k $ (s + from_nat p))" by (rule sum.insert, simp_all) also have "... = U $ s $ s * K $ s $ (s + from_nat p) + U $ s $ (s + from_nat p) * K $ (s + from_nat p) $ (s + from_nat p)" unfolding sum_rw sum_0 by simp finally have H_s_sp: "H $ s $ (s + from_nat p) = U $ s $ (s + from_nat p) * K $ (s + from_nat p) $ (s + from_nat p) + K $ s $ (s + from_nat p)" using Uii_1 by auto hence cong_HK: "cong (H $ s $ (s + from_nat p)) (K $ s $ (s + from_nat p)) (K $ (s+from_nat p) $ (s + from_nat p))" unfolding cong_def by auto have H_s_sp_residues: "(H $ s $ (s + from_nat p)) \ residues (K $ (s+from_nat p) $ (s + from_nat p))" using above_diagonal_in_residues[OF H inv_H upper_triangular_H s_less] unfolding diagonal_least_nonzero[OF H inv_H upper_triangular_H] by (metis Hii_Kii) have K_s_sp_residues: "(K $ s $ (s + from_nat p)) \ residues (K $ (s+from_nat p) $ (s + from_nat p))" using above_diagonal_in_residues[OF K inv_K upper_triangular_K s_less] unfolding diagonal_least_nonzero[OF K inv_K upper_triangular_K] . have Hs_sp_Ks_sp: "(H $ s $ (s + from_nat p)) = (K $ s $ (s + from_nat p))" using cong_HK in_Res_not_congruent[OF cs_residues H_s_sp_residues K_s_sp_residues] by fast have "is_unit (K $ (s + from_nat p) $ (s + from_nat p))" by (metis Hii_Kii inv_H invertible_iff_is_unit is_unit_diagonal upper_triangular_H) hence "K $ (s + from_nat p) $ (s + from_nat p) \ 0" by (metis not_is_unit_0) thus ?thesis unfolding from_nat_1 using H_s_sp unfolding Hs_sp_Ks_sp by auto qed qed qed have "U = mat 1" proof (unfold mat_def vec_eq_iff, auto) fix ia show "U $ ia $ ia = 1" using Uii_1 by simp fix i assume i_ia: "i \ ia" show "U $ i $ ia = 0" proof (cases "ia a" unfolding a_def by (metis diff_is_0_eq i_less_ia less_one not_less to_nat_mono) moreover have "a < ncols A - to_nat i" unfolding a_def ncols_def by (metis False diff_less_mono not_less to_nat_less_card to_nat_mono') ultimately show ?thesis using zero_above unfolding ia_eq by blast qed qed thus ?thesis using H_UK matrix_mul_lid by fast qed subsection\Examples of execution\ value[code] "let A = list_of_list_to_matrix ([[37,8,6],[5,4,-8],[3,24,-7]])::int^3^3 in matrix_to_list_of_list (Hermite_of A ass_function_euclidean res_function_euclidean euclid_ext2)" value[code] "let A = list_of_list_to_matrix ([[[:3,4,5:],[:-2,1:]],[[:-1,0,2:],[:0,1,4,1:]]])::real poly^2^2 in matrix_to_list_of_list (Hermite_of A ass_function_euclidean res_function_euclidean euclid_ext2)" end diff --git a/thys/Jordan_Normal_Form/Determinant_Impl.thy b/thys/Jordan_Normal_Form/Determinant_Impl.thy --- a/thys/Jordan_Normal_Form/Determinant_Impl.thy +++ b/thys/Jordan_Normal_Form/Determinant_Impl.thy @@ -1,962 +1,962 @@ (* Author: RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Code Equations for Determinants\ text \We compute determinants on arbitrary rings by applying elementary row-operations to bring a matrix on upper-triangular form. Then the determinant can be determined by multiplying all entries on the diagonal. Moreover the final result has to be divided by a factor which is determined by the row-operations that we performed. To this end, we require a division operation on the element type. The algorithm is parametric in a selection function for the pivot-element, e.g., for matrices over polynomials it turned out that selecting a polynomial of minimal degree is beneficial.\ theory Determinant_Impl imports Polynomial_Interpolation.Missing_Polynomial "HOL-Computational_Algebra.Polynomial_Factorial" Determinant begin type_synonym 'a det_selection_fun = "(nat \ 'a)list \ nat" definition det_selection_fun :: "'a det_selection_fun \ bool" where "det_selection_fun f = (\ xs. xs \ [] \ f xs \ fst ` set xs)" lemma det_selection_funD: "det_selection_fun f \ xs \ [] \ f xs \ fst ` set xs" unfolding det_selection_fun_def by auto definition mute_fun :: "('a :: comm_ring_1 \ 'a \ 'a \ 'a \ 'a) \ bool" where "mute_fun f = (\ x y x' y' g. f x y = (x',y',g) \ y \ 0 \ x = x' * g \ y * x' = x * y')" context fixes sel_fun :: "'a ::idom_divide det_selection_fun" begin subsection \Properties of triangular matrices\ text \ Each column of a triangular matrix should satisfy the following property. \ definition triangular_column::"nat \ 'a mat \ bool" where "triangular_column j A \ \i. j < i \ i < dim_row A \ A $$ (i,j) = 0" lemma triangular_columnD [dest]: "triangular_column j A \ j < i \ i < dim_row A \ A $$ (i,j) = 0" unfolding triangular_column_def by auto lemma triangular_columnI [intro]: "(\i. j < i \ i < dim_row A \ A $$ (i,j) = 0) \ triangular_column j A" unfolding triangular_column_def by auto text \ The following predicate states that the first $k$ columns satisfy triangularity. \ definition triangular_to:: "nat \ 'a mat \ bool" where "triangular_to k A == \j. j triangular_column j A" lemma triangular_to_triangular: "upper_triangular A = triangular_to (dim_row A) A" unfolding triangular_to_def triangular_column_def upper_triangular_def by auto lemma triangular_toD [dest]: "triangular_to k A \ j < k \ j < i \ i < dim_row A \ A $$ (i,j) = 0" unfolding triangular_to_def triangular_column_def by auto lemma triangular_toI [intro]: "(\i j. j < k \ j < i \ i < dim_row A \ A $$ (i,j) = 0) \ triangular_to k A" unfolding triangular_to_def triangular_column_def by auto lemma triangle_growth: assumes tri:"triangular_to k A" and col:"triangular_column k A" shows "triangular_to (Suc k) A" unfolding triangular_to_def proof (intro allI impI) fix i assume iSk:"i < Suc k" show "triangular_column i A" proof (cases "i = k") case True then show ?thesis using col by auto next case False then have "i < k" using iSk by auto thus ?thesis using tri unfolding triangular_to_def by auto qed qed lemma triangle_trans: "triangular_to k A \ k > k' \ triangular_to k' A" by (intro triangular_toI, elim triangular_toD, auto) subsection \Algorithms for Triangulization\ context fixes mf :: "'a \ 'a \ 'a \ 'a \ 'a" begin private fun mute :: "'a \ nat \ nat \ 'a \ 'a mat \ 'a \ 'a mat" where "mute A_ll k l (r,A) = (let p = A $$ (k,l) in if p = 0 then (r,A) else case mf A_ll p of (q',p',g) \ (r * q', addrow (-p') k l (multrow k q' A)))" lemma mute_preserves_dimensions: assumes "mute q k l (r,A) = (r',A')" shows [simp]: "dim_row A' = dim_row A" and [simp]: "dim_col A' = dim_col A" using assms by (auto simp: Let_def split: if_splits prod.splits) text \ Algorithm @{term "mute k l"} makes $k$-th row $l$-th column element to 0. \ lemma mute_makes_0 : assumes mute_fun: "mute_fun mf" assumes "mute (A $$ (l,l)) k l (r,A) = (r',A')" "l < dim_row A" "l < dim_col A" "k < dim_row A" "k \ l" shows "A' $$ (k,l) = 0" proof - define a where "a = A $$ (l, l)" define b where "b = A $$ (k, l)" let ?mf = "mf (A $$ (l, l)) (A $$ (k, l))" obtain q' p' g where id: "?mf = (q',p',g)" by (cases ?mf, auto) note mf = mute_fun[unfolded mute_fun_def, rule_format, OF id] from assms show ?thesis unfolding mat_addrow_def using mf id by (auto simp: ac_simps Let_def split: if_splits) qed text \It will not touch unexpected rows.\ lemma mute_preserves: "mute q k l (r,A) = (r',A') \ i < dim_row A \ j < dim_col A \ l < dim_row A \ k < dim_row A \ i \ k \ A' $$ (i,j) = A $$ (i,j)" by (auto simp: Let_def split: if_splits prod.splits) text \It preserves $0$s in the touched row.\ lemma mute_preserves_0: "mute q k l (r,A) = (r',A') \ i < dim_row A \ j < dim_col A \ l < dim_row A \ k < dim_row A \ A $$ (i,j) = 0 \ A $$ (l,j) = 0 \ A' $$ (i,j) = 0" by (auto simp: Let_def split: if_splits prod.splits) text \Hence, it will respect partially triangular matrix.\ lemma mute_preserves_triangle: assumes rA' : "mute q k l (r,A) = (r',A')" and triA: "triangular_to l A" and lk: "l < k" and kr: "k < dim_row A" and lr: "l < dim_row A" and lc: "l < dim_col A" shows "triangular_to l A'" proof (rule triangular_toI) fix i j assume jl: "j < l" and ji: "j < i" and ir': "i < dim_row A'" then have A0: "A $$ (i,j) = 0" using triA rA' by auto moreover have "A $$ (l,j) = 0" using triA jl jl lr by auto moreover have jc:"j < dim_col A" using jl lc by auto moreover have ir: "i < dim_row A" using ir' rA' by auto ultimately show "A' $$ (i,j) = 0" using mute_preserves_0[OF rA'] lr kr by auto qed text \Recursive application of @{const mute}\ private fun sub1 :: "'a \ nat \ nat \ 'a \ 'a mat \ 'a \ 'a mat" where "sub1 q 0 l rA = rA" | "sub1 q (Suc k) l rA = mute q (l + Suc k) l (sub1 q k l rA)" lemma sub1_preserves_dimensions[simp]: "sub1 q k l (r,A) = (r',A') \ dim_row A' = dim_row A" "sub1 q k l (r,A) = (r',A') \ dim_col A' = dim_col A" proof (induction k arbitrary: r' A') case (Suc k) moreover obtain r' A' where rA': "sub1 q k l (r, A) = (r', A')" by force moreover fix r'' A'' assume "sub1 q (Suc k) l (r, A) = (r'', A'')" ultimately show "dim_row A'' = dim_row A" "dim_col A'' = dim_col A" by auto qed auto lemma sub1_closed [simp]: "sub1 q k l (r,A) = (r',A') \ A \ carrier_mat m n \ A' \ carrier_mat m n" unfolding carrier_mat_def by auto lemma sub1_preserves_diagnal: assumes "sub1 q k l (r,A) = (r',A')" and "l < dim_col A" and "k + l < dim_row A" shows "A' $$ (l,l) = A $$ (l,l)" using assms proof - show "k + l < dim_row A \ sub1 q k l (r,A) = (r',A') \ A' $$ (l,l) = A $$ (l,l)" proof (induction k arbitrary: r' A') case (Suc k) obtain r'' A'' where rA''[simp]: "sub1 q k l (r,A) = (r'',A'')" by force have [simp]:"dim_row A'' = dim_row A" and [simp]:"dim_col A'' = dim_col A" using snd_conv sub1_preserves_dimensions[OF rA''] by auto have "A'' $$ (l,l) = A $$ (l,l)" using assms Suc by auto have rA': "mute q (l + Suc k) l (r'', A'') = (r',A')" using Suc by auto show ?case using subst mute_preserves[OF rA'] Suc assms by auto qed auto qed text \Triangularity is respected by @{const sub1}.\ lemma sub1_preserves_triangle: assumes "sub1 q k l (r,A) = (r',A')" and tri: "triangular_to l A" and lr: "l < dim_row A" and lc: "l < dim_col A" and lkr: "l + k < dim_row A" shows "triangular_to l A'" using assms proof - show "sub1 q k l (r,A) = (r',A') \ l + k < dim_row A \ triangular_to l A'" proof (induction k arbitrary: r' A') case (Suc k) then have "sub1 q (Suc k) l (r,A) = (r',A')" by auto moreover obtain r'' A'' where rA'': "sub1 q k l (r, A) = (r'',A'')" by force ultimately have rA': "mute q (Suc (l + k)) l (r'',A'') = (r',A')" by auto have "triangular_to l A''" using rA'' Suc by auto thus ?case using Suc assms mute_preserves_triangle[OF rA'] rA'' by auto qed (insert assms,auto) qed context assumes mf: "mute_fun mf" begin lemma sub1_makes_0s: assumes "sub1 (A $$ (l,l)) k l (r,A) = (r',A')" and lr: "l < dim_row A" and lc: "l < dim_col A" and li: "l < i" and "i \ k + l" and "k + l < dim_row A" shows "A' $$ (i,l) = 0" using assms proof - show "sub1 (A $$ (l,l)) k l (r,A) = (r',A') \ i \ k + l \ k + l < dim_row A \ A' $$ (i,l) = 0" using lr lc li proof (induction k arbitrary: r' A') case (Suc k) obtain r' A' where rA': "sub1 (A $$ (l,l)) k l (r, A) = (r',A')" by force fix r'' A'' from sub1_preserves_diagnal[OF rA'] have AA': "A $$ (l, l) = A' $$ (l, l)" using Suc(2-) by auto assume "sub1 (A $$ (l,l)) (Suc k) l (r, A) = (r'',A'')" then have rA'': "mute (A $$ (l,l)) (Suc (l + k)) l (r', A') = (r'', A'')" using rA' by simp have ir: "i < dim_row A" using Suc by auto have il: "i \ l" using li by auto have lr': "l < dim_row A'" using lr rA' by auto have lc': "l < dim_col A'" using lc rA' by auto have Slkr': "Suc (l+k) < dim_row A'" using Suc rA' by auto show "A'' $$ (i,l) = 0" proof (cases "Suc(l + k) = i") case True { have l: "Suc (l + k) \ l" by auto show ?thesis using mute_makes_0[OF mf rA''[unfolded AA'] lr' lc' Slkr' l] ir il rA' by (simp add:True) } next case False { then have ikl: "i \ k+l" using Suc by auto have ir': "i < dim_row A'" using ir rA' by auto have lc': "l < dim_col A'" using lc rA' by auto have IH: "A' $$ (i,l) = 0" using rA' Suc False by auto thus ?thesis using mute_preserves[OF rA'' ir' lc'] rA' False Suc by simp } qed qed auto qed lemma sub1_triangulizes_column: assumes rA': "sub1 (A $$ (l,l)) (dim_row A - Suc l) l (r,A) = (r',A')" and tri:"triangular_to l A" and r: "dim_row A > 0" and lr: "l < dim_row A" and lc: "l < dim_col A" shows "triangular_column l A'" proof (intro triangular_columnI) fix i assume li: "l < i" assume ir: "i < dim_row A'" also have "... = dim_row A" using sub1_preserves_dimensions[OF rA'] by auto also have "... = dim_row A - l + l" using lr li by auto finally have ir2: "i \ dim_row A - l + l" by auto show "A' $$ (i,l) = 0" apply (subst sub1_makes_0s[OF rA' lr lc]) using li ir assms by auto qed text \ The algorithm @{const sub1} increases the number of columns that form triangle. \ lemma sub1_grows_triangle: assumes rA': "sub1 (A $$ (l,l)) (dim_row A - Suc l) l (r,A) = (r',A')" and r: "dim_row A > 0" and tri:"triangular_to l A" and lr: "l < dim_row A" and lc: "l < dim_col A" shows "triangular_to (Suc l) A'" proof - have "triangular_to l A'" using sub1_preserves_triangle[OF rA'] assms by auto moreover have "triangular_column l A'" using sub1_triangulizes_column[OF rA'] assms by auto ultimately show ?thesis by (rule triangle_growth) qed end subsection \Finding Non-Zero Elements\ private definition find_non0 :: "nat \ 'a mat \ nat option" where "find_non0 l A = (let is = [Suc l ..< dim_row A]; Ais = filter (\ (i,Ail). Ail \ 0) (map (\ i. (i, A $$ (i,l))) is) in case Ais of [] \ None | _ \ Some (sel_fun Ais))" lemma find_non0: assumes sel_fun: "det_selection_fun sel_fun" and res: "find_non0 l A = Some m" shows "A $$ (m,l) \ 0" "l < m" "m < dim_row A" proof - let ?xs = "filter (\ (i,Ail). Ail \ 0) (map (\ i. (i, A $$ (i,l))) [Suc l.. []" and m: "m = sel_fun ?xs" by (cases ?xs, auto)+ from det_selection_funD[OF sel_fun xs, folded m] show "A $$ (m, l) \ 0" "l < m" "m < dim_row A" by auto qed text \ If @{term "find_non0 l A"} fails, then $A$ is already triangular to $l$-th column. \ lemma find_non0_all0: "find_non0 l A = None \ triangular_column l A" proof (intro triangular_columnI) fix i let ?xs = "filter (\ (i,Ail). Ail \ 0) (map (\ i. (i, A $$ (i,l))) [Suc l.. set (map (\ i. (i, A $$ (i,l))) [Suc l..Determinant Preserving Growth of Triangle\ text \ The algorithm @{const sub1} does not preserve determinants when it hits a $0$-valued diagonal element. To avoid this case, we introduce the following operation: \ private fun sub2 :: "nat \ nat \ 'a \ 'a mat \ 'a \ 'a mat" where "sub2 d l (r,A) = ( case find_non0 l A of None \ (r,A) | Some m \ let A' = swaprows m l A in sub1 (A' $$ (l,l)) (d - Suc l) l (-r, A'))" lemma sub2_preserves_dimensions[simp]: assumes rA': "sub2 d l (r,A) = (r',A')" shows "dim_row A' = dim_row A \ dim_col A' = dim_col A" proof (cases "find_non0 l A") case None then show ?thesis using rA' by auto next case (Some m) then show ?thesis using rA' by (cases "m = l", auto simp: Let_def) qed lemma sub2_closed [simp]: "sub2 d l (r,A) = (r',A') \ A \ carrier_mat m n \ A' \ carrier_mat m n" unfolding carrier_mat_def by auto context assumes sel_fun: "det_selection_fun sel_fun" begin lemma sub2_preserves_triangle: assumes rA': "sub2 d l (r,A) = (r',A')" and tri: "triangular_to l A" and lc: "l < dim_col A" and ld: "l < d" and dr: "d \ dim_row A" shows "triangular_to l A'" proof - have lr: "l < dim_row A" using ld dr by auto show ?thesis proof (cases "find_non0 l A") case None then show ?thesis using rA' tri by simp next case (Some m) { have lm : "l < m" and mr : "m < dim_row A" using find_non0[OF sel_fun Some] by auto let ?A1 = "swaprows m l A" have tri'': "triangular_to l ?A1" proof (intro triangular_toI) fix i j assume jl:"j < l" and ji:"j < i" and ir1: "i < dim_row ?A1" have jm: "j < m" using jl lm by auto have ir: "i < dim_row A" using ir1 by auto have jc: "j < dim_col A" using jl lc by auto show "?A1 $$ (i, j) = 0" proof (cases "m=i") case True { then have li: "l \ i" using lm by auto hence "?A1 $$ (i,j) = A $$ (l,j)" using ir jc \m=i\ by auto also have "... = 0" using tri jl lr by auto finally show ?thesis. } next case False show ?thesis proof (cases "l=i") case True { then have "?A1 $$ (i,j) = A $$ (m,j)" using ir jc \m\i\ by auto thus "?A1 $$ (i,j) = 0" using tri jl jm mr by auto } next case False { then have "?A1 $$ (i,j) = A $$ (i,j)" using ir jc \m\i\ by auto thus "?A1 $$ (i,j) = 0" using tri jl ji ir by auto } qed qed qed let ?rA3 = "sub1 (?A1 $$ (l,l)) (d - Suc l) l (-r,?A1)" have [simp]: "dim_row ?A1 = dim_row A \ dim_col ?A1 = dim_col A" by auto have rA'2: "?rA3 = (r',A')" using rA' Some by (simp add: Let_def) have "l + (d - Suc l) < dim_row A" using ld dr by auto thus ?thesis using sub1_preserves_triangle[OF rA'2 tri''] lr lc rA' by auto } qed qed lemma sub2_grows_triangle: assumes mf: "mute_fun mf" and rA': "sub2 (dim_row A) l (r,A) = (r',A')" and tri: "triangular_to l A" and lc: "l < dim_col A" and lr: "l < dim_row A" shows "triangular_to (Suc l) A'" proof (rule triangle_growth) show "triangular_to l A'" using sub2_preserves_triangle[OF rA' tri lc lr] by auto next have r0: "0 < dim_row A" using lr by auto show "triangular_column l A'" proof (cases "find_non0 l A") case None { then have "A' = A" using rA' by simp moreover have "triangular_column l A" using find_non0_all0[OF None]. ultimately show ?thesis by auto } next case (Some m) { have lm: "l < m" and mr: "m < dim_row A" using find_non0[OF sel_fun Some] by auto let ?A = "swaprows m l A" have tri2: "triangular_to l ?A" proof fix i j assume jl: "j < l" and ji:"j < i" and ir: "i < dim_row ?A" show "?A $$ (i,j) = 0" proof (cases "i = m") case True { then have "?A $$ (i,j) = A $$ (l,j)" using jl lc ir by simp also have "... = 0" using triangular_toD[OF tri jl jl] lr by auto finally show ?thesis by auto } next case False show ?thesis proof (cases "i = l") case True { then have "?A $$ (i,j) = A $$ (m,j)" using jl lc ir by auto also have "... = 0" using triangular_toD[OF tri jl] jl lm mr by auto finally show ?thesis by auto } next case False { then have "?A $$ (i,j) = A $$ (i,j)" using \i \ m\ jl lc ir by auto thus ?thesis using tri jl ji ir by auto } qed qed qed have rA'2: "sub1 (?A $$ (l,l)) (dim_row ?A - Suc l) l (-r, ?A) = (r',A')" using lm Some rA' by (simp add: Let_def) show ?thesis using sub1_triangulizes_column[OF mf rA'2 tri2] r0 lr lc by auto } qed qed end subsection \Recursive Triangulization of Columns\ text \ Now we recursively apply @{const sub2} to make the entire matrix to be triangular. \ private fun sub3 :: "nat \ nat \ 'a \ 'a mat \ 'a \ 'a mat" where "sub3 d 0 rA = rA" | "sub3 d (Suc l) rA = sub2 d l (sub3 d l rA)" lemma sub3_preserves_dimensions[simp]: "sub3 d l (r,A) = (r',A') \ dim_row A' = dim_row A" "sub3 d l (r,A) = (r',A') \ dim_col A' = dim_col A" proof (induction l arbitrary: r' A') case (Suc l) obtain r' A' where rA': "sub3 d l (r, A) = (r', A')" by force fix r'' A'' assume rA'': "sub3 d (Suc l) (r, A) = (r'', A'')" then show "dim_row A'' = dim_row A" "dim_col A'' = dim_col A" using Suc rA' by auto qed auto lemma sub3_closed[simp]: "sub3 k l (r,A) = (r',A') \ A \ carrier_mat m n \ A' \ carrier_mat m n" unfolding carrier_mat_def by auto lemma sub3_makes_triangle: assumes mf: "mute_fun mf" and sel_fun: "det_selection_fun sel_fun" and "sub3 (dim_row A) l (r,A) = (r',A')" and "l \ dim_row A" and "l \ dim_col A" shows "triangular_to l A'" using assms proof - show "sub3 (dim_row A) l (r,A) = (r',A') \ l \ dim_row A \ l \ dim_col A \ triangular_to l A'" proof (induction l arbitrary:r' A') case (Suc l) then have Slr: "Suc l \ dim_row A" and Slc: "Suc l \ dim_col A" by auto hence lr: "l < dim_row A" and lc: "l < dim_col A" by auto moreover obtain r'' A'' where rA'': "sub3 (dim_row A) l (r,A) = (r'',A'')" by force ultimately have IH: "triangular_to l A''" using Suc by auto have [simp]:"dim_row A'' = dim_row A" and [simp]:"dim_col A'' = dim_col A" using Slr Slc rA'' by auto fix r' A' assume "sub3 (dim_row A) (Suc l) (r, A) = (r', A')" then have rA': "sub2 (dim_row A'') l (r'',A'') = (r',A')" using rA'' by auto show "triangular_to (Suc l) A'" using sub2_grows_triangle[OF sel_fun mf rA'] lr lc rA'' IH by auto qed auto qed subsection \Triangulization\ definition triangulize :: "'a mat \ 'a \ 'a mat" where "triangulize A = sub3 (dim_row A) (dim_row A) (1,A)" lemma triangulize_preserves_dimensions[simp]: "triangulize A = (r',A') \ dim_row A' = dim_row A" "triangulize A = (r',A') \ dim_col A' = dim_col A" unfolding triangulize_def by auto lemma triangulize_closed[simp]: "triangulize A = (r',A') \ A \ carrier_mat m n \ A' \ carrier_mat m n" unfolding carrier_mat_def by auto context assumes mf: "mute_fun mf" and sel_fun: "det_selection_fun sel_fun" begin theorem triangulized: assumes "A \ carrier_mat n n" and "triangulize A = (r',A')" shows "upper_triangular A'" proof (cases "0Divisor will not be 0\ text \ Here we show that each sub-algorithm will not make $r$ of the input/output pair $(r,A)$ to 0. The algorithm @{term "sub1 A_ll k l (r,A)"} requires $A_{l,l} \neq 0$. \ lemma sub1_divisor [simp]: assumes rA': "sub1 q k l (r, A) = (r',A')" and r0: "r \ 0" and All: "q \ 0" and "k + l < dim_row A " and lc: "l < dim_col A" shows "r' \ 0" using assms proof - show "sub1 q k l (r,A) = (r',A') \ k + l < dim_row A \ r' \ 0" proof (induction k arbitrary: r' A') case (Suc k) obtain r'' A'' where rA'': "sub1 q k l (r, A) = (r'', A'')" by force then have IH: "r'' \ 0" using Suc by auto obtain q' l' g where mf_id: "mf q (A'' $$ (Suc (l + k), l)) = (q',l',g)" by (rule prod_cases3) define fact where "fact = (if A'' $$ (Suc (l+k),l) = 0 then 1 else q')" note mf = mf[unfolded mute_fun_def, rule_format, OF mf_id] have All: "q \ 0" using sub1_preserves_diagnal[OF rA'' lc] All Suc by auto moreover have "fact \ 0" unfolding fact_def using All mf by auto moreover assume "sub1 q (Suc k) l (r,A) = (r',A')" then have "mute q (Suc (l + k)) l (r'',A'') = (r',A')" using rA'' by auto hence "r'' * fact = r'" unfolding mute.simps fact_def Let_def mf_id using IH by (auto split: if_splits) ultimately show "r' \ 0" using IH by auto qed (insert r0, simp) qed text \The algorithm @{term "sub2"} will not require such a condition.\ lemma sub2_divisor [simp]: assumes rA': "sub2 k l (r, A) = (r',A')" and lk: "l < k" and kr: "k \ dim_row A" and lc: "l < dim_col A" and r0: "r \ 0" shows "r' \ 0" using assms proof (cases "find_non0 l A") { case (Some m) then have Aml0: "A $$ (m,l) \ 0" using find_non0[OF sel_fun] by auto have md: "m < dim_row A" using find_non0[OF sel_fun Some] lk kr by auto let ?A'' = "swaprows m l A" have rA'2: "sub1 (?A'' $$ (l,l)) (k - Suc l) l (-r, ?A'') = (r',A')" using rA' Some by (simp add: Let_def) have All0: "?A'' $$ (l,l) \ 0" using Aml0 md lk kr lc by auto show ?thesis using sub1_divisor[OF rA'2 _ All0] r0 lk kr lc by simp } qed auto lemma sub3_divisor [simp]: assumes "sub3 d l (r,A) = (r'',A'')" and "l \ d" and "d \ dim_row A" and "l \ dim_col A" and r0: "r \ 0" shows "r'' \ 0" using assms proof - show "sub3 d l (r,A) = (r'',A'') \ l \ d \ d \ dim_row A \ l \ dim_col A \ ?thesis" proof (induction l arbitrary: r'' A'') case 0 then show ?case using r0 by simp next case (Suc l) obtain r' A' where rA': "sub3 d l (r,A) = (r',A')" by force then have [simp]:"dim_row A' = dim_row A" and [simp]:"dim_col A' = dim_col A" by auto from rA' have "r' \ 0" using Suc r0 by auto moreover have "sub2 d l (r',A') = (r'',A'')" using rA' Suc by simp ultimately show ?case using sub2_divisor using Suc by simp qed qed theorem triangulize_divisor: assumes A: "A \ carrier_mat d d" shows "triangulize A = (r',A') \ r' \ 0" unfolding triangulize_def proof - assume rA': "sub3 (dim_row A) (dim_row A) (1, A) = (r', A')" show ?thesis using sub3_divisor[OF rA'] A by auto qed subsection \Determinant Preservation Results\ text \ For each sub-algorithm $f$, we show $f(r,A) = (r',A')$ implies @{term "r * det A' = r' * det A"}. \ lemma mute_det: assumes "A \ carrier_mat n n" and rA': "mute q k l (r,A) = (r',A')" and "k < n" and "l < n" and "k \ l" shows "r * det A' = r' * det A" proof (cases "A $$ (k,l) = 0") case True thus ?thesis using assms by auto next case False obtain p' q' g where mf_id: "mf q (A $$ (k,l)) = (q',p',g)" by (rule prod_cases3) let ?All = "q'" let ?Akl = "- p'" let ?B = "multrow k ?All A" let ?C = "addrow ?Akl k l ?B" have "r * det A' = r * det ?C" using assms by (simp add: Let_def mf_id False) also have "det ?C = det ?B" using assms by (auto simp: det_addrow) also have "\ = ?All * det A" using assms det_multrow by auto also have "r * \ = (r * ?All) * det A" by simp also have r: "r * ?All = r'" using assms by (simp add: Let_def mf_id False) finally show ?thesis. qed lemma sub1_det: assumes A: "A \ carrier_mat n n" and sub1: "sub1 q k l (r,A) = (r'',A'')" and r0: "r \ 0" and All0: "q \ 0" and l: "l + k < n" shows "r * det A'' = r'' * det A" using sub1 l proof (induction k arbitrary: A'' r'') case 0 then show ?case by auto next case (Suc k) let ?rA' = "sub1 q k l (r,A)" obtain r' A' where rA':"?rA' = (r',A')" by force have A':"A' \ carrier_mat n n" using sub1_closed[OF rA'] A by auto have IH: "r * det A' = r' * det A" using Suc assms rA' by auto assume "sub1 q (Suc k) l (r,A) = (r'',A'')" then have rA'':"mute q (Suc (l+k)) l (r',A') = (r'',A'')" using rA' by auto hence lem: "r' * det A'' = r'' * det A'" using assms Suc A' mute_det[OF A' rA''] by auto hence "r * r' * det A'' = r * r'' * det A'" by auto also from IH have "... = r'' * r' * det A" by auto finally have *: "r * r' * det A'' = r'' * r' * det A" . then have "r * r' * det A'' div r' = r'' * r' * det A div r'" by presburger moreover have "r' \ 0" using r0 sub1_divisor[OF rA'] All0 Suc A by auto ultimately show ?case using * by auto qed lemma sub2_det: assumes A: "A \ carrier_mat d d" and rA': "sub2 d l (r,A) = (r',A')" and r0: "r \ 0" and ld: "l < d" shows "r * det A' = r' * det A" proof (cases "find_non0 l A") case None then show ?thesis using assms by auto next case (Some m) { then have lm: "l < m" and md: "m < d" using A find_non0[OF sel_fun Some] ld by auto hence "m \ l" by auto let ?A'' = "swaprows m l A" have rA'2: "sub1 (?A'' $$ (l,l)) (d - Suc l) l (-r, ?A'') = (r',A')" using rA' Some by (simp add: Let_def) have A'': "?A'' \ carrier_mat d d" using A by auto hence A''ll0: "?A'' $$ (l,l) \ 0" using find_non0[OF sel_fun Some] ld by auto hence "-r * det A' = r' * det ?A''" using sub1_det[OF A'' rA'2] ld A r0 by auto also have "r * ... = -r * r' * det A" using det_swaprows[OF md ld \m\l\ A] by auto finally have "r * -r * det A' = -r * r' * det A" by auto thus ?thesis using r0 by auto } qed lemma sub3_det: assumes A:"A \ carrier_mat d d" and "sub3 d l (r,A) = (r'',A'')" and r0: "r \ 0" and "l \ d" shows "r * det A'' = r'' * det A" using assms proof - have d: "d = dim_row A" using A by auto show "sub3 d l (r,A) = (r'',A'') \ l \ d \ r * det A'' = r'' * det A" proof (induction l arbitrary: r'' A'') case (Suc l) let ?rA' = "sub3 d l (r,A)" obtain r' A' where rA':"?rA' = (r',A')" by force then have rA'': "sub2 d l (r',A') = (r'',A'')" using Suc by auto have A': "A' \ carrier_mat d d" using A rA' rA'' by auto have r'0: "r' \ 0" using r0 sub3_divisor[OF rA'] A Suc by auto have "r' * det A'' = r'' * det A'" using Suc r'0 A by(subst sub2_det[OF A' rA''],auto) also have "r * ... = r'' * (r * det A')" by auto also have "r * det A' = r' * det A" using Suc rA' by auto also have "r'' * ... div r' = r'' * r' * det A div r'" by (simp add: ac_simps) finally show "r * det A'' = r'' * det A" using r'0 by (metis \r * det A' = r' * det A\ \r' * det A'' = r'' * det A'\ mult.left_commute mult_cancel_left) qed simp qed theorem triangulize_det: assumes A: "A \ carrier_mat d d" and rA': "triangulize A = (r',A')" shows "det A * r' = det A'" proof - have rA'2: "sub3 d d (1,A) = (r',A')" using A rA' unfolding triangulize_def by auto show ?thesis proof (cases "d = 0") case True then have A': "A' \ carrier_mat 0 0" using A rA'2 by auto have rA'3: "(r',A') = (1,A)" using True rA'2 by simp thus ?thesis by auto next case False then show ?thesis using sub3_det[OF A rA'2] assms by auto qed qed end subsection \Determinant Computation\ definition det_code :: "'a mat \ 'a" where "det_code A = (if dim_row A = dim_col A then case triangulize A of (m,A') \ prod_list (diag_mat A') div m else 0)" lemma det_code[simp]: assumes sel_fun: "det_selection_fun sel_fun" and mf: "mute_fun mf" shows "det_code A = det A" using det_code_def[simp] proof (cases "dim_row A = dim_col A") case True then have A: "A \ carrier_mat (dim_row A) (dim_row A)" unfolding carrier_mat_def by auto obtain r' A' where rA': "triangulize A = (r',A')" by force from triangulize_divisor[OF mf sel_fun A] rA' have r'0: "r' \ 0" by auto from triangulize_det[OF mf sel_fun A rA'] have det': "det A * r' = det A'" by auto from triangulized[OF mf sel_fun A, unfolded rA'] have tri': "upper_triangular A'" by simp have A': "A' \ carrier_mat (dim_row A') (dim_row A')" using triangulize_closed[OF rA' A] by auto from tri' have tr: "triangular_to (dim_row A') A'" by auto have "det_code A = prod_list (diag_mat A') div r'" using rA' True by simp also have "prod_list (diag_mat A') = det A'" unfolding det_upper_triangular[OF tri' A'] .. also have "\ = det A * r'" by (simp add: det') also have "\ div r' = det A" using r'0 by auto finally show ?thesis . qed (simp add: det_def) end end text \Now we can select an arbitrary selection and mute function. This will be important for computing resultants over polynomials, where usually a polynomial with small degree is preferable. The default however is to use the first element.\ definition trivial_mute_fun :: "'a :: comm_ring_1 \ 'a \ 'a \ 'a \ 'a" where "trivial_mute_fun x y = (x,y,1)" lemma trivial_mute_fun[simp,intro]: "mute_fun trivial_mute_fun" unfolding mute_fun_def trivial_mute_fun_def by auto definition fst_sel_fun :: "'a det_selection_fun" where "fst_sel_fun x = fst (hd x)" lemma fst_sel_fun[simp]: "det_selection_fun fst_sel_fun" unfolding det_selection_fun_def fst_sel_fun_def by auto context fixes measure :: "'a \ nat" begin private fun select_min_main where "select_min_main m i ((j,p) # xs) = (let n = measure p in if n < m then select_min_main n j xs else select_min_main m i xs)" | "select_min_main m i [] = i" definition select_min :: "(nat \ 'a) list \ nat" where "select_min xs = (case xs of ((i,p) # ys) \ (select_min_main (measure p) i ys))" lemma select_min[simp]: "det_selection_fun select_min" unfolding det_selection_fun_def proof (intro allI impI) fix xs :: "(nat \ 'a)list" assume "xs \ []" then obtain i p ys where xs: "xs = ((i,p) # ys)" by (cases xs, auto) then obtain m where id: "select_min xs = select_min_main m i ys" unfolding select_min_def by auto have "i \ fst ` set xs" "set ys \ set xs" unfolding xs by auto thus "select_min xs \ fst ` set xs" unfolding id proof (induct ys arbitrary: m i ) case (Cons jp ys m i) obtain j p where jp: "jp = (j,p)" by force obtain k n where res: "select_min_main m i (jp # ys) = select_min_main n k ys" and k: "k \ fst ` set xs" using Cons(2-) unfolding jp by (cases "measure p < m"; force simp: Let_def) from Cons(1)[OF k, of n] Cons(3) show ?case unfolding res by auto qed simp qed end text \For the code equation we use the trivial mute and selection function as this does not impose any further class restrictions.\ lemma det_code_fst_sel_fun[code]: "det A = det_code fst_sel_fun trivial_mute_fun A" by simp text \But we also provide specialiced functions for more specific carriers.\ definition field_mute_fun :: "'a :: field \ 'a \ 'a \ 'a \ 'a" where "field_mute_fun x y = (x/y,1,y)" lemma field_mute_fun[simp,intro]: "mute_fun field_mute_fun" unfolding mute_fun_def field_mute_fun_def by auto definition det_field :: "'a :: field mat \ 'a" where "det_field A = det_code fst_sel_fun field_mute_fun A" lemma det_field[simp]: "det_field = det" by (intro ext, auto simp: det_field_def) definition gcd_mute_fun :: "'a :: ring_gcd \ 'a \ 'a \ 'a \ 'a" where "gcd_mute_fun x y = (let g = gcd x y in (x div g, y div g,g))" lemma gcd_mute_fun[simp,intro]: "mute_fun gcd_mute_fun" unfolding mute_fun_def gcd_mute_fun_def by (auto simp: Let_def div_mult_swap mult.commute) definition det_int :: "int mat \ int" where "det_int A = det_code (select_min (\ x. nat (abs x))) gcd_mute_fun A" lemma det_int[simp]: "det_int = det" by (intro ext, auto simp: det_int_def) -definition det_field_poly :: "'a :: {field,euclidean_ring_gcd} poly mat \ 'a poly" where +definition det_field_poly :: "'a :: {field,field_gcd} poly mat \ 'a poly" where "det_field_poly A = det_code (select_min degree) gcd_mute_fun A" lemma det_field_poly[simp]: "det_field_poly = det" by (intro ext, auto simp: det_field_poly_def) end diff --git a/thys/LLL_Basis_Reduction/Missing_Lemmas.thy b/thys/LLL_Basis_Reduction/Missing_Lemmas.thy --- a/thys/LLL_Basis_Reduction/Missing_Lemmas.thy +++ b/thys/LLL_Basis_Reduction/Missing_Lemmas.thy @@ -1,835 +1,835 @@ (* Authors: Jose DivasĂ³n Sebastiaan Joosten RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Missing lemmas\ text \This theory contains many results that are important but not specific for our development. They could be moved to the stardard library and some other AFP entries.\ theory Missing_Lemmas imports Berlekamp_Zassenhaus.Sublist_Iteration (* for thm upt_append *) Berlekamp_Zassenhaus.Square_Free_Int_To_Square_Free_GFp (* for thm large_mod_0 *) Algebraic_Numbers.Resultant Jordan_Normal_Form.Conjugate Jordan_Normal_Form.Missing_VectorSpace Jordan_Normal_Form.VS_Connect Berlekamp_Zassenhaus.Finite_Field_Factorization_Record_Based (* for transfer rules for thm vec_of_list_Nil *) Berlekamp_Zassenhaus.Berlekamp_Hensel (* for unique_factorization_m_factor *) begin no_notation test_bit (infixl "!!" 100) hide_const(open) module.smult up_ring.monom up_ring.coeff (**** Could be merged to HOL/Rings.thy ****) lemma (in zero_less_one) zero_le_one [simp]: "0 \ 1" by (rule less_imp_le, simp) subclass (in zero_less_one) zero_neq_one by (unfold_locales, simp add: less_imp_neq) class ordered_semiring_1 = Rings.ordered_semiring_0 + monoid_mult + zero_less_one begin subclass semiring_1.. lemma of_nat_ge_zero[intro!]: "of_nat n \ 0" using add_right_mono[of _ _ 1] by (induct n, auto) (* Following lemmas are moved from @{class ordered_idom}. *) lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" shows "1 < a * a ^ n" proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) from gt1 have "1 * 1 < a * 1" by simp also from gt1 have "\ \ a * a ^ n" by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" proof (induction N) case (Suc N) then have "a * a^N \ 1 * a^n" if "n \ N" using that by (intro mult_mono) auto then show ?case using Suc by (auto simp add: le_Suc_eq) qed (auto) lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induction N) case (Suc N) then have "1 * a^n \ a * a^N" if "n \ N" using that by (intro mult_mono) (auto simp add: order_trans[OF zero_le_one]) then show ?case using Suc by (auto simp add: le_Suc_eq) qed (auto) lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp end lemma prod_list_nonneg: "(\ x. (x :: 'a :: ordered_semiring_1) \ set xs \ x \ 0) \ prod_list xs \ 0" by (induct xs, auto) subclass (in ordered_idom) ordered_semiring_1 by unfold_locales auto (**** End of lemmas that could be moved to HOL/Rings.thy ****) (* missing lemma on logarithms *) lemma log_prod: assumes "0 < a" "a \ 1" "\ x. x \ X \ 0 < f x" shows "log a (prod f X) = sum (log a o f) X" using assms(3) proof (induct X rule: infinite_finite_induct) case (insert x F) have "log a (prod f (insert x F)) = log a (f x * prod f F)" using insert by simp also have "\ = log a (f x) + log a (prod f F)" by (rule log_mult[OF assms(1-2) insert(4) prod_pos], insert insert, auto) finally show ?case using insert by auto qed auto (* TODO: Jordan_Normal_Form/Missing_Ring.ordered_idom should be redefined *) subclass (in ordered_idom) zero_less_one by (unfold_locales, auto) hide_fact Missing_Ring.zero_less_one (**** The following lemmas could be part of the standard library ****) instance real :: ordered_semiring_strict by (intro_classes, auto) instance real :: linordered_idom.. (*This is a generalisation of thm less_1_mult*) lemma less_1_mult': fixes a::"'a::linordered_semidom" shows "1 < a \ 1 \ b \ 1 < a * b" by (metis le_less less_1_mult mult.right_neutral) lemma upt_minus_eq_append: "i\j \ i\j-k \ [i.. [0.. A" and ff: "\a. a \ A \ f (f a) = a" shows "bij_betw f A A" by (intro bij_betwI[OF f f], simp_all add: ff) lemma range_subsetI: assumes "\x. f x = g (h x)" shows "range f \ range g" using assms by auto lemma Gcd_uminus: fixes A::"int set" assumes "finite A" shows "Gcd A = Gcd (uminus ` A)" using assms by (induct A, auto) lemma aux_abs_int: fixes c :: int assumes "c \ 0" shows "\x\ \ \x * c\" proof - have "abs x = abs x * 1" by simp also have "\ \ abs x * abs c" by (rule mult_left_mono, insert assms, auto) finally show ?thesis unfolding abs_mult by auto qed lemma mod_0_abs_less_imp_0: fixes a::int assumes a1: "[a = 0] (mod m)" and a2: "abs(a)0" using assms by auto thus ?thesis using assms unfolding cong_def using int_mod_pos_eq large_mod_0 zless_imp_add1_zle by (metis abs_of_nonneg le_less not_less zabs_less_one_iff zmod_trival_iff) qed (* an intro version of sum_list_0 *) lemma sum_list_zero: assumes "set xs \ {0}" shows "sum_list xs = 0" using assms by (induct xs, auto) (* About @{const max} *) lemma max_idem [simp]: shows "max a a = a" by (simp add: max_def) lemma hom_max: assumes "a \ b \ f a \ f b" shows "f (max a b) = max (f a) (f b)" using assms by (auto simp: max_def) lemma le_max_self: fixes a b :: "'a :: preorder" assumes "a \ b \ b \ a" shows "a \ max a b" and "b \ max a b" using assms by (auto simp: max_def) lemma le_max: fixes a b :: "'a :: preorder" assumes "c \ a \ c \ b" and "a \ b \ b \ a" shows "c \ max a b" using assms(1) le_max_self[OF assms(2)] by (auto dest: order_trans) fun max_list where "max_list [] = (THE x. False)" (* more convenient than "undefined" *) | "max_list [x] = x" | "max_list (x # y # xs) = max x (max_list (y # xs))" declare max_list.simps(1) [simp del] declare max_list.simps(2-3)[code] lemma max_list_Cons: "max_list (x#xs) = (if xs = [] then x else max x (max_list xs))" by (cases xs, auto) lemma max_list_mem: "xs \ [] \ max_list xs \ set xs" by (induct xs, auto simp: max_list_Cons max_def) lemma mem_set_imp_le_max_list: fixes xs :: "'a :: preorder list" assumes "\a b. a \ set xs \ b \ set xs \ a \ b \ b \ a" and "a \ set xs" shows "a \ max_list xs" proof (insert assms, induct xs arbitrary:a) case Nil with assms show ?case by auto next case (Cons x xs) show ?case proof (cases "xs = []") case False have "x \ max_list xs \ max_list xs \ x" apply (rule Cons(2)) using max_list_mem[of xs] False by auto note 1 = le_max_self[OF this] from Cons have "a = x \ a \ set xs" by auto then show ?thesis proof (elim disjE) assume a: "a = x" show ?thesis by (unfold a max_list_Cons, auto simp: False intro!: 1) next assume "a \ set xs" then have "a \ max_list xs" by (intro Cons, auto) with 1 have "a \ max x (max_list xs)" by (auto dest: order_trans) then show ?thesis by (unfold max_list_Cons, auto simp: False) qed qed (insert Cons, auto) qed lemma le_max_list: fixes xs :: "'a :: preorder list" assumes ord: "\a b. a \ set xs \ b \ set xs \ a \ b \ b \ a" and ab: "a \ b" and b: "b \ set xs" shows "a \ max_list xs" proof- note ab also have "b \ max_list xs" by (rule mem_set_imp_le_max_list, fact ord, fact b) finally show ?thesis. qed lemma max_list_le: fixes xs :: "'a :: preorder list" assumes a: "\x. x \ set xs \ x \ a" and xs: "xs \ []" shows "max_list xs \ a" using max_list_mem[OF xs] a by auto lemma max_list_as_Greatest: assumes "\x y. x \ set xs \ y \ set xs \ x \ y \ y \ x" shows "max_list xs = (GREATEST a. a \ set xs)" proof (cases "xs = []") case True then show ?thesis by (unfold Greatest_def, auto simp: max_list.simps(1)) next case False from assms have 1: "x \ set xs \ x \ max_list xs" for x by (auto intro: le_max_list) have 2: "max_list xs \ set xs" by (fact max_list_mem[OF False]) have "\!x. x \ set xs \ (\y. y \ set xs \ y \ x)" (is "\!x. ?P x") proof (intro ex1I) from 1 2 show "?P (max_list xs)" by auto next fix x assume 3: "?P x" with 1 have "x \ max_list xs" by auto moreover from 2 3 have "max_list xs \ x" by auto ultimately show "x = max_list xs" by auto qed note 3 = theI_unique[OF this,symmetric] from 1 2 show ?thesis by (unfold Greatest_def Cons 3, auto) qed lemma hom_max_list_commute: assumes "xs \ []" and "\x y. x \ set xs \ y \ set xs \ h (max x y) = max (h x) (h y)" shows "h (max_list xs) = max_list (map h xs)" by (insert assms, induct xs, auto simp: max_list_Cons max_list_mem) (*Efficient rev [i.. nat \ nat list" ("(1[_>.._])") where rev_upt_0: "[0>..j] = []" | rev_upt_Suc: "[(Suc i)>..j] = (if i \ j then i # [i>..j] else [])" lemma rev_upt_rec: "[i>..j] = (if i>j then [i>..Suc j] @ [j] else [])" by (induct i, auto) definition rev_upt_aux :: "nat \ nat \ nat list \ nat list" where "rev_upt_aux i j js = [i>..j] @ js" lemma upt_aux_rec [code]: "rev_upt_aux i j js = (if j\i then js else rev_upt_aux i (Suc j) (j#js))" by (induct j, auto simp add: rev_upt_aux_def rev_upt_rec) lemma rev_upt_code[code]: "[i>..j] = rev_upt_aux i j []" by(simp add: rev_upt_aux_def) lemma upt_rev_upt: "rev [j>..i] = [i....i]" by (induct j, auto) lemma length_rev_upt [simp]: "length [i>..j] = i - j" by (induct i) (auto simp add: Suc_diff_le) lemma nth_rev_upt [simp]: "j + k < i \ [i>..j] ! k = i - 1 - k" proof - assume jk_i: "j + k < i" have "[i>..j] = rev [j....n]) ! i = f (m - 1 - i)" proof - have "(map f [m>..n]) ! i = f ([m>..n] ! i)" by (rule nth_map, auto simp add: i) also have "... = f (m - 1 - i)" proof (rule arg_cong[of _ _ f], rule nth_rev_upt) show "n + i < m" using i by linarith qed finally show ?thesis . qed lemma coeff_mult_monom: "coeff (p * monom a d) i = (if d \ i then a * coeff p (i - d) else 0)" using coeff_monom_mult[of a d p] by (simp add: ac_simps) (**** End of the lemmas which may be part of the standard library ****) (**** The following lemmas could be moved to Algebraic_Numbers/Resultant.thy ****) lemma vec_of_poly_0 [simp]: "vec_of_poly 0 = 0\<^sub>v 1" by (auto simp: vec_of_poly_def) lemma vec_index_vec_of_poly [simp]: "i \ degree p \ vec_of_poly p $ i = coeff p (degree p - i)" by (simp add: vec_of_poly_def Let_def) lemma poly_of_vec_vec: "poly_of_vec (vec n f) = Poly (rev (map f [0.. Suc) [0..) = Poly (rev (map (f \ Suc) [0.. = poly_of_vec (vec n (f \ Suc)) + monom (f 0) n" by (fold Suc, simp) also have "\ = poly_of_vec (vec (Suc n) f)" apply (unfold poly_of_vec_def Let_def dim_vec sum.lessThan_Suc) by (auto simp add: Suc_diff_Suc) finally show ?case.. qed lemma sum_list_map_dropWhile0: assumes f0: "f 0 = 0" shows "sum_list (map f (dropWhile ((=) 0) xs)) = sum_list (map f xs)" by (induct xs, auto simp add: f0) lemma coeffs_poly_of_vec: "coeffs (poly_of_vec v) = rev (dropWhile ((=) 0) (list_of_vec v))" proof- obtain n f where v: "v = vec n f" by transfer auto show ?thesis by (simp add: v poly_of_vec_vec) qed lemma poly_of_vec_vCons: "poly_of_vec (vCons a v) = monom a (dim_vec v) + poly_of_vec v" (is "?l = ?r") by (auto intro: poly_eqI simp: coeff_poly_of_vec vec_index_vCons) lemma poly_of_vec_as_Poly: "poly_of_vec v = Poly (rev (list_of_vec v))" by (induct v, auto simp:poly_of_vec_vCons Poly_snoc ac_simps) lemma poly_of_vec_add: assumes "dim_vec a = dim_vec b" shows "poly_of_vec (a + b) = poly_of_vec a + poly_of_vec b" using assms by (auto simp add: poly_eq_iff coeff_poly_of_vec) (*TODO: replace the one in Resultant.thy*) lemma degree_poly_of_vec_less: assumes "0 < dim_vec v" and "dim_vec v \ n" shows "degree (poly_of_vec v) < n" using degree_poly_of_vec_less assms by (auto dest: less_le_trans) lemma (in vec_module) poly_of_vec_finsum: assumes "f \ X \ carrier_vec n" shows "poly_of_vec (finsum V f X) = (\i\X. poly_of_vec (f i))" proof (cases "finite X") case False then show ?thesis by auto next case True show ?thesis proof (insert True assms, induct X rule: finite_induct) case IH: (insert a X) have [simp]: "f x \ carrier_vec n" if x: "x \ X" for x using x IH.prems unfolding Pi_def by auto have [simp]: "f a \ carrier_vec n" using IH.prems unfolding Pi_def by auto have [simp]: "dim_vec (finsum V f X) = n" by simp have [simp]: "dim_vec (f a) = n" by simp show ?case proof (cases "a \ X") case True then show ?thesis by (auto simp: insert_absorb IH) next case False then have "(finsum V f (insert a X)) = f a + (finsum V f X)" by (auto intro: finsum_insert IH) also have "poly_of_vec ... = poly_of_vec (f a) + poly_of_vec (finsum V f X)" by (rule poly_of_vec_add, simp) also have "... = (\i\insert a X. poly_of_vec (f i))" using IH False by (subst sum.insert, auto) finally show ?thesis . qed qed auto qed (*This function transforms a polynomial to a vector of dimension n*) definition "vec_of_poly_n p n = vec n (\i. if i < n - degree p - 1 then 0 else coeff p (n - i - 1))" (* TODO: make it abbreviation? *) lemma vec_of_poly_as: "vec_of_poly_n p (Suc (degree p)) = vec_of_poly p" by (induct p, auto simp: vec_of_poly_def vec_of_poly_n_def) lemma vec_of_poly_n_0 [simp]: "vec_of_poly_n p 0 = vNil" by (auto simp: vec_of_poly_n_def) lemma vec_dim_vec_of_poly_n [simp]: "dim_vec (vec_of_poly_n p n) = n" "vec_of_poly_n p n \ carrier_vec n" unfolding vec_of_poly_n_def by auto lemma dim_vec_of_poly [simp]: "dim_vec (vec_of_poly f) = degree f + 1" by (simp add: vec_of_poly_as[symmetric]) lemma vec_index_of_poly_n: assumes "i < n" shows "vec_of_poly_n p n $ i = (if i < n - Suc (degree p) then 0 else coeff p (n - i - 1))" using assms by (auto simp: vec_of_poly_n_def Let_def) lemma vec_of_poly_n_pCons[simp]: shows "vec_of_poly_n (pCons a p) (Suc n) = vec_of_poly_n p n @\<^sub>v vec_of_list [a]" (is "?l = ?r") proof (unfold vec_eq_iff, intro conjI allI impI) show "dim_vec ?l = dim_vec ?r" by auto show "i < dim_vec ?r \ ?l $ i = ?r $ i" for i by (cases "n - i", auto simp: coeff_pCons less_Suc_eq_le vec_index_of_poly_n) qed lemma vec_of_poly_pCons: shows "vec_of_poly (pCons a p) = (if p = 0 then vec_of_list [a] else vec_of_poly p @\<^sub>v vec_of_list [a])" by (cases "degree p", auto simp: vec_of_poly_as[symmetric]) lemma list_of_vec_of_poly [simp]: "list_of_vec (vec_of_poly p) = (if p = 0 then [0] else rev (coeffs p))" by (induct p, auto simp: vec_of_poly_pCons) lemma poly_of_vec_of_poly_n: assumes p: "degree p n" for i by (rule coeff_eq_0, insert i2 p, simp) ultimately show ?thesis using assms unfolding poly_eq_iff unfolding coeff_poly_of_vec by auto qed lemma vec_of_poly_n0[simp]: "vec_of_poly_n 0 n = 0\<^sub>v n" unfolding vec_of_poly_n_def by auto lemma vec_of_poly_n_add: "vec_of_poly_n (a + b) n = vec_of_poly_n a n + vec_of_poly_n b n" proof (induct n arbitrary: a b) case 0 then show ?case by auto next case (Suc n) then show ?case by (cases a, cases b, auto) qed lemma vec_of_poly_n_poly_of_vec: assumes n: "dim_vec g = n" shows "vec_of_poly_n (poly_of_vec g) n = g" proof (auto simp add: poly_of_vec_def vec_of_poly_n_def assms vec_eq_iff Let_def) have d: "degree (\ii degree (poly_of_vec g)" using n by linarith then show "g $ i = 0" using i1 i2 i3 by (metis (no_types, lifting) Suc_diff_Suc coeff_poly_of_vec diff_Suc_less diff_diff_cancel leD le_degree less_imp_le_nat n neq0_conv) next fix i assume "i < n" thus "coeff (\i\<^sub>v (vec_of_poly_n b n)) = smult a b" using assms by (auto simp add: poly_eq_iff coeff_poly_of_vec vec_of_poly_n_def coeff_eq_0) (*TODO: replace the one in Resultant.thy*) definition vec_of_poly_rev_shifted where "vec_of_poly_rev_shifted p n s j \ vec n (\i. if i \ j \ j \ s + i then coeff p (s + i - j) else 0)" lemma vec_of_poly_rev_shifted_dim[simp]: "dim_vec (vec_of_poly_rev_shifted p n s j) = n" unfolding vec_of_poly_rev_shifted_def by auto lemma col_sylvester_sub: (* TODO: from this directly derive col_sylvester *) assumes j: "j < m + n" shows "col (sylvester_mat_sub m n p q) j = vec_of_poly_rev_shifted p n m j @\<^sub>v vec_of_poly_rev_shifted q m n j" (is "?l = ?r") proof show "dim_vec ?l = dim_vec ?r" by simp fix i assume "i < dim_vec ?r" then have i: "i < m+n" by auto show "?l $ i = ?r $ i" unfolding vec_of_poly_rev_shifted_def apply (subst index_col) using i apply simp using j apply simp apply (subst sylvester_mat_sub_index) using i apply simp using j apply simp apply (cases "i < n") using i apply force using i apply (auto simp: not_less not_le intro!: coeff_eq_0) done qed lemma vec_of_poly_rev_shifted_scalar_prod: fixes p v defines "q \ poly_of_vec v" assumes m: "degree p \ m" and n: "dim_vec v = n" assumes j: "j < m+n" shows "vec_of_poly_rev_shifted p n m (n+m-Suc j) \ v = coeff (p * q) j" (is "?l = ?r") proof - have id1: "\ i. m + i - (n + m - Suc j) = i + Suc j - n" using j by auto let ?g = "\ i. if i \ n + m - Suc j \ n - Suc j \ i then coeff p (i + Suc j - n) * v $ i else 0" have "?thesis = ((\i = 0..i\j. coeff p i * (if j - i < n then v $ (n - Suc (j - i)) else 0)))" (is "_ = (?l = ?r)") unfolding vec_of_poly_rev_shifted_def coeff_mult m scalar_prod_def n q_def coeff_poly_of_vec by (subst sum.cong, insert id1, auto) also have "..." proof - have "?r = (\i\j. (if j - i < n then coeff p i * v $ (n - Suc (j - i)) else 0))" (is "_ = sum ?f _") by (rule sum.cong, auto) also have "sum ?f {..j} = sum ?f ({i. i \ j \ j - i < n} \ {i. i \ j \ \ j - i < n})" (is "_ = sum _ (?R1 \ ?R2)") by (rule sum.cong, auto) also have "\ = sum ?f ?R1 + sum ?f ?R2" by (subst sum.union_disjoint, auto) also have "sum ?f ?R2 = 0" by (rule sum.neutral, auto) also have "sum ?f ?R1 + 0 = sum (\ i. coeff p i * v $ (i + n - Suc j)) ?R1" (is "_ = sum ?F _") by (subst sum.cong, auto simp: ac_simps) also have "\ = sum ?F ((?R1 \ {..m}) \ (?R1 - {..m}))" (is "_ = sum _ (?R \ ?R')") by (rule sum.cong, auto) also have "\ = sum ?F ?R + sum ?F ?R'" by (subst sum.union_disjoint, auto) also have "sum ?F ?R' = 0" proof - { fix x assume "x > m" with m have "?F x = 0" by (subst coeff_eq_0, auto) } thus ?thesis by (subst sum.neutral, auto) qed finally have r: "?r = sum ?F ?R" by simp have "?l = sum ?g ({i. i < n \ i \ n + m - Suc j \ n - Suc j \ i} \ {i. i < n \ \ (i \ n + m - Suc j \ n - Suc j \ i)})" (is "_ = sum _ (?L1 \ ?L2)") by (rule sum.cong, auto) also have "\ = sum ?g ?L1 + sum ?g ?L2" by (subst sum.union_disjoint, auto) also have "sum ?g ?L2 = 0" by (rule sum.neutral, auto) also have "sum ?g ?L1 + 0 = sum (\ i. coeff p (i + Suc j - n) * v $ i) ?L1" (is "_ = sum ?G _") by (subst sum.cong, auto) also have "\ = sum ?G (?L1 \ {i. i + Suc j - n \ m} \ (?L1 - {i. i + Suc j - n \ m}))" (is "_ = sum _ (?L \ ?L')") by (subst sum.cong, auto) also have "\ = sum ?G ?L + sum ?G ?L'" by (subst sum.union_disjoint, auto) also have "sum ?G ?L' = 0" proof - { fix x assume "x + Suc j - n > m" with m have "?G x = 0" by (subst coeff_eq_0, auto) } thus ?thesis by (subst sum.neutral, auto) qed finally have l: "?l = sum ?G ?L" by simp let ?bij = "\ i. i + n - Suc j" { fix x assume x: "j < m + n" "Suc (x + j) - n \ m" "x < n" "n - Suc j \ x" define y where "y = x + Suc j - n" from x have "x + Suc j \ n" by auto with x have xy: "x = ?bij y" unfolding y_def by auto from x have y: "y \ ?R" unfolding y_def by auto have "x \ ?bij ` ?R" unfolding xy using y by blast } note tedious = this show ?thesis unfolding l r by (rule sum.reindex_cong[of ?bij], insert j, auto simp: inj_on_def tedious) qed finally show ?thesis by simp qed lemma sylvester_sub_poly: fixes p q :: "'a :: comm_semiring_0 poly" assumes m: "degree p \ m" assumes n: "degree q \ n" assumes v: "v \ carrier_vec (m+n)" shows "poly_of_vec ((sylvester_mat_sub m n p q)\<^sup>T *\<^sub>v v) = poly_of_vec (vec_first v n) * p + poly_of_vec (vec_last v m) * q" (is "?l = ?r") proof (rule poly_eqI) fix i let ?Tv = "(sylvester_mat_sub m n p q)\<^sup>T *\<^sub>v v" have dim: "dim_vec (vec_first v n) = n" "dim_vec (vec_last v m) = m" "dim_vec ?Tv = n + m" using v by auto have if_distrib: "\ x y z. (if x then y else (0 :: 'a)) * z = (if x then y * z else 0)" by auto show "coeff ?l i = coeff ?r i" proof (cases "i < m+n") case False hence i_mn: "i \ m+n" and i_n: "\x. x \ i \ x < n \ x < n" and i_m: "\x. x \ i \ x < m \ x < m" by auto have "coeff ?r i = (\ x < n. vec_first v n $ (n - Suc x) * coeff p (i - x)) + (\ x < m. vec_last v m $ (m - Suc x) * coeff q (i - x))" (is "_ = sum ?f _ + sum ?g _") unfolding coeff_add coeff_mult Let_def unfolding coeff_poly_of_vec dim if_distrib unfolding atMost_def apply(subst sum.inter_filter[symmetric],simp) apply(subst sum.inter_filter[symmetric],simp) unfolding mem_Collect_eq unfolding i_n i_m unfolding lessThan_def by simp also { fix x assume x: "x < n" have "coeff p (i-x) = 0" apply(rule coeff_eq_0) using i_mn x m by auto hence "?f x = 0" by auto } hence "sum ?f {..T *\<^sub>v v) $ (n + m - Suc i)" unfolding coeff_poly_of_vec dim sum.distrib[symmetric] by auto also have "... = coeff (p * poly_of_vec (vec_first v n) + q * poly_of_vec (vec_last v m)) i" apply(subst index_mult_mat_vec) using True apply simp apply(subst row_transpose) using True apply simp apply(subst col_sylvester_sub) using True apply simp apply(subst vec_first_last_append[of v n m,symmetric]) using v apply(simp add: add.commute) apply(subst scalar_prod_append) apply (rule carrier_vecI,simp)+ apply (subst vec_of_poly_rev_shifted_scalar_prod[OF m],simp) using True apply simp apply (subst add.commute[of n m]) apply (subst vec_of_poly_rev_shifted_scalar_prod[OF n]) apply simp using True apply simp by simp also have "... = (\x\i. (if x < n then vec_first v n $ (n - Suc x) else 0) * coeff p (i - x)) + (\x\i. (if x < m then vec_last v m $ (m - Suc x) else 0) * coeff q (i - x))" unfolding coeff_poly_of_vec[of "vec_first v n",unfolded dim_vec_first,symmetric] unfolding coeff_poly_of_vec[of "vec_last v m",unfolded dim_vec_last,symmetric] unfolding coeff_mult[symmetric] by (simp add: mult.commute) also have "... = coeff ?r i" unfolding coeff_add coeff_mult Let_def unfolding coeff_poly_of_vec dim.. finally show ?thesis. qed qed (**** End of the lemmas which could be moved to Algebraic_Numbers/Resultant.thy ****) (**** The following lemmas could be moved to Computational_Algebra/Polynomial.thy ****) lemma normalize_field [simp]: "normalize (a :: 'a :: {field, semiring_gcd}) = (if a = 0 then 0 else 1)" using unit_factor_normalize by fastforce lemma content_field [simp]: "content (p :: 'a :: {field,semiring_gcd} poly) = (if p = 0 then 0 else 1)" by (induct p, auto simp: content_def) lemma primitive_part_field [simp]: "primitive_part (p :: 'a :: {field,semiring_gcd} poly) = p" by (cases "p = 0", auto intro!: primitive_part_prim) lemma primitive_part_dvd: "primitive_part a dvd a" by (metis content_times_primitive_part dvd_def dvd_refl mult_smult_right) lemma degree_abs [simp]: "degree \p\ = degree p" by (auto simp: abs_poly_def) lemma degree_gcd1: assumes a_not0: "a \ 0" shows "degree (gcd a b) \ degree a" proof - let ?g = "gcd a b" have gcd_dvd_b: "?g dvd a" by simp from this obtain c where a_gc: "a = ?g * c" unfolding dvd_def by auto have g_not0: "?g \0" using a_not0 a_gc by auto have c0: "c \ 0" using a_not0 a_gc by auto have "degree ?g \ degree (?g * c)" by (rule degree_mult_right_le[OF c0]) also have "... = degree a" using a_gc by auto finally show ?thesis . qed lemma primitive_part_neg [simp]: - fixes a::"'a :: factorial_ring_gcd poly" + fixes a::"'a :: {factorial_ring_gcd,factorial_semiring_multiplicative} poly" shows "primitive_part (-a) = - primitive_part a" proof - have "primitive_part (-a) = primitive_part (smult (-1) a)" by auto then show ?thesis unfolding primitive_part_smult by (simp add: is_unit_unit_factor) qed lemma content_uminus[simp]: fixes f::"int poly" shows "content (-f) = content f" proof - have "-f = - (smult 1 f)" by auto also have "... = smult (-1) f" using smult_minus_left by auto finally have "content (-f) = content (smult (-1) f)" by auto also have "... = normalize (- 1) * content f" unfolding content_smult .. finally show ?thesis by auto qed lemma pseudo_mod_monic: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes monic_g: "monic g" shows "\q. 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 have g: "g \ 0" using monic_g by 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) have a1: "a = 1" unfolding a_def using monic_g by auto hence id2: "f = g * q + r" using id by auto show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id2 show "\q. f = g * q + r" by auto qed lemma monic_imp_div_mod_int_poly_degree: fixes p :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" assumes m: "monic u" shows "\q r. p = q*u + r \ (r = 0 \ degree r < degree u)" using pseudo_mod_monic[OF m] using mult.commute by metis corollary monic_imp_div_mod_int_poly_degree2: fixes p :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" assumes m: "monic u" and deg_u: "degree u > 0" shows "\q r. p = q*u + r \ (degree r < degree u)" proof - obtain q r where "p = q * u + r" and r: "(r = 0 \ degree r < degree u)" using monic_imp_div_mod_int_poly_degree[OF m, of p] by auto moreover have "degree r < degree u" using deg_u r by auto ultimately show ?thesis by auto qed (**** End of the lemmas that could be moved to Computational_Algebra/Polynomial.thy ****) (* To be categorized *) lemma (in zero_hom) hom_upper_triangular: "A \ carrier_mat n n \ upper_triangular A \ upper_triangular (map_mat hom A)" by (auto simp: upper_triangular_def) end diff --git a/thys/Linear_Recurrences/Factorizations.thy b/thys/Linear_Recurrences/Factorizations.thy --- a/thys/Linear_Recurrences/Factorizations.thy +++ b/thys/Linear_Recurrences/Factorizations.thy @@ -1,253 +1,253 @@ (* File: Factorizations.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Factorizations of polynomials\ theory Factorizations imports Complex_Main Linear_Recurrences_Misc "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin text \ We view a factorisation of a polynomial as a pair consisting of the leading coefficient and a list of roots with multiplicities. This gives us a factorization into factors of the form $(X - c) ^ {n+1}$. \ definition interp_factorization where "interp_factorization = (\(a,cs). Polynomial.smult a (\(c,n)\cs. [:-c,1:] ^ Suc n))" text \ An alternative way to factorise is as a pair of the leading coefficient and factors of the form $(1 - cX) ^ {n+1}$. \ definition interp_alt_factorization where "interp_alt_factorization = (\(a,cs). Polynomial.smult a (\(c,n)\cs. [:1,-c:] ^ Suc n))" definition is_factorization_of where "is_factorization_of fctrs p = (interp_factorization fctrs = p \ distinct (map fst (snd fctrs)))" definition is_alt_factorization_of where "is_alt_factorization_of fctrs p = (interp_alt_factorization fctrs = p \ 0 \ set (map fst (snd fctrs)) \ distinct (map fst (snd fctrs)))" text \ Regular and alternative factorisations are related by reflecting the polynomial. \ lemma interp_factorization_reflect: assumes "(0::'a::idom) \ fst ` set (snd fctrs)" shows "reflect_poly (interp_factorization fctrs) = interp_alt_factorization fctrs" proof - have "reflect_poly (interp_factorization fctrs) = Polynomial.smult (fst fctrs) (\x\snd fctrs. reflect_poly [:- fst x, 1:] ^ Suc (snd x))" by (simp add: interp_factorization_def interp_alt_factorization_def case_prod_unfold reflect_poly_smult reflect_poly_prod_list reflect_poly_power o_def del: power_Suc) also have "map (\x. reflect_poly [:- fst x, 1:] ^ Suc (snd x)) (snd fctrs) = map (\x. [:1, - fst x:] ^ Suc (snd x)) (snd fctrs)" using assms by (intro list.map_cong0, subst reflect_poly_pCons) auto also have "Polynomial.smult (fst fctrs) (prod_list \) = interp_alt_factorization fctrs" by (simp add: interp_alt_factorization_def case_prod_unfold) finally show ?thesis . qed lemma interp_alt_factorization_reflect: assumes "(0::'a::idom) \ fst ` set (snd fctrs)" shows "reflect_poly (interp_alt_factorization fctrs) = interp_factorization fctrs" proof - have "reflect_poly (interp_alt_factorization fctrs) = Polynomial.smult (fst fctrs) (\x\snd fctrs. reflect_poly [:1, - fst x:] ^ Suc (snd x))" by (simp add: interp_factorization_def interp_alt_factorization_def case_prod_unfold reflect_poly_smult reflect_poly_prod_list reflect_poly_power o_def del: power_Suc) also have "map (\x. reflect_poly [:1, - fst x:] ^ Suc (snd x)) (snd fctrs) = map (\x. [:- fst x, 1:] ^ Suc (snd x)) (snd fctrs)" proof (intro list.map_cong0, clarsimp simp del: power_Suc, goal_cases) fix c n assume "(c, n) \ set (snd fctrs)" with assms have "c \ 0" by force thus "reflect_poly [:1, -c:] ^ Suc n = [:-c, 1:] ^ Suc n" by (simp add: reflect_poly_pCons del: power_Suc) qed also have "Polynomial.smult (fst fctrs) (prod_list \) = interp_factorization fctrs" by (simp add: interp_factorization_def case_prod_unfold) finally show ?thesis . qed lemma coeff_0_interp_factorization: "coeff (interp_factorization fctrs) 0 = (0 :: 'a :: idom) \ fst fctrs = 0 \ 0 \ fst ` set (snd fctrs)" by (force simp: interp_factorization_def case_prod_unfold coeff_0_prod_list o_def coeff_0_power prod_list_zero_iff simp del: power_Suc) lemma reflect_factorization: assumes "coeff p 0 \ (0::'a::idom)" assumes "is_factorization_of fctrs p" shows "is_alt_factorization_of fctrs (reflect_poly p)" using assms by (force simp: interp_factorization_reflect is_factorization_of_def is_alt_factorization_of_def coeff_0_interp_factorization) lemma reflect_factorization': assumes "coeff p 0 \ (0::'a::idom)" assumes "is_alt_factorization_of fctrs p" shows "is_factorization_of fctrs (reflect_poly p)" using assms by (force simp: interp_alt_factorization_reflect is_factorization_of_def is_alt_factorization_of_def coeff_0_interp_factorization) lemma zero_in_factorization_iff: assumes "is_factorization_of fctrs p" shows "coeff p 0 = 0 \ p = 0 \ (0::'a::idom) \ fst ` set (snd fctrs)" proof (cases "p = 0") assume "p \ 0" with assms have [simp]: "fst fctrs \ 0" by (auto simp: is_factorization_of_def interp_factorization_def case_prod_unfold) from assms have "p = interp_factorization fctrs" by (simp add: is_factorization_of_def) also have "coeff \ 0 = 0 \ 0 \ fst ` set (snd fctrs)" by (force simp add: interp_factorization_def case_prod_unfold coeff_0_prod_list prod_list_zero_iff o_def coeff_0_power) finally show ?thesis using \p \ 0\ by blast next assume p: "p = 0" with assms have "interp_factorization fctrs = 0" by (simp add: is_factorization_of_def) also have "interp_factorization fctrs = 0 \ fst fctrs = 0 \ (\(c,n)\snd fctrs. [:-c,1:]^Suc n) = 0" by (simp add: interp_factorization_def case_prod_unfold) also have "(\(c,n)\snd fctrs. [:-c,1:]^Suc n) = 0 \ False" by (auto simp: prod_list_zero_iff simp del: power_Suc) finally show ?thesis by (simp add: \p = 0\) qed lemma poly_prod_list [simp]: "poly (prod_list ps) x = prod_list (map (\p. poly p x) ps)" by (induction ps) auto lemma is_factorization_of_roots: fixes a :: "'a :: idom" assumes "is_factorization_of (a, fctrs) p" "p \ 0" shows "set (map fst fctrs) = {x. poly p x = 0}" using assms by (force simp: is_factorization_of_def interp_factorization_def o_def case_prod_unfold prod_list_zero_iff simp del: power_Suc) lemma (in monoid_mult) prod_list_prod_nth: "prod_list xs = (\ix. x \ A \ f x \ 0" assumes "\x y. x \ A \ y \ A \ x \ y \ coprime (f x) (f y)" shows "order c (prod f A) = (\x\A. order c (f x))" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert.hyps have "order c (prod f (insert x A)) = order c (f x * prod f A)" by simp also have "\ = order c (f x) + order c (prod f A)" using insert.prems and insert.hyps by (intro order_mult) auto also have "order c (prod f A) = (\x\A. order c (f x))" using insert.prems and insert.hyps by (intro insert.IH) auto finally show ?case using insert.hyps by simp qed auto lemma is_factorization_of_order: - fixes p :: "'a :: {field,factorial_ring_gcd} poly" + fixes p :: "'a :: field_gcd poly" assumes "p \ 0" assumes "is_factorization_of (a, fctrs) p" assumes "(c, n) \ set fctrs" shows "order c p = Suc n" proof - from assms have distinct: "distinct (map fst (fctrs))" by (simp add: is_factorization_of_def) from assms have [simp]: "a \ 0" by (auto simp: is_factorization_of_def interp_factorization_def) from assms(2) have "p = interp_factorization (a, fctrs)" unfolding is_factorization_of_def by simp also have "order c \ = order c (\(c,n)\fctrs. [:-c, 1:] ^ Suc n)" unfolding interp_factorization_def by (simp add: order_smult) also have "(\(c,n)\fctrs. [:-c, 1:] ^ Suc n) = (\i\{.. = (\x {.. 0" by (simp only: power_eq_0_iff) simp next fix i j :: nat assume "i \ j" "i \ {.. {.. fst (fctrs ! j)" using nth_eq_iff_index_eq [OF distinct, of i j] by simp then show "coprime ([:- fst (fctrs ! i), 1:] ^ Suc (snd (fctrs ! i))) ([:- fst (fctrs ! j), 1:] ^ Suc (snd (fctrs ! j)))" by (simp only: coprime_power_left_iff coprime_power_right_iff) (auto simp add: coprime_linear_poly) qed also have "\ = (\(c',n')\fctrs. order c ([:-c', 1:] ^ Suc n'))" by (simp add: sum_list_sum_nth case_prod_unfold atLeast0LessThan) also have "\ = (\(c',n')\fctrs. if c = c' then Suc n' else 0)" by (intro arg_cong[OF map_cong]) (auto simp add: order_power_n_n order_0I simp del: power_Suc) also have "\ = (\x\fctrs. if x = (c, n) then Suc (snd x) else 0)" using distinct assms by (intro arg_cong[OF map_cong]) (force simp: distinct_map inj_on_def)+ also from distinct have "\ = (\x\set fctrs. if x = (c, n) then Suc (snd x) else 0)" by (intro sum_list_distinct_conv_sum_set) (simp_all add: distinct_map) also from assms have "\ = Suc n" by simp finally show ?thesis . qed text \ For complex polynomials, a factorisation in the above sense always exists. \ lemma complex_factorization_exists: "\fctrs. is_factorization_of fctrs (p :: complex poly)" proof (cases "p = 0") case True thus ?thesis by (intro exI[of _ "(0, [])"]) (auto simp: is_factorization_of_def interp_factorization_def) next case False hence "\xs. set xs = {x. poly p x = 0} \ distinct xs" by (intro finite_distinct_list poly_roots_finite) then obtain xs where [simp]: "set xs = {x. poly p x = 0}" "distinct xs" by blast have "interp_factorization (lead_coeff p, map (\x. (x, order x p - 1)) xs) = smult (lead_coeff p) (\x\xs. [:- x, 1:] ^ Suc (order x p - 1))" by (simp add: interp_factorization_def o_def) also have "(\x\xs. [:- x, 1:] ^ Suc (order x p - 1)) = (\x|poly p x = 0. [:- x, 1:] ^ Suc (order x p - 1))" by (subst prod.distinct_set_conv_list [symmetric]) simp_all also have "\ = (\x|poly p x = 0. [:- x, 1:] ^ order x p)" proof (intro prod.cong refl, goal_cases) case (1 x) with False have "order x p \ 0" by (subst (asm) order_root) auto hence *: "Suc (order x p - 1) = order x p" by simp show ?case by (simp only: *) qed also have "smult (lead_coeff p) \ = p" by (rule complex_poly_decompose) finally have "is_factorization_of (lead_coeff p, map (\x. (x, order x p - 1)) xs) p" by (auto simp: is_factorization_of_def o_def) thus ?thesis .. qed text \ By reflecting the polynomial, this means that for complex polynomials with non-zero constant coefficient, the alternative factorisation also exists. \ corollary complex_alt_factorization_exists: assumes "coeff p 0 \ 0" shows "\fctrs. is_alt_factorization_of fctrs (p :: complex poly)" proof - from assms have "coeff (reflect_poly p) 0 \ 0" by auto moreover from complex_factorization_exists [of "reflect_poly p"] obtain fctrs where "is_factorization_of fctrs (reflect_poly p)" .. ultimately have "is_alt_factorization_of fctrs (reflect_poly (reflect_poly p))" by (rule reflect_factorization) also from assms have "reflect_poly (reflect_poly p) = p" by simp finally show ?thesis .. qed end \ No newline at end of file diff --git a/thys/Linear_Recurrences/Linear_Homogenous_Recurrences.thy b/thys/Linear_Recurrences/Linear_Homogenous_Recurrences.thy --- a/thys/Linear_Recurrences/Linear_Homogenous_Recurrences.thy +++ b/thys/Linear_Recurrences/Linear_Homogenous_Recurrences.thy @@ -1,295 +1,295 @@ (* File: Linear_Homogenous_Recurrences.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Homogenous linear recurrences\ theory Linear_Homogenous_Recurrences imports Complex_Main RatFPS Rational_FPS_Solver Linear_Recurrences_Common begin text \ The following is the numerator of the rational generating function of a linear homogenous recurrence. \ definition lhr_fps_numerator where "lhr_fps_numerator m cs f = (let N = length cs - 1 in Poly [(\i\min N k. cs ! (N - i) * f (k - i)). k \ [0..i\min N k. cs ! (N - i) * f (k - i)). k \ [0.. 'a :: field" assumes "\n. n \ m \ (\k\N. c k * f (n + k)) = 0" assumes cN: "c N \ 0" defines "p \ Poly [c (N - k). k \ [0.. Poly [(\i\min N k. c (N - i) * f (k - i)). k \ [0.. N + m") case True let ?f = "\i. N - i" have "(fps_of_poly p * F) $ n = (\i\n. coeff p i * f (n - i))" by (simp add: fps_mult_nth atLeast0AtMost) also from True have "\ = (\i\N. coeff p i * f (n - i))" by (intro sum.mono_neutral_right) (auto simp: nth_default_def p_def) also have "\ = (\i\N. c (N - i) * f (n - i))" by (intro sum.cong) (auto simp: nth_default_def p_def simp del: upt_Suc) also from True have "\ = (\i\N. c i * f (n - N + i))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) auto also from True have "\ = 0" by (intro assms) simp_all also from True have "\ = coeff q n" by (simp add: q_def nth_default_def del: upt_Suc) finally show ?thesis . next case False hence "(fps_of_poly p * F) $ n = (\i\n. coeff p i * f (n - i))" by (simp add: fps_mult_nth atLeast0AtMost) also have "\ = (\i\min N n. coeff p i * f (n - i))" by (intro sum.mono_neutral_right) (auto simp: p_def nth_default_def simp del: upt_Suc) also have "\ = (\i\min N n. c (N - i) * f (n - i))" by (intro sum.cong) (simp_all add: p_def nth_default_def del: upt_Suc) also from False have "\ = coeff q n" by (simp add: q_def nth_default_def) finally show ?thesis . qed hence "fps_of_poly p * F = fps_of_poly q" by (intro fps_ext) simp with cN show "F = fps_of_poly q / fps_of_poly p" by (subst unit_eq_div2) (simp_all add: mult_ac) qed lemma lhr_fps: fixes f :: "nat \ 'a :: field" and cs :: "'a list" defines "N \ length cs - 1" assumes cs: "cs \ []" assumes "\n. n \ m \ (\k\N. cs ! k * f (n + k)) = 0" assumes cN: "last cs \ 0" shows "Abs_fps f = fps_of_poly (lhr_fps_numerator m cs f) / fps_of_poly (lr_fps_denominator cs)" proof - define p and q where "p = Poly (map (\k. \i\min N k. cs ! (N - i) * f (k - i)) [0..k. cs ! (N - k)) [0.. last cs = 0 \ length fs < length cs - 1 then undefined else (if n < length fs then fs ! n else (\k []" "last cs \ 0" "length fs \ length cs - 1" "n \ length fs" shows "(\kk\\ . cs ! k * lhr cs fs (n + 1 - length cs + k)) = (\k = 0" by (subst (2) lhr.simps) (simp_all add: field_simps) finally show ?thesis . qed lemma lhrI: assumes "cs \ []" "last cs \ 0" "length fs \ length cs - 1" assumes "\n. n < length fs \ f n = fs ! n" assumes "\n. n \ length fs \ (\kkk\\ . cs ! k * f (n + 1 - length cs + k)) = (\kkkklast cs \ 0\ by (simp add: field_simps eq_neg_iff_add_eq_0) also from 1(2-4) False have "\ = lhr cs fs n" by (subst lhr.simps) simp finally show ?thesis . qed (insert 1(2-5), simp add: lhr.simps) qed (* END TODO *) locale linear_homogenous_recurrence = fixes f :: "nat \ 'a :: comm_semiring_0" and cs fs :: "'a list" assumes base: "n < length fs \ f n = fs ! n" assumes cs_not_null [simp]: "cs \ []" and last_cs [simp]: "last cs \ 0" and hd_cs [simp]: "hd cs \ 0" and enough_base: "length fs + 1 \ length cs" assumes rec: "n \ length fs - length cs \ (\kk. (\i\min N k. cs ! (N - i) * f (k - i))) [0.. {0.. min N k" for i using enough_base that by (intro base) (auto simp: Suc_le_eq N_def m_def algebra_simps) hence "(\i\min N k. cs ! (N - i) * f (k - i)) = (\i\min N k. cs ! (N - i) * fs ! (k - i))" by simp } hence "map (\k. (\i\min N k. cs ! (N - i) * f (k - i))) [0..k. (\i\min N k. cs ! (N - i) * fs ! (k - i))) [0.. = lhr_fps_numerator m cs ((!) fs)" using enough_base by (cases cs) (simp_all add: lhr_fps_numerator_def Let_def m_def N_def) finally show ?thesis unfolding m_def . qed end (* TODO Duplication *) lemma solve_lhr_aux: assumes "linear_homogenous_recurrence f cs fs" assumes "is_factorization_of fctrs (lr_fps_denominator' cs)" shows "f = interp_ratfps_solution (solve_factored_ratfps' (lhr_fps_numerator (length fs + 1 - length cs) cs ((!) fs)) fctrs)" proof - interpret linear_homogenous_recurrence f cs fs by fact note assms(2) hence "is_alt_factorization_of fctrs (reflect_poly (lr_fps_denominator' cs))" by (intro reflect_factorization) (simp_all add: lr_fps_denominator'_def nth_default_def hd_conv_nth [symmetric]) also have "reflect_poly (lr_fps_denominator' cs) = lr_fps_denominator cs" unfolding lr_fps_denominator_def lr_fps_denominator'_def by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly strip_while_rev [symmetric] no_trailing_unfold last_rev del: strip_while_rev) finally have factorization: "is_alt_factorization_of fctrs (lr_fps_denominator cs)" . define m where "m = length fs + 1 - length cs" obtain a ds where fctrs: "fctrs = (a, ds)" by (cases fctrs) simp_all define p and p' where "p = lhr_fps_numerator m cs ((!) fs)" and "p' = smult (inverse a) p" obtain b es where sol: "solve_factored_ratfps' p fctrs = (b, es)" by (cases "solve_factored_ratfps' p fctrs") simp_all have sol': "(b, es) = solve_factored_ratfps p' ds" by (subst sol [symmetric]) (simp add: fctrs p'_def solve_factored_ratfps_def solve_factored_ratfps'_def case_prod_unfold) have factorization': "lr_fps_denominator cs = interp_alt_factorization fctrs" using factorization by (simp add: is_alt_factorization_of_def) from assms(2) have distinct: "distinct (map fst ds)" by (simp add: fctrs is_factorization_of_def) have coeff_0_denom: "coeff (lr_fps_denominator cs) 0 \ 0" by (simp add: lr_fps_denominator_def nth_default_def hd_conv_nth [symmetric] hd_rev) have "coeff (lr_fps_denominator' cs) 0 \ 0" by (simp add: lr_fps_denominator'_def nth_default_def hd_conv_nth [symmetric]) with assms(2) have no_zero: "0 \ fst ` set ds" by (simp add: zero_in_factorization_iff fctrs) from assms(2) have a_nz [simp]: "a \ 0" by (auto simp: fctrs interp_factorization_def is_factorization_of_def lr_fps_denominator'_nz) hence unit1: "is_unit (fps_const a)" by simp moreover have "is_unit (fps_of_poly (interp_alt_factorization fctrs))" by (simp add: coeff_0_denom factorization' [symmetric]) ultimately have unit2: "is_unit (fps_of_poly (\p\ds. [:1, - fst p:] ^ Suc (snd p)))" by (simp add: fctrs case_prod_unfold interp_alt_factorization_def del: power_Suc) have "Abs_fps f = fps_of_poly (lhr_fps_numerator m cs f) / fps_of_poly (lr_fps_denominator cs)" proof (intro lhr_fps) fix n assume n: "n \ m" have "{..length cs - 1} = {..k\\ . cs ! k * f (n + k)) = 0" by (intro rec) (simp_all add: m_def algebra_simps) finally show "(\k\length cs - 1. cs ! k * f (n + k)) = 0" . qed (simp_all add: m_def) also have "lhr_fps_numerator m cs f = lhr_fps_numerator m cs ((!) fs)" unfolding lhr_fps_numerator_def using enough_base by (auto simp: Let_def poly_eq_iff nth_default_def base m_def Suc_le_eq intro!: sum.cong) also have "fps_of_poly \ / fps_of_poly (lr_fps_denominator cs) = fps_of_poly (lhr_fps_numerator m cs ((!) fs)) / (fps_const (fst fctrs) * fps_of_poly (\p\snd fctrs. [:1, - fst p:] ^ Suc (snd p)))" unfolding assms factorization' interp_alt_factorization_def by (simp add: case_prod_unfold Let_def fps_of_poly_smult) also from unit1 unit2 have "\ = fps_of_poly p / fps_const a / fps_of_poly (\(c,n)\ds. [:1, -c:]^Suc n)" by (subst is_unit_div_mult2_eq) (simp_all add: fctrs case_prod_unfold p_def) also from unit1 have "fps_of_poly p / fps_const a = fps_of_poly p'" by (simp add: fps_divide_unit fps_of_poly_smult fps_const_inverse p'_def) also from distinct no_zero have "\ / fps_of_poly (\(c,n)\ds. [:1, -c:]^Suc n) = Abs_fps (interp_ratfps_solution (solve_factored_ratfps' p fctrs))" by (subst solve_factored_ratfps) (simp_all add: case_prod_unfold sol' sol) finally show ?thesis unfolding p_def m_def by (intro ext) (simp add: fps_eq_iff) qed definition "lhr_fps as fs = ( let m = length fs + 1 - length as; p = lhr_fps_numerator m as (\n. fs ! n); q = lr_fps_denominator as in ratfps_of_poly p / ratfps_of_poly q)" lemma lhr_fps_correct: - fixes f :: "nat \ 'a :: {field_char_0,factorial_ring_gcd}" + fixes f :: "nat \ 'a :: {field_char_0,field_gcd}" assumes "linear_homogenous_recurrence f cs fs" shows "fps_of_ratfps (lhr_fps cs fs) = Abs_fps f" proof - interpret linear_homogenous_recurrence f cs fs by fact define m where "m = length fs + 1 - length cs" let ?num = "lhr_fps_numerator m cs f" let ?num' = "lhr_fps_numerator m cs ((!) fs)" let ?denom = "lr_fps_denominator cs" have "{..length cs - 1} = {.. 1" by (cases cs) auto ultimately have "Abs_fps f = fps_of_poly ?num / fps_of_poly ?denom" by (intro lhr_fps) (insert rec, simp_all add: m_def) also have "?num = ?num'" by (rule lhr_fps_numerator_altdef [folded m_def]) also have "fps_of_poly ?num' / fps_of_poly ?denom = fps_of_ratfps (ratfps_of_poly ?num' / ratfps_of_poly ?denom)" by simp also from enough_base have "\ = fps_of_ratfps (lhr_fps cs fs)" by (cases cs) (simp_all add: base fps_of_ratfps_def case_prod_unfold lhr_fps_def m_def) finally show ?thesis .. qed end diff --git a/thys/Linear_Recurrences/Linear_Inhomogenous_Recurrences.thy b/thys/Linear_Recurrences/Linear_Inhomogenous_Recurrences.thy --- a/thys/Linear_Recurrences/Linear_Inhomogenous_Recurrences.thy +++ b/thys/Linear_Recurrences/Linear_Inhomogenous_Recurrences.thy @@ -1,241 +1,241 @@ (* File: Linear_Inhomogenous_Recurrences.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Inhomogenous linear recurrences\ theory Linear_Inhomogenous_Recurrences imports Complex_Main Linear_Homogenous_Recurrences Eulerian_Polynomials RatFPS begin definition lir_fps_numerator where "lir_fps_numerator m cs f g = (let N = length cs - 1 in Poly [(\i\min N k. cs ! (N - i) * f (k - i)) - g k. k \ [0..i\min N k. cs ! (N - i) * f (k - i)) - g k. k \ [0.. 'a :: comm_ring" and cs fs :: "'a list" assumes base: "n < length fs \ f n = fs ! n" assumes cs_not_null [simp]: "cs \ []" and last_cs [simp]: "last cs \ 0" and hd_cs [simp]: "hd cs \ 0" and enough_base: "length fs + 1 \ length cs" assumes rec: "n \ length fs + 1 - length cs \ (\kk. (\i\min N k. cs ! (N - i) * f (k - i)) - g k) [0.. {0.. min N k" for i using enough_base that by (intro base) (auto simp: Suc_le_eq N_def m_def algebra_simps) hence "(\i\min N k. cs ! (N - i) * f (k - i)) = (\i\min N k. cs ! (N - i) * fs ! (k - i))" by simp } hence "map (\k. (\i\min N k. cs ! (N - i) * f (k - i)) - g k) [0..k. (\i\min N k. cs ! (N - i) * fs ! (k - i)) - g k) [0.. = lir_fps_numerator m cs ((!) fs) g" using enough_base by (cases cs) (simp_all add: lir_fps_numerator_def Let_def m_def N_def) finally show ?thesis unfolding m_def . qed end context begin private lemma lir_fps_aux: fixes f :: "nat \ 'a :: field" assumes rec: "\n. n \ m \ (\k\N. c k * f (n + k)) = g (n + N)" assumes cN: "c N \ 0" defines "p \ Poly [c (N - k). k \ [0.. Poly [(\i\min N k. c (N - i) * f (k - i)) - g k. k \ [0.. N + m") case True let ?f = "\i. N - i" have "(fps_of_poly p * F) $ n = (\i\n. coeff p i * f (n - i))" by (simp add: fps_mult_nth atLeast0AtMost) also from True have "\ = (\i\N. coeff p i * f (n - i))" by (intro sum.mono_neutral_right) (auto simp: nth_default_def p_def) also have "\ = (\i\N. c (N - i) * f (n - i))" by (intro sum.cong) (auto simp: nth_default_def p_def simp del: upt_Suc) also from True have "\ = (\i\N. c i * f (n - N + i))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) auto also from True have "\ = g (n - N + N)" by (intro rec) simp_all also from True have "\ = coeff q n + g n" by (simp add: q_def nth_default_def del: upt_Suc) finally show ?thesis . next case False hence "(fps_of_poly p * F) $ n = (\i\n. coeff p i * f (n - i))" by (simp add: fps_mult_nth atLeast0AtMost) also have "\ = (\i\min N n. coeff p i * f (n - i))" by (intro sum.mono_neutral_right) (auto simp: p_def nth_default_def simp del: upt_Suc) also have "\ = (\i\min N n. c (N - i) * f (n - i))" by (intro sum.cong) (simp_all add: p_def nth_default_def del: upt_Suc) also from False have "\ = coeff q n + g n" by (simp add: q_def nth_default_def) finally show ?thesis . qed hence "fps_of_poly p * F = fps_of_poly q + Abs_fps g" by (intro fps_ext) (simp add:) with cN show "F = (fps_of_poly q + Abs_fps g) / fps_of_poly p" by (subst unit_eq_div2) (simp_all add: mult_ac) qed lemma lir_fps: fixes f g :: "nat \ 'a :: field" and cs :: "'a list" defines "N \ length cs - 1" assumes cs: "cs \ []" assumes "\n. n \ m \ (\k\N. cs ! k * f (n + k)) = g (n + N)" assumes cN: "last cs \ 0" shows "Abs_fps f = (fps_of_poly (lir_fps_numerator m cs f g) + Abs_fps g) / fps_of_poly (lr_fps_denominator cs)" proof - define p and q where "p = Poly [(\i\min N k. cs ! (N - i) * f (k - i)) - g k. k \ [0..k. cs ! (N - k)) [0.. nat \ 'a" where "eval_polyexp xs = (\n. \(a,k,b)\xs. a * of_nat n ^ k * b ^ n)" lemma eval_polyexp_Nil [simp]: "eval_polyexp [] = (\_. 0)" by (simp add: eval_polyexp_def) lemma eval_polyexp_Cons: "eval_polyexp (x#xs) = (\n. (case x of (a,k,b) \ a * of_nat n ^ k * b ^ n) + eval_polyexp xs n)" by (simp add: eval_polyexp_def) definition polyexp_fps :: "('a :: field) polyexp \ 'a fps" where "polyexp_fps xs = (\(a,k,b)\xs. fps_of_poly (Polynomial.smult a (fps_monom_poly b k)) / (1 - fps_const b * fps_X) ^ (k + 1))" lemma polyexp_fps_Nil [simp]: "polyexp_fps [] = 0" by (simp add: polyexp_fps_def) lemma polyexp_fps_Cons: "polyexp_fps (x#xs) = (case x of (a,k,b) \ fps_of_poly (Polynomial.smult a (fps_monom_poly b k)) / (1 - fps_const b * fps_X) ^ (k + 1)) + polyexp_fps xs" by (simp add: polyexp_fps_def) -definition polyexp_ratfps :: "('a :: {field,factorial_ring_gcd}) polyexp \ 'a ratfps" where +definition polyexp_ratfps :: "('a :: field_gcd) polyexp \ 'a ratfps" where "polyexp_ratfps xs = (\(a,k,b)\xs. ratfps_of_poly (Polynomial.smult a (fps_monom_poly b k)) / ratfps_of_poly ([:1, -b:] ^ (k + 1)))" lemma polyexp_ratfps_Nil [simp]: "polyexp_ratfps [] = 0" by (simp add: polyexp_ratfps_def) lemma polyexp_ratfps_Cons: "polyexp_ratfps (x#xs) = (case x of (a,k,b) \ ratfps_of_poly (Polynomial.smult a (fps_monom_poly b k)) / ratfps_of_poly ([:1, -b:] ^ (k + 1))) + polyexp_ratfps xs" by (simp add: polyexp_ratfps_def) lemma polyexp_fps: "Abs_fps (eval_polyexp xs) = polyexp_fps xs" proof (induction xs) case (Cons x xs) obtain a k b where [simp]: "x = (a, k, b)" by (metis prod.exhaust) have "Abs_fps (eval_polyexp (x#xs)) = fps_const a * Abs_fps (\n. of_nat n ^ k * b ^ n) + Abs_fps (eval_polyexp xs)" by (simp add: eval_polyexp_Cons fps_plus_def mult_ac) also have "Abs_fps (\n. of_nat n ^ k * b ^ n) = fps_of_poly (fps_monom_poly b k) / (1 - fps_const b * fps_X) ^ (k + 1)" (is "_ = ?A / ?B") by (rule fps_monom) also have "fps_const a * (?A / ?B) = (fps_const a * ?A) / ?B" by (intro unit_div_mult_swap) simp_all also have "fps_const a * ?A = fps_of_poly (Polynomial.smult a (fps_monom_poly b k))" by simp also note Cons.IH finally show ?case by (simp add: polyexp_fps_Cons) qed (simp_all add: fps_zero_def) lemma polyexp_ratfps [simp]: "fps_of_ratfps (polyexp_ratfps xs) = polyexp_fps xs" by (induction xs) (auto simp del: power_Suc fps_const_neg simp: coeff_0_power fps_of_poly_power fps_of_poly_smult fps_of_poly_pCons fps_const_neg [symmetric] mult_ac polyexp_ratfps_Cons polyexp_fps_Cons) definition lir_fps :: - "'a :: {field,factorial_ring_gcd} list \ 'a list \ 'a polyexp \ ('a ratfps) option" where + "'a :: field_gcd list \ 'a list \ 'a polyexp \ ('a ratfps) option" where "lir_fps cs fs g = (if cs = [] \ length fs < length cs - 1 then None else let m = length fs + 1 - length cs; p = lir_fps_numerator m cs (\n. fs ! n) (eval_polyexp g); q = lr_fps_denominator cs in Some ((ratfps_of_poly p + polyexp_ratfps g) * inverse (ratfps_of_poly q)))" lemma lir_fps_correct: - fixes f :: "nat \ 'a :: {field,factorial_ring_gcd}" + fixes f :: "nat \ 'a :: field_gcd" assumes "linear_inhomogenous_recurrence f (eval_polyexp g) cs fs" shows "map_option fps_of_ratfps (lir_fps cs fs g) = Some (Abs_fps f)" proof - interpret linear_inhomogenous_recurrence f "eval_polyexp g" cs fs by fact define m where "m = length fs + 1 - length cs" let ?num = "lir_fps_numerator m cs f (eval_polyexp g)" let ?num' = "lir_fps_numerator m cs ((!) fs) (eval_polyexp g)" let ?denom = "lr_fps_denominator cs" have "{..length cs - 1} = {.. 1" by (cases cs) auto ultimately have "Abs_fps f = (fps_of_poly ?num + Abs_fps (eval_polyexp g)) / fps_of_poly ?denom" by (intro lir_fps) (insert rec, simp_all add: m_def) also have "?num = ?num'" by (rule lir_fps_numerator_altdef [folded m_def]) also have "(fps_of_poly ?num' + Abs_fps (eval_polyexp g)) / fps_of_poly ?denom = fps_of_ratfps ((ratfps_of_poly ?num' + polyexp_ratfps g) * inverse (ratfps_of_poly ?denom))" by (simp add: polyexp_fps fps_divide_unit) also from enough_base have "Some \ = map_option fps_of_ratfps (lir_fps cs fs g)" by (cases cs) (simp_all add: base fps_of_ratfps_def case_prod_unfold lir_fps_def m_def) finally show ?thesis .. qed end diff --git a/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy b/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy --- a/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy +++ b/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy @@ -1,227 +1,227 @@ (* File: Linear_Recurrences_Misc.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Miscellaneous material required for linear recurrences\ theory Linear_Recurrences_Misc imports Complex_Main "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin fun zip_with where "zip_with f (x#xs) (y#ys) = f x y # zip_with f xs ys" | "zip_with f _ _ = []" lemma length_zip_with [simp]: "length (zip_with f xs ys) = min (length xs) (length ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_altdef: "zip_with f xs ys = map (\(x,y). f x y) (zip xs ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_nth [simp]: "n < length xs \ n < length ys \ zip_with f xs ys ! n = f (xs!n) (ys!n)" by (simp add: zip_with_altdef) lemma take_zip_with: "take n (zip_with f xs ys) = zip_with f (take n xs) (take n ys)" proof (induction f xs ys arbitrary: n rule: zip_with.induct) case (1 f x xs y ys n) thus ?case by (cases n) simp_all qed simp_all lemma drop_zip_with: "drop n (zip_with f xs ys) = zip_with f (drop n xs) (drop n ys)" proof (induction f xs ys arbitrary: n rule: zip_with.induct) case (1 f x xs y ys n) thus ?case by (cases n) simp_all qed simp_all lemma map_zip_with: "map f (zip_with g xs ys) = zip_with (\x y. f (g x y)) xs ys" by (induction g xs ys rule: zip_with.induct) simp_all lemma zip_with_map: "zip_with f (map g xs) (map h ys) = zip_with (\x y. f (g x) (h y)) xs ys" by (induction "\x y. f (g x) (h y)" xs ys rule: zip_with.induct) simp_all lemma zip_with_map_left: "zip_with f (map g xs) ys = zip_with (\x y. f (g x) y) xs ys" using zip_with_map[of f g xs "\x. x" ys] by simp lemma zip_with_map_right: "zip_with f xs (map g ys) = zip_with (\x y. f x (g y)) xs ys" using zip_with_map[of f "\x. x" xs g ys] by simp lemma zip_with_swap: "zip_with (\x y. f y x) xs ys = zip_with f ys xs" by (induction f ys xs rule: zip_with.induct) simp_all lemma set_zip_with: "set (zip_with f xs ys) = (\(x,y). f x y) ` set (zip xs ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_Pair: "zip_with Pair (xs :: 'a list) (ys :: 'b list) = zip xs ys" by (induction "Pair :: 'a \ 'b \ _" xs ys rule: zip_with.induct) simp_all lemma zip_with_altdef': "zip_with f xs ys = [f (xs!i) (ys!i). i \ [0.. [0.. 0" shows "card {x. poly p x = 0} \ degree p" using assms proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "\x. poly p x = 0") case False hence "{x. poly p x = 0} = {}" by blast thus ?thesis by simp next case True then obtain x where x: "poly p x = 0" by blast hence "[:-x, 1:] dvd p" by (subst (asm) poly_eq_0_iff_dvd) then obtain q where q: "p = [:-x, 1:] * q" by (auto simp: dvd_def) with \p \ 0\ have [simp]: "q \ 0" by auto have deg: "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) auto have "card {x. poly p x = 0} \ card (insert x {x. poly q x = 0})" by (intro card_mono) (auto intro: poly_roots_finite simp: q) also have "\ \ Suc (card {x. poly q x = 0})" by (rule card_insert_le_m1) auto also from deg have "card {x. poly q x = 0} \ degree q" using \p \ 0\ and q by (intro less) auto also have "Suc \ = degree p" by (simp add: deg) finally show ?thesis by - simp_all qed qed lemma poly_eqI_degree: fixes p q :: "'a :: {comm_ring_1, ring_no_zero_divisors} poly" assumes "\x. x \ A \ poly p x = poly q x" assumes "card A > degree p" "card A > degree q" shows "p = q" proof (rule ccontr) assume neq: "p \ q" have "degree (p - q) \ max (degree p) (degree q)" by (rule degree_diff_le_max) also from assms have "\ < card A" by linarith also have "\ \ card {x. poly (p - q) x = 0}" using neq and assms by (intro card_mono poly_roots_finite) auto finally have "degree (p - q) < card {x. poly (p - q) x = 0}" . moreover have "degree (p - q) \ card {x. poly (p - q) x = 0}" using neq by (intro card_poly_roots_bound) auto ultimately show False by linarith qed lemma poly_root_order_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "P 0" "\p. (\x. poly p x \ 0) \ P p" "\p x n. n > 0 \ poly p x \ 0 \ P p \ P ([:-x, 1:] ^ n * p)" shows "P p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) consider "p = 0" | "p \ 0" "\x. poly p x = 0" | "\x. poly p x \ 0" by blast thus ?case proof cases case 3 with assms(2)[of p] show ?thesis by simp next case 2 then obtain x where x: "poly p x = 0" by auto have "[:-x, 1:] ^ order x p dvd p" by (intro order_1) then obtain q where q: "p = [:-x, 1:] ^ order x p * q" by (auto simp: dvd_def) with 2 have [simp]: "q \ 0" by auto have order_pos: "order x p > 0" using \p \ 0\ and x by (auto simp: order_root) have "order x p = order x p + order x q" by (subst q, subst order_mult) (auto simp: order_power_n_n) hence [simp]: "order x q = 0" by simp have deg: "degree p = order x p + degree q" by (subst q, subst degree_mult_eq) (auto simp: degree_power_eq) with order_pos have "degree q < degree p" by simp hence "P q" by (rule less) with order_pos have "P ([:-x, 1:] ^ order x p * q)" by (intro assms(3)) (auto simp: order_root) with q show ?thesis by simp qed (simp_all add: assms(1)) qed lemma complex_poly_decompose: "smult (lead_coeff p) (\z|poly p z = 0. [:-z, 1:] ^ order z p) = (p :: complex poly)" proof (induction p rule: poly_root_order_induct) case (no_roots p) show ?case proof (cases "degree p = 0") case False hence "\constant (poly p)" by (subst constant_degree) with fundamental_theorem_of_algebra and no_roots show ?thesis by blast qed (auto elim!: degree_eq_zeroE) next case (root p x n) from root have *: "{z. poly ([:- x, 1:] ^ n * p) z = 0} = insert x {z. poly p z = 0}" by auto have "smult (lead_coeff ([:-x, 1:] ^ n * p)) (\z|poly ([:-x,1:] ^ n * p) z = 0. [:-z, 1:] ^ order z ([:- x, 1:] ^ n * p)) = [:- x, 1:] ^ order x ([:- x, 1:] ^ n * p) * smult (lead_coeff p) (\z\{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p))" by (subst *, subst prod.insert) (insert root, auto intro: poly_roots_finite simp: mult_ac lead_coeff_mult lead_coeff_power) also have "order x ([:- x, 1:] ^ n * p) = n" using root by (subst order_mult) (auto simp: order_power_n_n order_0I) also have "(\z\{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p)) = (\z\{z. poly p z = 0}. [:- z, 1:] ^ order z p)" proof (intro prod.cong refl, goal_cases) case (1 y) with root have "order y ([:-x,1:] ^ n) = 0" by (intro order_0I) auto thus ?case using root by (subst order_mult) auto qed also note root.IH finally show ?case . qed simp_all lemma normalize_field: "normalize (x :: 'a :: {normalization_semidom,field}) = (if x = 0 then 0 else 1)" by (auto simp: normalize_1_iff dvd_field_iff) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {normalization_semidom,field}) = x" using unit_factor_mult_normalize[of x] normalize_field[of x] by (simp split: if_splits) lemma coprime_linear_poly: - fixes c :: "'a :: {field,factorial_ring_gcd}" + fixes c :: "'a :: field_gcd" assumes "c \ c'" shows "coprime [:c,1:] [:c',1:]" proof - have "gcd [:c,1:] [:c',1:] = gcd ([:c,1:] - [:c',1:]) [:c',1:]" by (rule gcd_diff1 [symmetric]) also have "[:c,1:] - [:c',1:] = [:c-c':]" by simp also from assms have "gcd \ [:c',1:] = normalize [:c-c':]" by (intro gcd_proj1_if_dvd) (auto simp: const_poly_dvd_iff dvd_field_iff) also from assms have "\ = 1" by (simp add: normalize_poly_def) finally show "coprime [:c,1:] [:c',1:]" by (simp add: gcd_eq_1_imp_coprime) qed lemma coprime_linear_poly': - fixes c :: "'a :: {field,factorial_ring_gcd,normalization_euclidean_semiring}" + fixes c :: "'a :: field_gcd" assumes "c \ c'" "c \ 0" "c' \ 0" shows "coprime [:1,c:] [:1,c':]" proof - have "gcd [:1,c:] [:1,c':] = gcd ([:1,c:] mod [:1,c':]) [:1,c':]" by simp also have "eucl_rel_poly [:1, c:] [:1, c':] ([:c/c':], [:1-c/c':])" using assms by (auto simp add: eucl_rel_poly_iff one_pCons) hence "[:1,c:] mod [:1,c':] = [:1 - c / c':]" by (rule mod_poly_eq) also from assms have "gcd \ [:1,c':] = normalize ([:1 - c / c':])" by (intro gcd_proj1_if_dvd) (auto simp: const_poly_dvd_iff dvd_field_iff) also from assms have "\ = 1" by (auto simp: normalize_poly_def) finally show ?thesis by (rule gcd_eq_1_imp_coprime) qed end diff --git a/thys/Linear_Recurrences/Partial_Fraction_Decomposition.thy b/thys/Linear_Recurrences/Partial_Fraction_Decomposition.thy --- a/thys/Linear_Recurrences/Partial_Fraction_Decomposition.thy +++ b/thys/Linear_Recurrences/Partial_Fraction_Decomposition.thy @@ -1,384 +1,384 @@ (* File: Partial_Fraction_Decomposition.thy Author: Manuel Eberl Partial fraction decomposition on Euclidean rings, i.e. decomposing a quotient into a sum of quotients where each denominator is a power of an irreducible element. (and possibly one summand that is an entire element) The most interesting setting is when the Euclidean ring is a polynomial ring. *) section \Partial Fraction Decomposition\ theory Partial_Fraction_Decomposition imports Main "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" "HOL-Library.Sublist" Linear_Recurrences_Misc begin subsection \Decomposition on general Euclidean rings\ text \ Consider elements $x, y_1, \ldots, y_n$ of a ring $R$, where the $y_i$ are pairwise coprime. A \emph{Partial Fraction Decomposition} of these elements (or rather the formal quotient $x / (y_1 \ldots y_n)$ that they represent) is a finite sum of summands of the form $a / y_i ^ k$. Obviously, the sum can be arranged such that there is at most one summand with denominator $y_i ^ n$ for any combination of $i$ and $n$; in particular, there is at most one summand with denominator 1. We can decompose the summands further by performing division with remainder until in all quotients, the numerator's Euclidean size is less than that of the denominator. \ text \ The following function performs the first step of the above process: it takes the values $x$ and $y_1,\ldots, y_n$ and returns the numerators of the summands in the decomposition. (the denominators are simply the $y_i$ from the input) \ fun decompose :: "('a :: euclidean_ring_gcd) \ 'a list \ 'a list" where "decompose x [] = []" | "decompose x [y] = [x]" | "decompose x (y#ys) = (case bezout_coefficients y (prod_list ys) of (a, b) \ (b*x) # decompose (a*x) ys)" lemma decompose_rec: "ys \ [] \ decompose x (y#ys) = (case bezout_coefficients y (prod_list ys) of (a, b) \ (b*x) # decompose (a*x) ys)" by (cases ys) simp_all lemma length_decompose [simp]: "length (decompose x ys) = length ys" proof (induction x ys rule: decompose.induct) case (3 x y z ys) obtain a b where ab: "(a,b) = bezout_coefficients y (prod_list (z#ys))" by (cases "bezout_coefficients y (z * prod_list ys)") simp_all from 3[OF ab] ab[symmetric] show ?case by simp qed simp_all fun decompose' :: "('a :: euclidean_ring_gcd) \ 'a list \ 'a list \ 'a list" where "decompose' x [] _ = []" | "decompose' x [y] _ = [x]" | "decompose' _ _ [] = []" | "decompose' x (y#ys) (p#ps) = (case bezout_coefficients y p of (a, b) \ (b*x) # decompose' (a*x) ys ps)" primrec decompose_aux :: "'a :: {ab_semigroup_mult,monoid_mult} \ _" where "decompose_aux acc [] = [acc]" | "decompose_aux acc (x#xs) = acc # decompose_aux (x * acc) xs" lemma decompose_code [code]: "decompose x ys = decompose' x ys (tl (rev (decompose_aux 1 (rev ys))))" proof (induction x ys rule: decompose.induct) case (3 x y1 y2 ys) have [simp]: "decompose_aux acc xs = map (\x. prod_list x * acc) (prefixes xs)" for acc :: 'a and xs by (induction xs arbitrary: acc) (simp_all add: mult_ac) show ?case using 3[of "fst (bezout_coefficients y1 (y2 * prod_list ys))" "snd (bezout_coefficients y1 (y2 * prod_list ys))"] by (simp add: case_prod_unfold rev_map prefixes_conv_suffixes o_def mult_ac) qed simp_all text \ The next function performs the second step: Given a quotient of the form $x / y^n$, it returns a list of $x_0, \ldots, x_n$ such that $x / y^n = x_0 / y^n + \ldots + x_{n-1} / y + x_n$ and all $x_i$ have a Euclidean size less than that of $y$. \ fun normalise_decomp :: "('a :: semiring_modulo) \ 'a \ nat \ 'a \ ('a list)" where "normalise_decomp x y 0 = (x, [])" | "normalise_decomp x y (Suc n) = ( case normalise_decomp (x div y) y n of (z, rs) \ (z, x mod y # rs))" lemma length_normalise_decomp [simp]: "length (snd (normalise_decomp x y n)) = n" by (induction x y n rule: normalise_decomp.induct) (auto split: prod.split) text \ The following constant implements the full process of partial fraction decomposition: The input is a quotient $x / (y_1 ^ {k_1} \ldots y_n ^ {k_n})$ and the output is a sum of an entire element and terms of the form $a / y_i ^ k$ where $a$ has a Euclidean size less than $y_i$. \ definition partial_fraction_decomposition :: "'a :: euclidean_ring_gcd \ ('a \ nat) list \ 'a \ 'a list list" where "partial_fraction_decomposition x ys = (if ys = [] then (x, []) else (let zs = [let (y, n) = ys ! i in normalise_decomp (decompose x (map (\(y,n). y ^ Suc n) ys) ! i) y (Suc n). i \ [0.. length (snd (partial_fraction_decomposition x ys) ! i) = snd (ys ! i) + 1" by (auto simp: partial_fraction_decomposition_def case_prod_unfold Let_def) lemma size_normalise_decomp: "a \ set (snd (normalise_decomp x y n)) \ y \ 0 \ euclidean_size a < euclidean_size y" by (induction x y n rule: normalise_decomp.induct) (auto simp: case_prod_unfold Let_def mod_size_less) lemma size_partial_fraction_decomposition: "i < length xs \ fst (xs ! i) \ 0 \ x \ set (snd (partial_fraction_decomposition y xs) ! i) \ euclidean_size x < euclidean_size (fst (xs ! i))" by (auto simp: partial_fraction_decomposition_def Let_def case_prod_unfold simp del: normalise_decomp.simps split: if_split_asm intro!: size_normalise_decomp) text \ A homomorphism $\varphi$ from a Euclidean ring $R$ into another ring $S$ with a notion of division. We will show that for any $x,y\in R$ such that $\phi(y)$ is a unit, we can perform partial fraction decomposition on the quotient $\varphi(x) / \varphi(y)$. The obvious choice for $S$ is the fraction field of $R$, but other choices may also make sense: If, for example, $R$ is a ring of polynomials $K[X]$, then one could let $S = K$ and $\varphi$ the evaluation homomorphism. Or one could let $S = K[[X]]$ (the ring of formal power series) and $\varphi$ the canonical homomorphism from polynomials to formal power series. \ locale pfd_homomorphism = fixes lift :: "('a :: euclidean_ring_gcd) \ ('b :: euclidean_semiring_cancel)" assumes lift_add: "lift (a + b) = lift a + lift b" assumes lift_mult: "lift (a * b) = lift a * lift b" assumes lift_0 [simp]: "lift 0 = 0" assumes lift_1 [simp]: "lift 1 = 1" begin lemma lift_power: "lift (a ^ n) = lift a ^ n" by (induction n) (simp_all add: lift_mult) definition from_decomp :: "'a \ 'a \ nat \ 'b" where "from_decomp x y n = lift x div lift y ^ n" lemma decompose: assumes "ys \ []" "pairwise coprime (set ys)" "distinct ys" "\y. y \ set ys \ is_unit (lift y)" shows "(\iy\set ys. is_unit (lift y)" by simp hence unit': "is_unit (lift (prod_list ys))" by (induction ys) (auto simp: lift_mult) ultimately have unit: "lift y dvd b" "lift (prod_list ys) dvd b" for b by auto obtain s t where st: "bezout_coefficients y (prod_list ys) = (s, t)" by (cases "bezout_coefficients y (prod_list ys)") simp_all from \pairwise coprime (set (y#ys))\ have coprime:"pairwise coprime (set ys)" by (rule pairwise_subset) auto have "(\i = lift (prod_list ys * t * x + y * s * x)" using assms unit by (simp add: lift_mult lift_add algebra_simps) finally have "(\i = 1" using \coprime (prod_list ys) y\ by simp finally show ?case by simp qed simp_all lemma normalise_decomp: fixes x y :: 'a and n :: nat assumes "is_unit (lift y)" defines "xs \ snd (normalise_decomp x y n)" shows "lift (fst (normalise_decomp x y n)) + (\iiii + from_decomp (x mod y) y (Suc n)) * lift y = (lift ((x div y) * y + x mod y) div lift y ^ n)" (is "?A * _ = ?B div _") unfolding lift_add lift_mult apply (subst div_add) apply (auto simp add: from_decomp_def algebra_simps dvd_div_mult2_eq unit_div_mult_swap dvd_div_mult2_eq[OF unit_imp_dvd] is_unit_mult_iff) done with 2(2) have "?A = \ div lift y" by (subst eq_commute, subst dvd_div_eq_mult) auto also from 2(2) unit have "\ = ?B div (lift y ^ Suc n)" by (subst is_unit_div_mult2_eq [symmetric]) (auto simp: mult_ac) also have "x div y * y + x mod y = x" by (rule div_mult_mod_eq) finally show ?case . qed simp_all lemma lift_prod_list: "lift (prod_list xs) = prod_list (map lift xs)" by (induction xs) (simp_all add: lift_mult) lemma lift_sum: "lift (sum f A) = sum (\x. lift (f x)) A" by (cases "finite A", induction A rule: finite_induct) (simp_all add: lift_add) lemma partial_fraction_decomposition: fixes ys :: "('a \ nat) list" defines "ys' \ map (\(x,n). x ^ Suc n) ys :: 'a list" assumes unit: "\y. y \ fst ` set ys \ is_unit (lift y)" assumes coprime: "pairwise coprime (set ys')" assumes distinct: "distinct ys'" assumes "partial_fraction_decomposition x ys = (a, zs)" shows "lift a + (\ij\snd (ys!i). from_decomp (zs!i!j) (fst (ys!i)) (snd (ys!i)+1 - j)) = lift x div lift (prod_list ys')" proof (cases "ys = []") assume [simp]: "ys \ []" define n where "n = length ys" have "lift x div lift (prod_list ys') = (\i = (\iij\snd (ys!i). from_decomp (zs!i!j) (fst (ys!i)) (snd (ys!i)+1 - j)))" (is "_ = ?A + ?B") proof (subst sum.distrib [symmetric], intro sum.cong refl, goal_cases) case (1 i) from 1 have "lift (ys' ! i) = lift (fst (ys ! i)) ^ Suc (snd (ys ! i))" by (simp add: ys'_def n_def lift_power lift_mult split: prod.split) also from 1 have "lift (decompose x ys' ! i) div \ = lift (fst (normalise_decomp (decompose x ys' ! i) (fst (ys!i)) (snd (ys!i)+1))) + (\jj\snd (ys!i). from_decomp (zs!i!j) (fst (ys!i)) (snd (ys!i)+1 - j))" using assms 1 by (intro sum.cong refl) (auto simp: partial_fraction_decomposition_def case_prod_unfold Let_def o_def n_def simp del: normalise_decomp.simps) finally show ?case . qed also from assms have "?A = lift a" by (auto simp: partial_fraction_decomposition_def o_def sum_list_sum_nth atLeast0LessThan case_prod_unfold Let_def lift_sum n_def intro!: sum.cong) finally show ?thesis by (simp add: n_def) qed (insert assms, simp add: partial_fraction_decomposition_def) end subsection \Specific results for polynomials\ definition divmod_field_poly :: "'a :: field poly \ 'a poly \ 'a poly \ 'a poly" where "divmod_field_poly p q = (p div q, p mod q)" lemma divmod_field_poly_code [code]: "divmod_field_poly p q = (let cg = coeffs q in if cg = [] then (0, p) else let cf = coeffs p; ilc = inverse (last cg); ch = map ((*) ilc) cg; (q, r) = divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) ilc) q), poly_of_list (rev r)))" unfolding divmod_field_poly_def by (rule pdivmod_via_divmod_list) -definition normalise_decomp_poly :: "'a::field poly \ 'a poly \ nat \ 'a poly \ 'a poly list" +definition normalise_decomp_poly :: "'a::field_gcd poly \ 'a poly \ nat \ 'a poly \ 'a poly list" where [simp]: "normalise_decomp_poly (p :: _ poly) q n = normalise_decomp p q n" lemma normalise_decomp_poly_code [code]: "normalise_decomp_poly x y 0 = (x, [])" "normalise_decomp_poly x y (Suc n) = ( let (x', r) = divmod_field_poly x y; (z, rs) = normalise_decomp_poly x' y n in (z, r # rs))" by (simp_all add: divmod_field_poly_def) definition poly_pfd_simple where "poly_pfd_simple x cs = (if cs = [] then (x, []) else (let zs = [let (c, n) = cs ! i in normalise_decomp_poly (decompose x (map (\(c,n). [:1,-c:] ^ Suc n) cs) ! i) [:1,-c:] (n+1). i \ [0..p. coeff p 0) \ snd) zs)))" lemma poly_pfd_simple_code [code]: "poly_pfd_simple x cs = (if cs = [] then (x, []) else let zs = zip_with (\(c,n) decomp. normalise_decomp_poly decomp [:1,-c:] (n+1)) cs (decompose x (map (\(c,n). [:1,-c:] ^ Suc n) cs)) in (sum_list (map fst zs), map (map (\p. coeff p 0) \ snd) zs))" unfolding poly_pfd_simple_def zip_with_altdef' by (simp add: Let_def case_prod_unfold) lemma fst_poly_pfd_simple: "fst (poly_pfd_simple x cs) = fst (partial_fraction_decomposition x (map (\(c,n). ([:1,-c:],n)) cs))" by (auto simp: poly_pfd_simple_def partial_fraction_decomposition_def o_def case_prod_unfold Let_def sum_list_sum_nth intro!: sum.cong) lemma const_polyI: "degree p = 0 \ [:coeff p 0:] = p" by (elim degree_eq_zeroE) simp_all lemma snd_poly_pfd_simple: - "map (map (\c. [:c :: 'a :: {field,factorial_ring_gcd,normalization_euclidean_semiring}:])) (snd (poly_pfd_simple x cs)) = + "map (map (\c. [:c :: 'a :: field_gcd:])) (snd (poly_pfd_simple x cs)) = (snd (partial_fraction_decomposition x (map (\(c,n). ([:1,-c:],n)) cs)))" proof - have "snd (poly_pfd_simple x cs) = map (map (\p. coeff p 0)) (snd (partial_fraction_decomposition x (map (\(c,n). ([:1,-c:],n)) cs)))" (is "_ = map ?f ?B") by (auto simp: poly_pfd_simple_def partial_fraction_decomposition_def o_def case_prod_unfold Let_def sum_list_sum_nth intro!: sum.cong) also have "map (map (\c. [:c:])) (map ?f ?B) = map (map (\x. x)) ?B" unfolding map_map o_def proof (intro map_cong refl const_polyI, goal_cases) case (1 ys y) from 1 obtain i where i: "i < length cs" "ys = snd (partial_fraction_decomposition x (map (\(c,n). ([:1,-c:],n)) cs)) ! i" by (auto simp: in_set_conv_nth) with 1 have "euclidean_size y < euclidean_size (fst (map (\(c,n). ([:1,-c:],n)) cs ! i))" by (intro size_partial_fraction_decomposition[of i _ _ x]) (auto simp: case_prod_unfold Let_def) with i(1) have "euclidean_size y < 2" by (auto simp: case_prod_unfold Let_def euclidean_size_poly_def split: if_split_asm) thus ?case by (cases y rule: pCons_cases) (auto simp: euclidean_size_poly_def split: if_split_asm) qed finally show ?thesis by simp qed lemma poly_pfd_simple: "partial_fraction_decomposition x (map (\(c,n). ([:1,-c:],n)) cs) = (fst (poly_pfd_simple x cs), map (map (\c. [:c:])) (snd (poly_pfd_simple x cs)))" by (simp add: fst_poly_pfd_simple snd_poly_pfd_simple) end diff --git a/thys/Linear_Recurrences/RatFPS.thy b/thys/Linear_Recurrences/RatFPS.thy --- a/thys/Linear_Recurrences/RatFPS.thy +++ b/thys/Linear_Recurrences/RatFPS.thy @@ -1,936 +1,945 @@ (* File: RatFPS.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Rational formal power series\ theory RatFPS imports Complex_Main "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin subsection \Some auxiliary\ abbreviation constant_term :: "'a poly \ 'a::zero" where "constant_term p \ coeff p 0" lemma coeff_0_mult: "coeff (p * q) 0 = coeff p 0 * coeff q 0" by (simp add: coeff_mult) lemma coeff_0_div: assumes "coeff p 0 \ 0" assumes "(q :: 'a :: field poly) dvd p" shows "coeff (p div q) 0 = coeff p 0 div coeff q 0" proof (cases "q = 0") case False from assms have "p = p div q * q" by simp also have "coeff \ 0 = coeff (p div q) 0 * coeff q 0" by (simp add: coeff_0_mult) finally show ?thesis using assms by auto qed simp_all lemma coeff_0_add_fract_nonzero: assumes "coeff (snd (quot_of_fract x)) 0 \ 0" "coeff (snd (quot_of_fract y)) 0 \ 0" shows "coeff (snd (quot_of_fract (x + y))) 0 \ 0" proof - define num where "num = fst (quot_of_fract x) * snd (quot_of_fract y) + snd (quot_of_fract x) * fst (quot_of_fract y)" define denom where "denom = snd (quot_of_fract x) * snd (quot_of_fract y)" define z where "z = (num, denom)" from assms have "snd z \ 0" by (auto simp: denom_def z_def) from normalize_quotE'[OF this] guess d . note d = this from assms have z: "coeff (snd z) 0 \ 0" by (simp add: z_def denom_def coeff_0_mult) have "coeff (snd (quot_of_fract (x + y))) 0 = coeff (snd (normalize_quot z)) 0" by (simp add: quot_of_fract_add Let_def case_prod_unfold z_def num_def denom_def) also from z have "\ \ 0" using d by (simp add: d coeff_0_mult) finally show ?thesis . qed lemma coeff_0_normalize_quot_nonzero [simp]: assumes "coeff (snd x) 0 \ 0" shows "coeff (snd (normalize_quot x)) 0 \ 0" proof - from assms have "snd x \ 0" by auto from normalize_quotE'[OF this] guess d . with assms show ?thesis by (auto simp: coeff_0_mult) qed -abbreviation numerator :: "'a fract \ 'a::{ring_gcd,idom_divide}" +abbreviation numerator :: "'a fract \ 'a::{ring_gcd,idom_divide,semiring_gcd_mult_normalize}" where "numerator x \ fst (quot_of_fract x)" -abbreviation denominator :: "'a fract \ 'a::{ring_gcd,idom_divide}" +abbreviation denominator :: "'a fract \ 'a::{ring_gcd,idom_divide,semiring_gcd_mult_normalize}" where "denominator x \ snd (quot_of_fract x)" declare unit_factor_snd_quot_of_fract [simp] normalize_snd_quot_of_fract [simp] lemma constant_term_denominator_nonzero_imp_constant_term_denominator_div_gcd_nonzero: "constant_term (denominator x div gcd a (denominator x)) \ 0" if "constant_term (denominator x) \ 0" using that coeff_0_normalize_quot_nonzero [of "(a, denominator x)"] normalize_quot_proj(2) [of "denominator x" a] by simp subsection \The type of rational formal power series\ -typedef (overloaded) 'a :: "{field,factorial_ring_gcd}" ratfps = +typedef (overloaded) 'a :: field_gcd ratfps = "{x :: 'a poly fract. constant_term (denominator x) \ 0}" by (rule exI [of _ 0]) simp setup_lifting type_definition_ratfps -instantiation ratfps :: ("{field,factorial_ring_gcd}") idom +instantiation ratfps :: (field_gcd) idom begin lift_definition zero_ratfps :: "'a ratfps" is "0" by simp lift_definition one_ratfps :: "'a ratfps" is "1" by simp lift_definition uminus_ratfps :: "'a ratfps \ 'a ratfps" is "uminus" by (simp add: quot_of_fract_uminus case_prod_unfold Let_def) lift_definition plus_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(+)" by (rule coeff_0_add_fract_nonzero) lift_definition minus_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(-)" by (simp only: diff_conv_add_uminus, rule coeff_0_add_fract_nonzero) (simp_all add: quot_of_fract_uminus Let_def case_prod_unfold) lift_definition times_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(*)" by (simp add: quot_of_fract_mult Let_def case_prod_unfold coeff_0_mult constant_term_denominator_nonzero_imp_constant_term_denominator_div_gcd_nonzero) instance by (standard; transfer) (simp_all add: ring_distribs) end fun ratfps_nth_aux :: "('a::field) poly \ nat \ 'a" where "ratfps_nth_aux p 0 = inverse (coeff p 0)" | "ratfps_nth_aux p n = - inverse (coeff p 0) * sum (\i. coeff p i * ratfps_nth_aux p (n - i)) {1..n}" lemma ratfps_nth_aux_correct: "ratfps_nth_aux p n = natfun_inverse (fps_of_poly p) n" by (induction p n rule: ratfps_nth_aux.induct) simp_all -lift_definition ratfps_nth :: "'a :: {field,factorial_ring_gcd} ratfps \ nat \ 'a" is +lift_definition ratfps_nth :: "'a :: field_gcd ratfps \ nat \ 'a" is "\x n. let (a,b) = quot_of_fract x in (\i = 0..n. coeff a i * ratfps_nth_aux b (n - i))" . -lift_definition ratfps_subdegree :: "'a :: {field,factorial_ring_gcd} ratfps \ nat" is +lift_definition ratfps_subdegree :: "'a :: field_gcd ratfps \ nat" is "\x. poly_subdegree (fst (quot_of_fract x))" . context includes lifting_syntax begin lemma RatFPS_parametric: "(rel_prod (=) (=) ===> (=)) (\(p,q). if coeff q 0 = 0 then 0 else quot_to_fract (p, q)) (\(p,q). if coeff q 0 = 0 then 0 else quot_to_fract (p, q))" by transfer_prover end lemma normalize_quot_quot_of_fract [simp]: "normalize_quot (quot_of_fract x) = quot_of_fract x" by (rule normalize_quot_id, rule quot_of_fract_in_normalized_fracts) context -assumes "SORT_CONSTRAINT('a::{field,factorial_ring_gcd})" +assumes "SORT_CONSTRAINT('a::field_gcd)" begin lift_definition quot_of_ratfps :: "'a ratfps \ ('a poly \ 'a poly)" is "quot_of_fract :: 'a poly fract \ ('a poly \ 'a poly)" . lift_definition quot_to_ratfps :: "('a poly \ 'a poly) \ 'a ratfps" is "\(x,y). let (x',y') = normalize_quot (x,y) in if coeff y' 0 = 0 then 0 else quot_to_fract (x',y')" by (simp add: case_prod_unfold Let_def quot_of_fract_quot_to_fract) lemma quot_to_ratfps_quot_of_ratfps [code abstype]: "quot_to_ratfps (quot_of_ratfps x) = x" by transfer (simp add: case_prod_unfold Let_def) lemma coeff_0_snd_quot_of_ratfps_nonzero [simp]: "coeff (snd (quot_of_ratfps x)) 0 \ 0" by transfer simp lemma quot_of_ratfps_quot_to_ratfps: "coeff (snd x) 0 \ 0 \ x \ normalized_fracts \ quot_of_ratfps (quot_to_ratfps x) = x" by transfer (simp add: Let_def case_prod_unfold coeff_0_normalize_quot_nonzero quot_of_fract_quot_to_fract normalize_quot_id) lemma quot_of_ratfps_0 [simp, code abstract]: "quot_of_ratfps 0 = (0, 1)" by transfer simp_all lemma quot_of_ratfps_1 [simp, code abstract]: "quot_of_ratfps 1 = (1, 1)" by transfer simp_all lift_definition ratfps_of_poly :: "'a poly \ 'a ratfps" is "to_fract :: 'a poly \ _" by transfer simp lemma ratfps_of_poly_code [code abstract]: "quot_of_ratfps (ratfps_of_poly p) = (p, 1)" by transfer' simp lemmas zero_ratfps_code = quot_of_ratfps_0 lemmas one_ratfps_code = quot_of_ratfps_1 lemma uminus_ratfps_code [code abstract]: "quot_of_ratfps (- x) = (let (a, b) = quot_of_ratfps x in (-a, b))" by transfer (rule quot_of_fract_uminus) lemma plus_ratfps_code [code abstract]: "quot_of_ratfps (x + y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y in normalize_quot (a * d + b * c, b * d))" by transfer' (rule quot_of_fract_add) lemma minus_ratfps_code [code abstract]: "quot_of_ratfps (x - y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y in normalize_quot (a * d - b * c, b * d))" by transfer' (rule quot_of_fract_diff) -definition ratfps_cutoff :: "nat \ 'a :: {field,factorial_ring_gcd} ratfps \ 'a poly" where +definition ratfps_cutoff :: "nat \ 'a :: field_gcd ratfps \ 'a poly" where "ratfps_cutoff n x = poly_of_list (map (ratfps_nth x) [0.. 'a :: {field,factorial_ring_gcd} ratfps \ 'a ratfps" where +definition ratfps_shift :: "nat \ 'a :: field_gcd ratfps \ 'a ratfps" where "ratfps_shift n x = (let (a, b) = quot_of_ratfps (x - ratfps_of_poly (ratfps_cutoff n x)) in quot_to_ratfps (poly_shift n a, b))" lemma times_ratfps_code [code abstract]: "quot_of_ratfps (x * y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y; (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b) in (e*g, f*h))" by transfer' (rule quot_of_fract_mult) lemma ratfps_nth_code [code]: "ratfps_nth x n = (let (a,b) = quot_of_ratfps x in \i = 0..n. coeff a i * ratfps_nth_aux b (n - i))" by transfer' simp lemma ratfps_subdegree_code [code]: "ratfps_subdegree x = poly_subdegree (fst (quot_of_ratfps x))" by transfer simp end -instantiation ratfps :: ("{field,factorial_ring_gcd}") inverse +instantiation ratfps :: ("field_gcd") inverse begin lift_definition inverse_ratfps :: "'a ratfps \ 'a ratfps" is "\x. let (a,b) = quot_of_fract x in if coeff a 0 = 0 then 0 else inverse x" by (auto simp: case_prod_unfold Let_def quot_of_fract_inverse) lift_definition divide_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "\f g. (if g = 0 then 0 else let n = ratfps_subdegree g; h = ratfps_shift n g in ratfps_shift n (f * inverse h))" . instance .. end lemma ratfps_inverse_code [code abstract]: "quot_of_ratfps (inverse x) = (let (a,b) = quot_of_ratfps x in if coeff a 0 = 0 then (0, 1) else let u = unit_factor a in (b div u, a div u))" by transfer' (simp_all add: Let_def case_prod_unfold quot_of_fract_inverse) instantiation ratfps :: (equal) equal begin definition equal_ratfps :: "'a ratfps \ 'a ratfps \ bool" where [simp]: "equal_ratfps x y \ x = y" instance by standard simp end lemma quot_of_fract_eq_iff [simp]: "quot_of_fract x = quot_of_fract y \ x = y" by transfer (auto simp: normalize_quot_eq_iff) lemma equal_ratfps_code [code]: "HOL.equal x y \ quot_of_ratfps x = quot_of_ratfps y" unfolding equal_ratfps_def by transfer simp lemma fps_of_poly_quot_normalize_quot [simp]: "fps_of_poly (fst (normalize_quot x)) / fps_of_poly (snd (normalize_quot x)) = fps_of_poly (fst x) / fps_of_poly (snd x)" - if "(snd x :: 'a :: {field, factorial_ring_gcd} poly) \ 0" + if "(snd x :: 'a :: field_gcd poly) \ 0" proof - from that obtain d where "fst x = fst (normalize_quot x) * d" and "snd x = snd (normalize_quot x) * d" and "d \ 0" by (rule normalize_quotE') then show ?thesis by (simp add: fps_of_poly_mult) qed lemma fps_of_poly_quot_normalize_quot' [simp]: "fps_of_poly (fst (normalize_quot x)) / fps_of_poly (snd (normalize_quot x)) = fps_of_poly (fst x) / fps_of_poly (snd x)" - if "coeff (snd x) 0 \ (0 :: 'a :: {field,factorial_ring_gcd})" + if "coeff (snd x) 0 \ (0 :: 'a :: field_gcd)" using that by (auto intro: fps_of_poly_quot_normalize_quot) -lift_definition fps_of_ratfps :: "'a :: {field,factorial_ring_gcd} ratfps \ 'a fps" is +lift_definition fps_of_ratfps :: "'a :: field_gcd ratfps \ 'a fps" is "\x. fps_of_poly (numerator x) / fps_of_poly (denominator x)" . lemma fps_of_ratfps_altdef: "fps_of_ratfps x = (case quot_of_ratfps x of (a, b) \ fps_of_poly a / fps_of_poly b)" by transfer (simp add: case_prod_unfold) lemma fps_of_ratfps_ratfps_of_poly [simp]: "fps_of_ratfps (ratfps_of_poly p) = fps_of_poly p" by transfer simp lemma fps_of_ratfps_0 [simp]: "fps_of_ratfps 0 = 0" by transfer simp lemma fps_of_ratfps_1 [simp]: "fps_of_ratfps 1 = 1" by transfer simp lemma fps_of_ratfps_uminus [simp]: "fps_of_ratfps (-x) = - fps_of_ratfps x" by transfer (simp add: quot_of_fract_uminus case_prod_unfold Let_def fps_of_poly_simps dvd_neg_div) lemma fps_of_ratfps_add [simp]: "fps_of_ratfps (x + y) = fps_of_ratfps x + fps_of_ratfps y" by transfer (simp add: quot_of_fract_add Let_def case_prod_unfold fps_of_poly_simps) lemma fps_of_ratfps_diff [simp]: "fps_of_ratfps (x - y) = fps_of_ratfps x - fps_of_ratfps y" by transfer (simp add: quot_of_fract_diff Let_def case_prod_unfold fps_of_poly_simps) lemma is_unit_div_div_commute: "is_unit b \ is_unit c \ a div b div c = a div c div b" by (metis is_unit_div_mult2_eq mult.commute) lemma fps_of_ratfps_mult [simp]: "fps_of_ratfps (x * y) = fps_of_ratfps x * fps_of_ratfps y" proof (transfer, goal_cases) case (1 x y) moreover define x' y' where "x' = quot_of_fract x" and "y' = quot_of_fract y" ultimately have assms: "coeff (snd x') 0 \ 0" "coeff (snd y') 0 \ 0" by simp_all moreover define w z where "w = normalize_quot (fst x', snd y')" and "z = normalize_quot (fst y', snd x')" ultimately have unit: "coeff (snd x') 0 \ 0" "coeff (snd y') 0 \ 0" "coeff (snd w) 0 \ 0" "coeff (snd z) 0 \ 0" by (simp_all add: coeff_0_normalize_quot_nonzero) have "fps_of_poly (fst w * fst z) / fps_of_poly (snd w * snd z) = (fps_of_poly (fst w) / fps_of_poly (snd w)) * (fps_of_poly (fst z) / fps_of_poly (snd z))" (is "_ = ?A * ?B") by (simp add: is_unit_div_mult2_eq fps_of_poly_mult unit_div_mult_swap unit_div_commute unit) also have "\ = (fps_of_poly (fst x') / fps_of_poly (snd x')) * (fps_of_poly (fst y') / fps_of_poly (snd y'))" using unit by (simp add: w_def z_def unit_div_commute unit_div_mult_swap is_unit_div_div_commute) finally show ?case by (simp add: w_def z_def x'_def y'_def Let_def case_prod_unfold quot_of_fract_mult mult_ac) qed lemma div_const_unit_poly: "is_unit c \ p div [:c:] = smult (1 div c) p" by (simp add: is_unit_const_poly_iff unit_eq_div1) lemma normalize_field: "normalize (x :: 'a :: {normalization_semidom,field}) = (if x = 0 then 0 else 1)" by (auto simp: normalize_1_iff dvd_field_iff) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {normalization_semidom,field}) = x" using unit_factor_mult_normalize[of x] normalize_field[of x] by (simp split: if_splits) lemma fps_of_poly_normalize_field: "fps_of_poly (normalize (p :: 'a :: {field, normalization_semidom} poly)) = fps_of_poly p * fps_const (inverse (lead_coeff p))" by (cases "p = 0") (simp_all add: normalize_poly_def div_const_unit_poly divide_simps dvd_field_iff) lemma unit_factor_poly_altdef: "unit_factor p = monom (unit_factor (lead_coeff p)) 0" by (simp add: unit_factor_poly_def monom_altdef) lemma div_const_poly: "p div [:c::'a::field:] = smult (inverse c) p" by (cases "c = 0") (simp_all add: unit_eq_div1 is_unit_triv) lemma fps_of_ratfps_inverse [simp]: "fps_of_ratfps (inverse x) = inverse (fps_of_ratfps x)" proof (transfer, goal_cases) case (1 x) hence "smult (lead_coeff (fst (quot_of_fract x))) (snd (quot_of_fract x)) div unit_factor (fst (quot_of_fract x)) = snd (quot_of_fract x)" if "fst (quot_of_fract x) \ 0" using that by (simp add: unit_factor_poly_altdef monom_0 div_const_poly) with 1 show ?case by (auto simp: Let_def case_prod_unfold fps_divide_unit fps_inverse_mult quot_of_fract_inverse mult_ac fps_of_poly_simps fps_const_inverse fps_of_poly_normalize_field div_smult_left [symmetric]) qed context includes fps_notation begin lemma ratfps_nth_altdef: "ratfps_nth x n = fps_of_ratfps x $ n" by transfer (simp_all add: case_prod_unfold fps_divide_unit fps_times_def fps_inverse_def ratfps_nth_aux_correct Let_def) lemma fps_of_ratfps_is_unit: "fps_of_ratfps a $ 0 \ 0 \ ratfps_nth a 0 \ 0" by (simp add: ratfps_nth_altdef) lemma ratfps_nth_0 [simp]: "ratfps_nth 0 n = 0" by (simp add: ratfps_nth_altdef) lemma fps_of_ratfps_cases: obtains p q where "coeff q 0 \ 0" "fps_of_ratfps f = fps_of_poly p / fps_of_poly q" by (rule that[of "snd (quot_of_ratfps f)" "fst (quot_of_ratfps f)"]) (simp_all add: fps_of_ratfps_altdef case_prod_unfold) lemma fps_of_ratfps_cutoff [simp]: "fps_of_poly (ratfps_cutoff n x) = fps_cutoff n (fps_of_ratfps x)" by (simp add: fps_eq_iff ratfps_cutoff_def nth_default_def ratfps_nth_altdef) lemma subdegree_fps_of_ratfps: "subdegree (fps_of_ratfps x) = ratfps_subdegree x" by transfer (simp_all add: case_prod_unfold subdegree_div_unit poly_subdegree_def) lemma ratfps_subdegree_altdef: "ratfps_subdegree x = subdegree (fps_of_ratfps x)" using subdegree_fps_of_ratfps .. end code_datatype fps_of_ratfps lemma fps_zero_code [code]: "0 = fps_of_ratfps 0" by simp lemma fps_one_code [code]: "1 = fps_of_ratfps 1" by simp lemma fps_const_code [code]: "fps_const c = fps_of_poly [:c:]" by simp lemma fps_of_poly_code [code]: "fps_of_poly p = fps_of_ratfps (ratfps_of_poly p)" by simp lemma fps_X_code [code]: "fps_X = fps_of_ratfps (ratfps_of_poly [:0,1:])" by simp lemma fps_nth_code [code]: "fps_nth (fps_of_ratfps x) n = ratfps_nth x n" by (simp add: ratfps_nth_altdef) lemma fps_uminus_code [code]: "- fps_of_ratfps x = fps_of_ratfps (-x)" by simp lemma fps_add_code [code]: "fps_of_ratfps x + fps_of_ratfps y = fps_of_ratfps (x + y)" by simp lemma fps_diff_code [code]: "fps_of_ratfps x - fps_of_ratfps y = fps_of_ratfps (x - y)" by simp lemma fps_mult_code [code]: "fps_of_ratfps x * fps_of_ratfps y = fps_of_ratfps (x * y)" by simp lemma fps_inverse_code [code]: "inverse (fps_of_ratfps x) = fps_of_ratfps (inverse x)" by simp lemma fps_cutoff_code [code]: "fps_cutoff n (fps_of_ratfps x) = fps_of_poly (ratfps_cutoff n x)" by simp lemmas subdegree_code [code] = subdegree_fps_of_ratfps lemma fractrel_normalize_quot: "fractrel p p \ fractrel q q \ fractrel (normalize_quot p) (normalize_quot q) \ fractrel p q" by (subst fractrel_normalize_quot_left fractrel_normalize_quot_right, simp)+ (rule refl) lemma fps_of_ratfps_eq_iff [simp]: "fps_of_ratfps p = fps_of_ratfps q \ p = q" proof - { fix p q :: "'a poly fract" assume "fractrel (quot_of_fract p) (quot_of_fract q)" hence "p = q" by transfer (simp only: fractrel_normalize_quot) } note A = this show ?thesis by transfer (auto simp: case_prod_unfold unit_eq_div1 unit_eq_div2 unit_div_commute intro: A) qed lemma fps_of_ratfps_eq_zero_iff [simp]: "fps_of_ratfps p = 0 \ p = 0" by (simp del: fps_of_ratfps_0 add: fps_of_ratfps_0 [symmetric]) lemma unit_factor_snd_quot_of_ratfps [simp]: "unit_factor (snd (quot_of_ratfps x)) = 1" by transfer simp lemma poly_shift_times_monom_le: "n \ m \ poly_shift n (monom c m * p) = monom c (m - n) * p" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma poly_shift_times_monom_ge: "n \ m \ poly_shift n (monom c m * p) = smult c (poly_shift (n - m) p)" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma poly_shift_times_monom: "poly_shift n (monom c n * p) = smult c p" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma monom_times_poly_shift: assumes "poly_subdegree p \ n" shows "monom c n * poly_shift n p = smult c p" (is "?lhs = ?rhs") proof (intro poly_eqI) fix k show "coeff ?lhs k = coeff ?rhs k" proof (cases "k < n") case True with assms have "k < poly_subdegree p" by simp hence "coeff p k = 0" by (simp add: coeff_less_poly_subdegree) thus ?thesis by (auto simp: coeff_monom_mult coeff_poly_shift) qed (auto simp: coeff_monom_mult coeff_poly_shift) qed lemma monom_times_poly_shift': assumes "poly_subdegree p \ n" shows "monom (1 :: 'a :: comm_semiring_1) n * poly_shift n p = p" by (simp add: monom_times_poly_shift[OF assms]) lemma subdegree_minus_cutoff_ge: assumes "f - fps_cutoff n (f :: 'a :: ab_group_add fps) \ 0" shows "subdegree (f - fps_cutoff n f) \ n" using assms by (rule subdegree_geI) simp_all lemma fps_shift_times_X_power'': "fps_shift n (fps_X ^ n * f :: 'a :: comm_ring_1 fps) = f" using fps_shift_times_fps_X_power'[of n f] by (simp add: mult.commute) lemma ratfps_shift_code [code abstract]: "quot_of_ratfps (ratfps_shift n x) = (let (a, b) = quot_of_ratfps (x - ratfps_of_poly (ratfps_cutoff n x)) in (poly_shift n a, b))" (is "?lhs1 = ?rhs1") and fps_of_ratfps_shift [simp]: "fps_of_ratfps (ratfps_shift n x) = fps_shift n (fps_of_ratfps x)" proof - include fps_notation define x' where "x' = ratfps_of_poly (ratfps_cutoff n x)" define y where "y = quot_of_ratfps (x - x')" have "coprime (fst y) (snd y)" unfolding y_def by transfer (rule coprime_quot_of_fract) also have fst_y: "fst y = monom 1 n * poly_shift n (fst y)" proof (cases "x = x'") case False have "poly_subdegree (fst y) = subdegree (fps_of_poly (fst y))" by (simp add: poly_subdegree_def) also have "\ = subdegree (fps_of_poly (fst y) / fps_of_poly (snd y))" by (subst subdegree_div_unit) (simp_all add: y_def) also have "fps_of_poly (fst y) / fps_of_poly (snd y) = fps_of_ratfps (x - x')" unfolding y_def by transfer (simp add: case_prod_unfold) also from False have "subdegree \ \ n" proof (intro subdegree_geI) fix k assume "k < n" thus "fps_of_ratfps (x - x') $ k = 0" by (simp add: x'_def) qed simp_all finally show ?thesis by (rule monom_times_poly_shift' [symmetric]) qed (simp_all add: y_def) finally have coprime: "coprime (poly_shift n (fst y)) (snd y)" by simp have "quot_of_ratfps (ratfps_shift n x) = quot_of_ratfps (quot_to_ratfps (poly_shift n (fst y), snd y))" by (simp add: ratfps_shift_def Let_def case_prod_unfold x'_def y_def) also from coprime have "\ = (poly_shift n (fst y), snd y)" by (intro quot_of_ratfps_quot_to_ratfps) (simp_all add: y_def normalized_fracts_def) also have "\ = ?rhs1" by (simp add: case_prod_unfold Let_def y_def x'_def) finally show eq: "?lhs1 = ?rhs1" . have "fps_shift n (fps_of_ratfps x) = fps_shift n (fps_of_ratfps (x - x'))" by (intro fps_ext) (simp_all add: x'_def) also have "fps_of_ratfps (x - x') = fps_of_poly (fst y) / fps_of_poly (snd y)" by (simp add: fps_of_ratfps_altdef y_def case_prod_unfold) also have "fps_shift n \ = fps_of_ratfps (ratfps_shift n x)" by (subst fst_y, subst fps_of_poly_mult, subst unit_div_mult_swap [symmetric]) (simp_all add: y_def fps_of_poly_monom fps_shift_times_X_power'' eq fps_of_ratfps_altdef case_prod_unfold Let_def x'_def) finally show "fps_of_ratfps (ratfps_shift n x) = fps_shift n (fps_of_ratfps x)" .. qed lemma fps_shift_code [code]: "fps_shift n (fps_of_ratfps x) = fps_of_ratfps (ratfps_shift n x)" by simp instantiation fps :: (equal) equal begin definition equal_fps :: "'a fps \ 'a fps \ bool" where [simp]: "equal_fps f g \ f = g" instance by standard simp_all end lemma equal_fps_code [code]: "HOL.equal (fps_of_ratfps f) (fps_of_ratfps g) \ f = g" by simp lemma fps_of_ratfps_divide [simp]: "fps_of_ratfps (f div g) = fps_of_ratfps f div fps_of_ratfps g" unfolding fps_divide_def Let_def by transfer' (simp add: Let_def ratfps_subdegree_altdef) lemma ratfps_eqI: "fps_of_ratfps x = fps_of_ratfps y \ x = y" by simp -instance ratfps :: ("{field,factorial_ring_gcd}") algebraic_semidom +instance ratfps :: ("field_gcd") algebraic_semidom by standard (auto intro: ratfps_eqI) lemma fps_of_ratfps_dvd [simp]: "fps_of_ratfps x dvd fps_of_ratfps y \ x dvd y" proof assume "fps_of_ratfps x dvd fps_of_ratfps y" hence "fps_of_ratfps y = fps_of_ratfps y div fps_of_ratfps x * fps_of_ratfps x" by simp also have "\ = fps_of_ratfps (y div x * x)" by simp finally have "y = y div x * x" by (subst (asm) fps_of_ratfps_eq_iff) thus "x dvd y" by (intro dvdI[of _ _ "y div x"]) (simp add: mult_ac) next assume "x dvd y" hence "y = y div x * x" by simp also have "fps_of_ratfps \ = fps_of_ratfps (y div x) * fps_of_ratfps x" by simp finally show "fps_of_ratfps x dvd fps_of_ratfps y" by (simp del: fps_of_ratfps_divide) qed lemma is_unit_ratfps_iff [simp]: "is_unit x \ ratfps_nth x 0 \ 0" proof assume "is_unit x" then obtain y where "1 = x * y" by (auto elim!: dvdE) hence "1 = fps_of_ratfps (x * y)" by (simp del: fps_of_ratfps_mult) also have "\ = fps_of_ratfps x * fps_of_ratfps y" by simp finally have "is_unit (fps_of_ratfps x)" by (rule dvdI[of _ _ "fps_of_ratfps y"]) thus "ratfps_nth x 0 \ 0" by (simp add: ratfps_nth_altdef) next assume "ratfps_nth x 0 \ 0" hence "fps_of_ratfps (x * inverse x) = 1" by (simp add: ratfps_nth_altdef inverse_mult_eq_1') also have "\ = fps_of_ratfps 1" by simp finally have "x * inverse x = 1" by (subst (asm) fps_of_ratfps_eq_iff) thus "is_unit x" by (intro dvdI[of _ _ "inverse x"]) simp_all qed -instantiation ratfps :: ("{field,factorial_ring_gcd}") normalization_semidom +instantiation ratfps :: ("field_gcd") normalization_semidom begin definition unit_factor_ratfps :: "'a ratfps \ 'a ratfps" where "unit_factor x = ratfps_shift (ratfps_subdegree x) x" definition normalize_ratfps :: "'a ratfps \ 'a ratfps" where "normalize x = (if x = 0 then 0 else ratfps_of_poly (monom 1 (ratfps_subdegree x)))" lemma fps_of_ratfps_unit_factor [simp]: "fps_of_ratfps (unit_factor x) = unit_factor (fps_of_ratfps x)" unfolding unit_factor_ratfps_def by (simp add: ratfps_subdegree_altdef) lemma fps_of_ratfps_normalize [simp]: "fps_of_ratfps (normalize x) = normalize (fps_of_ratfps x)" unfolding normalize_ratfps_def by (simp add: fps_of_poly_monom ratfps_subdegree_altdef) instance proof show "unit_factor x * normalize x = x" "normalize (0 :: 'a ratfps) = 0" "unit_factor (0 :: 'a ratfps) = 0" for x :: "'a ratfps" by (rule ratfps_eqI, simp add: ratfps_subdegree_code del: fps_of_ratfps_eq_iff fps_unit_factor_def fps_normalize_def)+ show "is_unit (unit_factor a)" if "a \ 0" for a :: "'a ratfps" using that by (auto simp: ratfps_nth_altdef) - show "unit_factor (a * b) = unit_factor a * unit_factor b" for a b :: "'a ratfps" - by (rule ratfps_eqI, insert unit_factor_mult[of "fps_of_ratfps a" "fps_of_ratfps b"]) - (simp del: fps_of_ratfps_eq_iff) + fix a b :: "'a ratfps" + assume "is_unit a" + thus "unit_factor (a * b) = a * unit_factor b" + by (intro ratfps_eqI, unfold fps_of_ratfps_unit_factor fps_of_ratfps_mult, + subst unit_factor_mult_unit_left) (auto simp: ratfps_nth_altdef) show "unit_factor a = a" if "is_unit a" for a :: "'a ratfps" by (rule ratfps_eqI) (insert that, auto simp: fps_of_ratfps_is_unit) qed end -instantiation ratfps :: ("{field,factorial_ring_gcd}") semidom_modulo +instance ratfps :: ("field_gcd") normalization_semidom_multiplicative +proof + show "unit_factor (a * b) = unit_factor a * unit_factor b" for a b :: "'a ratfps" + by (rule ratfps_eqI, insert unit_factor_mult[of "fps_of_ratfps a" "fps_of_ratfps b"]) + (simp del: fps_of_ratfps_eq_iff) +qed + +instantiation ratfps :: ("field_gcd") semidom_modulo begin lift_definition modulo_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "\f g. if g = 0 then f else let n = ratfps_subdegree g; h = ratfps_shift n g in ratfps_of_poly (ratfps_cutoff n (f * inverse h)) * h" . lemma fps_of_ratfps_mod [simp]: "fps_of_ratfps (f mod g :: 'a ratfps) = fps_of_ratfps f mod fps_of_ratfps g" unfolding fps_mod_def by transfer' (simp add: Let_def ratfps_subdegree_altdef) instance by standard (auto intro: ratfps_eqI) end -instantiation ratfps :: ("{field,factorial_ring_gcd}") euclidean_ring +instantiation ratfps :: ("field_gcd") euclidean_ring begin definition euclidean_size_ratfps :: "'a ratfps \ nat" where "euclidean_size_ratfps x = (if x = 0 then 0 else 2 ^ ratfps_subdegree x)" lemma fps_of_ratfps_euclidean_size [simp]: "euclidean_size x = euclidean_size (fps_of_ratfps x)" unfolding euclidean_size_ratfps_def fps_euclidean_size_def by (simp add: ratfps_subdegree_altdef) instance proof show "euclidean_size (0 :: 'a ratfps) = 0" by simp show "euclidean_size (a mod b) < euclidean_size b" "euclidean_size a \ euclidean_size (a * b)" if "b \ 0" for a b :: "'a ratfps" using that by (simp_all add: mod_size_less size_mult_mono) qed end -instantiation ratfps :: ("{field,factorial_ring_gcd}") euclidean_ring_cancel +instantiation ratfps :: ("field_gcd") euclidean_ring_cancel begin instance by standard (auto intro: ratfps_eqI) end lemma quot_of_ratfps_eq_iff [simp]: "quot_of_ratfps x = quot_of_ratfps y \ x = y" by transfer simp lemma ratfps_eq_0_code: "x = 0 \ fst (quot_of_ratfps x) = 0" proof assume "fst (quot_of_ratfps x) = 0" moreover have "coprime (fst (quot_of_ratfps x)) (snd (quot_of_ratfps x))" by transfer (simp add: coprime_quot_of_fract) moreover have "normalize (snd (quot_of_ratfps x)) = snd (quot_of_ratfps x)" by (simp add: div_unit_factor [symmetric] del: div_unit_factor) ultimately have "quot_of_ratfps x = (0,1)" by (simp add: prod_eq_iff normalize_idem_imp_is_unit_iff) also have "\ = quot_of_ratfps 0" by simp finally show "x = 0" by (subst (asm) quot_of_ratfps_eq_iff) qed simp_all lemma fps_dvd_code [code_unfold]: - "x dvd y \ y = 0 \ ((x::'a::{field,factorial_ring_gcd} fps) \ 0 \ subdegree x \ subdegree y)" + "x dvd y \ y = 0 \ ((x::'a::field_gcd fps) \ 0 \ subdegree x \ subdegree y)" using fps_dvd_iff[of x y] by (cases "x = 0") auto lemma ratfps_dvd_code [code_unfold]: "x dvd y \ y = 0 \ (x \ 0 \ ratfps_subdegree x \ ratfps_subdegree y)" using fps_dvd_code [of "fps_of_ratfps x" "fps_of_ratfps y"] by (simp add: ratfps_subdegree_altdef) -instance ratfps :: ("{field,factorial_ring_gcd}") normalization_euclidean_semiring .. +instance ratfps :: ("field_gcd") normalization_euclidean_semiring .. -instantiation ratfps :: ("{field,factorial_ring_gcd}") euclidean_ring_gcd +instantiation ratfps :: ("field_gcd") euclidean_ring_gcd begin definition "gcd_ratfps = (Euclidean_Algorithm.gcd :: 'a ratfps \ _)" definition "lcm_ratfps = (Euclidean_Algorithm.lcm :: 'a ratfps \ _)" definition "Gcd_ratfps = (Euclidean_Algorithm.Gcd :: 'a ratfps set \ _)" definition "Lcm_ratfps = (Euclidean_Algorithm.Lcm:: 'a ratfps set \ _)" instance by standard (simp_all add: gcd_ratfps_def lcm_ratfps_def Gcd_ratfps_def Lcm_ratfps_def) end lemma ratfps_eq_0_iff: "x = 0 \ fps_of_ratfps x = 0" using fps_of_ratfps_eq_iff[of x 0] unfolding fps_of_ratfps_0 by simp lemma ratfps_of_poly_eq_0_iff: "ratfps_of_poly x = 0 \ x = 0" by (auto simp: ratfps_eq_0_iff) lemma ratfps_gcd: assumes [simp]: "f \ 0" "g \ 0" shows "gcd f g = ratfps_of_poly (monom 1 (min (ratfps_subdegree f) (ratfps_subdegree g)))" by (rule sym, rule gcdI) (auto simp: ratfps_subdegree_altdef ratfps_dvd_code subdegree_fps_of_poly ratfps_of_poly_eq_0_iff normalize_ratfps_def) -lemma ratfps_gcd_altdef: "gcd (f :: 'a :: {field,factorial_ring_gcd} ratfps) g = +lemma ratfps_gcd_altdef: "gcd (f :: 'a :: field_gcd ratfps) g = (if f = 0 \ g = 0 then 0 else if f = 0 then ratfps_of_poly (monom 1 (ratfps_subdegree g)) else if g = 0 then ratfps_of_poly (monom 1 (ratfps_subdegree f)) else ratfps_of_poly (monom 1 (min (ratfps_subdegree f) (ratfps_subdegree g))))" by (simp add: ratfps_gcd normalize_ratfps_def) lemma ratfps_lcm: assumes [simp]: "f \ 0" "g \ 0" shows "lcm f g = ratfps_of_poly (monom 1 (max (ratfps_subdegree f) (ratfps_subdegree g)))" by (rule sym, rule lcmI) (auto simp: ratfps_subdegree_altdef ratfps_dvd_code subdegree_fps_of_poly ratfps_of_poly_eq_0_iff normalize_ratfps_def) -lemma ratfps_lcm_altdef: "lcm (f :: 'a :: {field,factorial_ring_gcd} ratfps) g = +lemma ratfps_lcm_altdef: "lcm (f :: 'a :: field_gcd ratfps) g = (if f = 0 \ g = 0 then 0 else ratfps_of_poly (monom 1 (max (ratfps_subdegree f) (ratfps_subdegree g))))" by (simp add: ratfps_lcm) lemma ratfps_Gcd: assumes "A - {0} \ {}" shows "Gcd A = ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f))" proof (rule sym, rule GcdI) fix f assume "f \ A" thus "ratfps_of_poly (monom 1 (INF f\A - {0}. ratfps_subdegree f)) dvd f" by (cases "f = 0") (auto simp: ratfps_dvd_code ratfps_of_poly_eq_0_iff ratfps_subdegree_altdef subdegree_fps_of_poly intro!: cINF_lower) next fix d assume d: "\f. f \ A \ d dvd f" from assms obtain f where "f \ A - {0}" by auto with d[of f] have [simp]: "d \ 0" by auto from d assms have "ratfps_subdegree d \ (INF f\A-{0}. ratfps_subdegree f)" by (intro cINF_greatest) (auto simp: ratfps_dvd_code) with d assms show "d dvd ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f))" by (simp add: ratfps_dvd_code ratfps_subdegree_altdef subdegree_fps_of_poly) qed (simp_all add: ratfps_subdegree_altdef subdegree_fps_of_poly normalize_ratfps_def) -lemma ratfps_Gcd_altdef: "Gcd (A :: 'a :: {field,factorial_ring_gcd} ratfps set) = +lemma ratfps_Gcd_altdef: "Gcd (A :: 'a :: field_gcd ratfps set) = (if A \ {0} then 0 else ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f)))" using ratfps_Gcd by auto lemma ratfps_Lcm: assumes "A \ {}" "0 \ A" "bdd_above (ratfps_subdegree`A)" shows "Lcm A = ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f))" proof (rule sym, rule LcmI) fix f assume "f \ A" moreover from assms(3) have "bdd_above (ratfps_subdegree ` A)" by auto ultimately show "f dvd ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f))" using assms(2) by (cases "f = 0") (auto simp: ratfps_dvd_code ratfps_of_poly_eq_0_iff subdegree_fps_of_poly ratfps_subdegree_altdef [abs_def] intro!: cSUP_upper) next fix d assume d: "\f. f \ A \ f dvd d" from assms obtain f where f: "f \ A" "f \ 0" by auto show "ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f)) dvd d" proof (cases "d = 0") assume "d \ 0" moreover from d have "\f. f \ A \ f \ 0 \ f dvd d" by blast ultimately have "ratfps_subdegree d \ (SUP f\A. ratfps_subdegree f)" using assms by (intro cSUP_least) (auto simp: ratfps_dvd_code) with \d \ 0\ show ?thesis by (simp add: ratfps_dvd_code ratfps_of_poly_eq_0_iff ratfps_subdegree_altdef subdegree_fps_of_poly) qed simp_all qed (simp_all add: ratfps_subdegree_altdef subdegree_fps_of_poly normalize_ratfps_def) lemma ratfps_Lcm_altdef: - "Lcm (A :: 'a :: {field,factorial_ring_gcd} ratfps set) = + "Lcm (A :: 'a :: field_gcd ratfps set) = (if 0 \ A \ \bdd_above (ratfps_subdegree`A) then 0 else if A = {} then 1 else ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f)))" proof (cases "bdd_above (ratfps_subdegree`A)") assume unbounded: "\bdd_above (ratfps_subdegree`A)" have "Lcm A = 0" proof (rule ccontr) assume "Lcm A \ 0" from unbounded obtain f where f: "f \ A" "ratfps_subdegree (Lcm A) < ratfps_subdegree f" unfolding bdd_above_def by (auto simp: not_le) moreover from this and \Lcm A \ 0\ have "ratfps_subdegree f \ ratfps_subdegree (Lcm A)" using dvd_Lcm[of f A] by (auto simp: ratfps_dvd_code) ultimately show False by simp qed with unbounded show ?thesis by simp qed (simp_all add: ratfps_Lcm Lcm_eq_0_I) lemma fps_of_ratfps_quot_to_ratfps: "coeff y 0 \ 0 \ fps_of_ratfps (quot_to_ratfps (x,y)) = fps_of_poly x / fps_of_poly y" proof (transfer, goal_cases) case (1 y x) define x' y' where "x' = fst (normalize_quot (x,y))" and "y' = snd (normalize_quot (x,y))" from 1 have nz: "y \ 0" by auto have eq: "normalize_quot (x', y') = (x', y')" by (simp add: x'_def y'_def) from normalize_quotE[OF nz, of x] guess d . note d = this [folded x'_def y'_def] have "(case quot_of_fract (if coeff y' 0 = 0 then 0 else quot_to_fract (x', y')) of (a, b) \ fps_of_poly a / fps_of_poly b) = fps_of_poly x / fps_of_poly y" using d eq 1 by (auto simp: case_prod_unfold fps_of_poly_simps quot_of_fract_quot_to_fract Let_def coeff_0_mult) thus ?case by (auto simp add: Let_def case_prod_unfold x'_def y'_def) qed lemma fps_of_ratfps_quot_to_ratfps_code_post1: "fps_of_ratfps (quot_to_ratfps (x,pCons 1 y)) = fps_of_poly x / fps_of_poly (pCons 1 y)" "fps_of_ratfps (quot_to_ratfps (x,pCons (-1) y)) = fps_of_poly x / fps_of_poly (pCons (-1) y)" by (simp_all add: fps_of_ratfps_quot_to_ratfps) lemma fps_of_ratfps_quot_to_ratfps_code_post2: - "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,factorial_ring_gcd} poly,pCons (numeral n) y')) = + "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,field_gcd} poly,pCons (numeral n) y')) = fps_of_poly x' / fps_of_poly (pCons (numeral n) y')" - "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,factorial_ring_gcd} poly,pCons (-numeral n) y')) = + "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,field_gcd} poly,pCons (-numeral n) y')) = fps_of_poly x' / fps_of_poly (pCons (-numeral n) y')" by (simp_all add: fps_of_ratfps_quot_to_ratfps) lemmas fps_of_ratfps_quot_to_ratfps_code_post [code_post] = fps_of_ratfps_quot_to_ratfps_code_post1 fps_of_ratfps_quot_to_ratfps_code_post2 lemma fps_dehorner: fixes a b c :: "'a :: semiring_1 fps" and d e f :: "'b :: ring_1 fps" shows "(b + c) * fps_X = b * fps_X + c * fps_X" "(a * fps_X) * fps_X = a * fps_X ^ 2" "a * fps_X ^ m * fps_X = a * fps_X ^ (Suc m)" "a * fps_X * fps_X ^ m = a * fps_X ^ (Suc m)" "a * fps_X^m * fps_X^n = a * fps_X^(m+n)" "a + (b + c) = a + b + c" "a * 1 = a" "1 * a = a" "d + - e = d - e" "(-d) * e = - (d * e)" "d + (e - f) = d + e - f" "(d - e) * fps_X = d * fps_X - e * fps_X" "fps_X * fps_X = fps_X^2" "fps_X * fps_X^m = fps_X^(Suc m)" "fps_X^m * fps_X = fps_X^Suc m" "fps_X^m * fps_X^n = fps_X^(m + n)" by (simp_all add: algebra_simps power2_eq_square power_add power_commutes) lemma fps_divide_1: "(a :: 'a :: field fps) / 1 = a" by simp lemmas fps_of_poly_code_post [code_post] = fps_of_poly_simps fps_const_0_eq_0 fps_const_1_eq_1 numeral_fps_const [symmetric] fps_const_neg [symmetric] fps_const_divide [symmetric] fps_dehorner Suc_numeral arith_simps fps_divide_1 definition (in term_syntax) valterm_ratfps :: - "'a ::{field,factorial_ring_gcd,typerep} poly \ (unit \ Code_Evaluation.term) \ + "'a ::{field_gcd, typerep} poly \ (unit \ Code_Evaluation.term) \ 'a poly \ (unit \ Code_Evaluation.term) \ 'a ratfps \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_ratfps k l = Code_Evaluation.valtermify (/) {\} (Code_Evaluation.valtermify ratfps_of_poly {\} k) {\} (Code_Evaluation.valtermify ratfps_of_poly {\} l)" notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) -instantiation ratfps :: ("{field,factorial_ring_gcd,random}") random +instantiation ratfps :: ("{field_gcd,random}") random begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\num::'a poly \ (unit \ term). Quickcheck_Random.random i \\ (\denom::'a poly \ (unit \ term). Pair (let denom = (if fst denom = 0 then Code_Evaluation.valtermify 1 else denom) in valterm_ratfps num denom)))" instance .. end no_notation fcomp (infixl "\>" 60) no_notation scomp (infixl "\\" 60) instantiation ratfps :: ("{field,factorial_ring_gcd,exhaustive}") exhaustive begin definition "exhaustive_ratfps f d = Quickcheck_Exhaustive.exhaustive (\num. Quickcheck_Exhaustive.exhaustive (\denom. f ( let denom = if denom = 0 then 1 else denom in ratfps_of_poly num / ratfps_of_poly denom)) d) d" instance .. end -instantiation ratfps :: ("{field,factorial_ring_gcd,full_exhaustive}") full_exhaustive +instantiation ratfps :: ("{field_gcd,full_exhaustive}") full_exhaustive begin definition "full_exhaustive_ratfps f d = Quickcheck_Exhaustive.full_exhaustive (\num::'a poly \ (unit \ term). Quickcheck_Exhaustive.full_exhaustive (\denom::'a poly \ (unit \ term). f (let denom = if fst denom = 0 then Code_Evaluation.valtermify 1 else denom in valterm_ratfps num denom)) d) d" instance .. end quickcheck_generator fps constructors: fps_of_ratfps end diff --git a/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy b/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy --- a/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy +++ b/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy @@ -1,398 +1,398 @@ (* File: Rational_FPS_Asymptotics.thy Author: Manuel Eberl, TU MĂ¼nchen *) theory Rational_FPS_Asymptotics imports "HOL-Library.Landau_Symbols" "Polynomial_Factorization.Square_Free_Factorization" "HOL-Real_Asymp.Real_Asymp" "Count_Complex_Roots.Count_Complex_Roots" Linear_Homogenous_Recurrences Linear_Inhomogenous_Recurrences RatFPS Rational_FPS_Solver "HOL-Library.Code_Target_Numeral" begin lemma poly_asymp_equiv: assumes "p \ 0" and "F \ at_infinity" shows "poly p \[F] (\x. lead_coeff p * x ^ degree p)" proof - have poly_pCons': "poly (pCons a q) = (\x. a + x * poly q x)" for a :: 'a and q by (simp add: fun_eq_iff) show ?thesis using assms(1) proof (induction p) case (pCons a p) define n where "n = Suc (degree p)" show ?case proof (cases "p = 0") case [simp]: False hence *: "poly p \[F] (\x. lead_coeff p * x ^ degree p)" by (intro pCons.IH) have "poly (pCons a p) = (\x. a + x * poly p x)" by (simp add: poly_pCons') moreover have "\ \[F] (\x. lead_coeff p * x ^ n)" proof (subst asymp_equiv_add_left) have "(\x. x * poly p x) \[F] (\x. x * (lead_coeff p * x ^ degree p))" by (intro asymp_equiv_intros *) also have "\ = (\x. lead_coeff p * x ^ n)" by (simp add: n_def mult_ac) finally show "(\x. x * poly p x) \[F] \" . next have "filterlim (\x. x) at_infinity F" by (simp add: filterlim_def assms) hence "(\x. x ^ n) \ \[F](\_. 1 :: 'a)" unfolding smallomega_1_conv_filterlim by (intro Limits.filterlim_power_at_infinity filterlim_ident) (auto simp: n_def) hence "(\x. a) \ o[F](\x. x ^ n)" unfolding smallomega_iff_smallo[symmetric] by (cases "a = 0") auto thus "(\x. a) \ o[F](\x. lead_coeff p * x ^ n)" by simp qed ultimately show ?thesis by (simp add: n_def) qed auto qed auto qed lemma poly_bigtheta: assumes "p \ 0" and "F \ at_infinity" shows "poly p \ \[F](\x. x ^ degree p)" proof - have "poly p \[F] (\x. lead_coeff p * x ^ degree p)" by (intro poly_asymp_equiv assms) thus ?thesis using assms by (auto dest!: asymp_equiv_imp_bigtheta) qed lemma poly_bigo: assumes "F \ at_infinity" and "degree p \ k" shows "poly p \ O[F](\x. x ^ k)" proof (cases "p = 0") case True hence "poly p = (\_. 0)" by (auto simp: fun_eq_iff) thus ?thesis by simp next case False have *: "(\x. x ^ (k - degree p)) \ \[F](\x. 1)" proof (cases "k = degree p") case False hence "(\x. x ^ (k - degree p)) \ \[F](\_. 1)" unfolding smallomega_1_conv_filterlim using assms False by (intro Limits.filterlim_power_at_infinity filterlim_ident) (auto simp: filterlim_def) thus ?thesis by (rule landau_omega.small_imp_big) qed auto have "poly p \ \[F](\x. x ^ degree p * 1)" using poly_bigtheta[OF False assms(1)] by simp also have "(\x. x ^ degree p * 1) \ O[F](\x. x ^ degree p * x ^ (k - degree p))" using * by (intro landau_o.big.mult landau_o.big_refl) (auto simp: bigomega_iff_bigo) also have "(\x::'a. x ^ degree p * x ^ (k - degree p)) = (\x. x ^ k)" using assms by (simp add: power_add [symmetric]) finally show ?thesis . qed lemma reflect_poly_dvdI: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "reflect_poly p dvd reflect_poly q" using assms by (auto simp: reflect_poly_mult) lemma smult_altdef: "smult c p = [:c:] * p" by (induction p) (auto simp: mult_ac) lemma smult_power: "smult (c ^ n) (p ^ n) = (smult c p) ^ n" proof - have "smult (c ^ n) (p ^ n) = [:c ^ n:] * p ^ n" by simp also have "[:c:] ^ n = [:c ^ n:]" by (induction n) (auto simp: mult_ac) hence "[:c ^ n:] = [:c:] ^ n" .. also have "\ * p ^ n = ([:c:] * p) ^ n" by (rule power_mult_distrib [symmetric]) also have "\ = (smult c p) ^ n" by simp finally show ?thesis . qed lemma order_reflect_poly_ge: fixes c :: "'a :: field" assumes "c \ 0" and "p \ 0" shows "order c (reflect_poly p) \ order (1 / c) p" proof - have "reflect_poly ([:-(1 / c), 1:] ^ order (1 / c) p) dvd reflect_poly p" by (intro reflect_poly_dvdI, subst order_divides) auto also have "reflect_poly ([:-(1 / c), 1:] ^ order (1 / c) p) = smult ((-1 / c) ^ order (1 / c) p) ([:-c, 1:] ^ order (1 / c) p)" using assms by (simp add: reflect_poly_power reflect_poly_pCons smult_power) finally have "([:-c, 1:] ^ order (1 / c) p) dvd reflect_poly p" by (rule smult_dvd_cancel) with \p \ 0\ show ?thesis by (subst (asm) order_divides) auto qed lemma order_reflect_poly: fixes c :: "'a :: field" assumes "c \ 0" and "coeff p 0 \ 0" shows "order c (reflect_poly p) = order (1 / c) p" proof (rule antisym) from assms show "order c (reflect_poly p) \ order (1 / c) p" by (intro order_reflect_poly_ge) auto next from assms have "order (1 / (1 / c)) (reflect_poly p) \ order (1 / c) (reflect_poly (reflect_poly p))" by (intro order_reflect_poly_ge) auto with assms show "order c (reflect_poly p) \ order (1 / c) p" by simp qed lemma poly_reflect_eq_0_iff: "poly (reflect_poly p) (x :: 'a :: field) = 0 \ p = 0 \ x \ 0 \ poly p (1 / x) = 0" by (cases "x = 0") (auto simp: poly_reflect_poly_nz inverse_eq_divide) theorem ratfps_nth_bigo: fixes q :: "complex poly" assumes "R > 0" assumes roots1: "\z. z \ ball 0 (1 / R) \ poly q z \ 0" assumes roots2: "\z. z \ sphere 0 (1 / R) \ poly q z = 0 \ order z q \ Suc k" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof - define q' where "q' = reflect_poly q" from roots1[of 0] and \R > 0\ have [simp]: "coeff q 0 \ 0" "q \ 0" by (auto simp: poly_0_coeff_0) from ratfps_closed_form_exists[OF this(1), of p] guess r rs . note closed_form = this have "fps_nth (fps_of_poly p / fps_of_poly q) = (\n. coeff r n + (\c | poly q' c = 0. poly (rs c) (of_nat n) * c ^ n))" by (intro ext, subst closed_form) (simp_all add: q'_def) also have "\ \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (intro sum_in_bigo big_sum_in_bigo) have "eventually (\n. coeff r n = 0) at_top" using MOST_nat coeff_eq_0 cofinite_eq_sequentially by force hence "coeff r \ \(\_. 0)" by (rule bigthetaI_cong) also have "(\_. 0 :: complex) \ O(\n. of_nat n ^ k * of_real R ^ n)" by simp finally show "coeff r \ O(\n. of_nat n ^ k * of_real R ^ n)" . next fix c assume c: "c \ {c. poly q' c = 0}" hence [simp]: "c \ 0" by (auto simp: q'_def) show "(\n. poly (rs c) n * c ^ n) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (cases "norm c = R") case True \\The case of a root at the border of the disc\ show ?thesis proof (intro landau_o.big.mult landau_o.big.compose[OF poly_bigo tendsto_of_nat]) have "degree (rs c) \ order c (reflect_poly q) - 1" using c by (intro closed_form(2)) (auto simp: q'_def) also have "order c (reflect_poly q) = order (1 / c) q" using c by (intro order_reflect_poly) (auto simp: q'_def) also { have "order (1 / c) q \ Suc k" using \R > 0\ and True and c by (intro roots2) (auto simp: q'_def norm_divide poly_reflect_eq_0_iff) moreover have "order (1 / c) q \ 0" using order_root[of q "1 / c"] c by (auto simp: q'_def poly_reflect_eq_0_iff) ultimately have "order (1 / c) q - 1 \ k" by simp } finally show "degree (rs c) \ k" . next have "(\n. norm (c ^ n)) \ O(\n. norm (complex_of_real R ^ n))" using True and \R > 0\ by (simp add: norm_power) thus "(\n. c ^ n) \ O(\n. complex_of_real R ^ n)" by (subst (asm) landau_o.big.norm_iff) qed auto next case False \\The case of a root in the interior of the disc\ hence "norm c < R" using c and roots1[of "1/c"] and \R > 0\ by (cases "norm c" R rule: linorder_cases) (auto simp: q'_def poly_reflect_eq_0_iff norm_divide field_simps) define l where "l = degree (rs c)" have "(\n. poly (rs c) (of_nat n) * c ^ n) \ O(\n. of_nat n ^ l * c ^ n)" by (intro landau_o.big.mult landau_o.big.compose[OF poly_bigo tendsto_of_nat]) (auto simp: l_def) also have "(\n. of_nat n ^ l * c ^ n) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (subst landau_o.big.norm_iff [symmetric]) have "(\n. real n ^ l) \ O(\n. real n ^ k * (R / norm c) ^ n)" using \norm c < R\ and \R > 0\ by real_asymp hence "(\n. real n ^ l * norm c ^ n) \ O(\n. real n ^ k * R ^ n)" by (simp add: power_divide landau_o.big.divide_eq1) thus "(\x. norm (of_nat x ^ l * c ^ x)) \ O(\x. norm (of_nat x ^ k * complex_of_real R ^ x))" unfolding norm_power norm_mult using \R > 0\ by simp qed finally show ?thesis . qed qed finally show ?thesis . qed lemma order_power: "p \ 0 \ order c (p ^ n) = n * order c p" by (induction n) (auto simp: order_mult) lemma same_root_imp_not_coprime: - assumes "poly p x = 0" and "poly q (x :: 'a :: {factorial_ring_gcd}) = 0" + assumes "poly p x = 0" and "poly q (x :: 'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize}) = 0" shows "\coprime p q" proof assume "coprime p q" from assms have "[:-x, 1:] dvd p" and "[:-x, 1:] dvd q" by (simp_all add: poly_eq_0_iff_dvd) hence "[:-x, 1:] dvd gcd p q" by (simp add: poly_eq_0_iff_dvd) also from \coprime p q\ have "gcd p q = 1" by (rule coprime_imp_gcd_eq_1) finally show False by (elim is_unit_polyE) auto qed lemma ratfps_nth_bigo_square_free_factorization: fixes p :: "complex poly" assumes "square_free_factorization q (b, cs)" assumes "q \ 0" and "R > 0" assumes roots1: "\c l. (c, l) \ set cs \ \x\ball 0 (1 / R). poly c x \ 0" assumes roots2: "\c l. (c, l) \ set cs \ l > k \ \x\sphere 0 (1 / R). poly c x \ 0" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof - from assms(1) have q: "q = smult b (\(c, l)\set cs. c ^ Suc l)" unfolding square_free_factorization_def prod.case by blast with \q \ 0\ have [simp]: "b \ 0" by auto from assms(1) have [simp]: "(0, x) \ set cs" for x by (auto simp: square_free_factorization_def) from assms(1) have coprime: "c1 = c2" "m = n" if "\coprime c1 c2" "(c1, m) \ set cs" "(c2, n) \ set cs" for c1 c2 m n using that by (auto simp: square_free_factorization_def case_prod_unfold) show ?thesis proof (rule ratfps_nth_bigo) fix z :: complex assume z: "z \ ball 0 (1 / R)" show "poly q z \ 0" proof assume "poly q z = 0" then obtain c l where cl: "(c, l) \ set cs" and "poly c z = 0" by (auto simp: q poly_prod image_iff) with roots1[of c l] and z show False by auto qed next fix z :: complex assume z: "z \ sphere 0 (1 / R)" have order: "order z q = order z (\(c, l)\set cs. c ^ Suc l)" by (simp add: order_smult q) also have "\ = (\x\set cs. order z (case x of (c, l) \ c ^ Suc l))" by (subst order_prod) (auto dest: coprime) also have "\ = (\(c, l)\set cs. Suc l * order z c)" unfolding case_prod_unfold by (intro sum.cong refl, subst order_power) auto finally have "order z q = \" . show "order z q \ Suc k" proof (cases "\c0 l0. (c0, l0) \ set cs \ poly c0 z = 0") case False have "order z q = (\(c, l)\set cs. Suc l * order z c)" by fact also have "order z c = 0" if "(c, l) \ set cs" for c l using False that by (auto simp: order_root) hence "(\(c, l)\set cs. Suc l * order z c) = 0" by (intro sum.neutral) auto finally show "order z q \ Suc k" by simp next case True \\The order of a root is determined by the unique polynomial in the square-free factorisation that contains it.\ then obtain c0 l0 where cl0: "(c0, l0) \ set cs" "poly c0 z = 0" by blast have "order z q = (\(c, l)\set cs. Suc l * order z c)" by fact also have "\ = Suc l0 * order z c0 + (\(c, l) \ set cs - {(c0, l0)}. Suc l * order z c)" using cl0 by (subst sum.remove[of _ "(c0, l0)"]) auto also have "(\(c, l) \ set cs - {(c0, l0)}. Suc l * order z c) = 0" proof (intro sum.neutral ballI, goal_cases) case (1 cl) then obtain c l where [simp]: "cl = (c, l)" and cl: "(c, l) \ set cs" "(c0, l0) \ (c, l)" by (cases cl) auto from cl and cl0 and coprime[of c c0 l l0] have "coprime c c0" by auto with same_root_imp_not_coprime[of c z c0] and cl0 have "poly c z \ 0" by auto thus ?case by (auto simp: order_root) qed also have "square_free c0" using cl0 assms(1) by (auto simp: square_free_factorization_def) hence "rsquarefree c0" by (rule square_free_rsquarefree) with cl0 have "order z c0 = 1" by (auto simp: rsquarefree_def' order_root intro: antisym) finally have "order z q = Suc l0" by simp also from roots2[of c0 l0] cl0 z have "l0 \ k" by (cases l0 k rule: linorder_cases) auto finally show "order z q \ Suc k" by simp qed qed fact+ qed find_consts name:"Count_Complex" term proots_ball_card term proots_sphere_card lemma proots_within_card_zero_iff: assumes "p \ (0 :: 'a :: idom poly)" shows "card (proots_within p A) = 0 \ (\x\A. poly p x \ 0)" using assms by (subst card_0_eq) (auto intro: finite_proots) lemma ratfps_nth_bigo_square_free_factorization': fixes p :: "complex poly" assumes "square_free_factorization q (b, cs)" assumes "q \ 0" and "R > 0" assumes roots1: "list_all (\cl. proots_ball_card (fst cl) 0 (1 / R) = 0) cs" assumes roots2: "list_all (\cl. proots_sphere_card (fst cl) 0 (1 / R) = 0) (filter (\cl. snd cl > k) cs)" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (rule ratfps_nth_bigo_square_free_factorization[OF assms(1)]) from assms(1) have q: "q = smult b (\(c, l)\set cs. c ^ Suc l)" unfolding square_free_factorization_def prod.case by blast with \q \ 0\ have [simp]: "b \ 0" by auto from assms(1) have [simp]: "(0, x) \ set cs" for x by (auto simp: square_free_factorization_def) show "\x\ball 0 (1 / R). poly c x \ 0" if "(c, l) \ set cs" for c l proof - from roots1 that have "card (proots_within c (ball 0 (1 / R))) = 0" by (auto simp: proots_ball_card_def list_all_def) with that show ?thesis by (subst (asm) proots_within_card_zero_iff) auto qed show "\x\sphere 0 (1 / R). poly c x \ 0" if "(c, l) \ set cs" "l > k" for c l proof - from roots2 that have "card (proots_within c (sphere 0 (1 / R))) = 0" by (auto simp: proots_sphere_card_def list_all_def) with that show ?thesis by (subst (asm) proots_within_card_zero_iff) auto qed qed fact+ definition ratfps_has_asymptotics where "ratfps_has_asymptotics q k R \ q \ 0 \ R > 0 \ (let cs = snd (yun_factorization gcd q) in list_all (\cl. proots_ball_card (fst cl) 0 (1 / R) = 0) cs \ list_all (\cl. proots_sphere_card (fst cl) 0 (1 / R) = 0) (filter (\cl. snd cl > k) cs))" lemma ratfps_has_asymptotics_correct: assumes "ratfps_has_asymptotics q k R" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (rule ratfps_nth_bigo_square_free_factorization') show "square_free_factorization q (fst (yun_factorization gcd q), snd (yun_factorization gcd q))" by (rule yun_factorization) simp qed (insert assms, auto simp: ratfps_has_asymptotics_def Let_def list_all_def) value "map (fps_nth (fps_of_poly [:0, 1:] / fps_of_poly [:1, -1, -1 :: real:])) [0..<5]" method ratfps_bigo = (rule ratfps_has_asymptotics_correct; eval) lemma "fps_nth (fps_of_poly [:0, 1:] / fps_of_poly [:1, -1, -1 :: complex:]) \ O(\n. of_nat n ^ 0 * complex_of_real 1.618034 ^ n)" by ratfps_bigo lemma "fps_nth (fps_of_poly 1 / fps_of_poly [:1, -3, 3, -1 :: complex:]) \ O(\n. of_nat n ^ 2 * complex_of_real 1 ^ n)" by ratfps_bigo lemma "fps_nth (fps_of_poly f / fps_of_poly [:5, 4, 3, 2, 1 :: complex:]) \ O(\n. of_nat n ^ 0 * complex_of_real 0.69202 ^ n)" by ratfps_bigo end diff --git a/thys/Linear_Recurrences/Rational_FPS_Solver.thy b/thys/Linear_Recurrences/Rational_FPS_Solver.thy --- a/thys/Linear_Recurrences/Rational_FPS_Solver.thy +++ b/thys/Linear_Recurrences/Rational_FPS_Solver.thy @@ -1,381 +1,381 @@ (* File: Rational_FPS_Solver.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Solver for rational formal power series\ theory Rational_FPS_Solver imports Complex_Main Pochhammer_Polynomials Partial_Fraction_Decomposition Factorizations "HOL-Computational_Algebra.Field_as_Ring" begin text \ We can determine the $k$-th coefficient of an FPS of the form $d/(1-cX)^n$, which is an important step in solving linear recurrences. The $k$-th coefficient of such an FPS is always of the form $p(k) c^k$ where $p$ is the following polynomial: \ definition inverse_irred_power_poly :: "'a :: field_char_0 \ nat \ 'a poly" where "inverse_irred_power_poly d n = Poly [(d * of_nat (stirling n (k+1))) / (fact (n - 1)). k \ [0.. 0" shows "fps_const d / ((1 - fps_const (c :: 'a :: field_char_0) * fps_X) ^ n) = Abs_fps (\k. poly (inverse_irred_power_poly d n) (of_nat k) * c^k)" (is "?lhs = ?rhs") proof (rule fps_ext) include fps_notation fix k :: nat let ?p = "smult (d / (fact (n - 1))) (pcompose (pochhammer_poly (n - 1)) [:1,1:])" from n have "?lhs = fps_const d * inverse ((1 - fps_const c * fps_X) ^ n)" by (subst fps_divide_unit) auto also have "inverse ((1 - fps_const c * fps_X) ^ n) = Abs_fps (\k. of_nat ((n + k - 1) choose k) * c^k)" by (intro one_minus_const_fps_X_neg_power' n) also have "(fps_const d * \) $ k = d * of_nat ((n + k - 1) choose k) * c^k" by simp also from n have "(n + k - 1 choose k) = (n + k - 1 choose (n - 1))" by (subst binomial_symmetric) simp_all also from n have "of_nat \ = (pochhammer (of_nat k + 1) (n - 1) / fact (n - 1) :: 'a)" by (simp_all add: binomial_gbinomial gbinomial_pochhammer' of_nat_diff) also have "d * \ = poly ?p (of_nat k)" by (simp add: divide_inverse eval_pochhammer_poly poly_pcompose add_ac) also { from assms have "pCons 0 (pcompose (pochhammer_poly (n-1)) [:1,1::'a:]) = pochhammer_poly n" by (subst pochhammer_poly_Suc' [symmetric]) simp also from assms have "\ = pCons 0 (Poly [of_nat (stirling n (k+1)). k \ [0.. [0.. [0.. (of_nat k) * c ^ k = ?rhs $ k" by simp finally show "?lhs $ k = ?rhs $ k" . qed lemma inverse_irred_power_poly_code [code abstract]: "coeffs (inverse_irred_power_poly d n) = (if n = 0 \ d = 0 then [] else let e = d / (fact (n - 1)) in [e * of_nat x. x \ tl (stirling_row n)])" proof (cases "n = 0 \ d = 0") case False define e where "e = d / (fact (n - 1))" from False have "coeffs (inverse_irred_power_poly d n) = [e * of_nat (stirling n (k+1)). k \ [0.. = [e * of_nat x. x \ tl (stirling_row n)]" by (simp add: stirling_row_def map_tl [symmetric] o_def tl_upt map_Suc_upt [symmetric] del: upt_Suc) finally show ?thesis using False by (simp add: Let_def e_def) qed (auto simp: inverse_irred_power_poly_def) lemma solve_rat_fps_aux: - fixes p :: "'a :: {field_char_0,factorial_ring_gcd,normalization_euclidean_semiring} poly" and cs :: "('a \ nat) list" + fixes p :: "'a :: {field_char_0,field_gcd} poly" and cs :: "('a \ nat) list" assumes distinct: "distinct (map fst cs)" assumes azs: "(a, zs) = poly_pfd_simple p cs" assumes nz: "0 \ fst ` set cs" shows "fps_of_poly p / fps_of_poly (\(c,n)\cs. [:1,-c:]^Suc n) = Abs_fps (\k. coeff a k + (\ij\snd (cs ! i). (inverse_irred_power_poly (zs ! i ! j) (snd (cs ! i)+1 - j))) (of_nat k) * (fst (cs ! i)) ^ k))" (is "_ = ?rhs") proof - interpret pfd_homomorphism "fps_of_poly :: 'a poly \ 'a fps" by standard (auto simp: fps_of_poly_add fps_of_poly_mult) from distinct have distinct': "(a, b1) \ set cs \ (a, b2) \ set cs \ b1 = b2" for a b1 b2 by (metis (no_types, hide_lams) Some_eq_map_of_iff image_set in_set_zipE insert_iff list.simps(15) map_of_Cons_code(2) map_of_SomeD nz snd_conv) from nz have nz': "(0, b) \ set cs" for b by (auto simp add: image_iff) define n where "n = length cs" let ?g = "\(c, n). [:1, - c:] ^ Suc n" have "inj_on ?g (set cs)" proof fix x y assume "x \ set cs" "y \ set cs" "?g x = ?g y" moreover obtain c1 n1 c2 n2 where [simp]: "x = (c1, n1)" "y = (c2, n2)" by (cases x, cases y) ultimately have in_cs: "(c1, n1) \ set cs" "(c2, n2) \ set cs" and eq: "[:1, - c1:] ^ Suc n1 = [:1, - c2:] ^ Suc n2" by simp_all with nz have [simp]: "c1 \ 0" "c2 \ 0" by (auto simp add: image_iff) have "Suc n1 = degree ([:1, - c1:] ^ Suc n1)" by (simp add: degree_power_eq del: power_Suc) also have "\ = degree ([:1, - c2:] ^ Suc n2)" using eq by simp also have "\ = Suc n2" by (simp add: degree_power_eq del: power_Suc) finally have "n1 = n2" by simp then have "0 = poly ([:1, - c1:] ^ Suc n1) (1 / c1)" by simp also have "\ = poly ([:1, - c2:] ^ Suc n2) (1 / c1)" using eq by simp finally show "x = y" using \n1 = n2\ by (auto simp: field_simps) qed with distinct have distinct': "distinct (map ?g cs)" by (simp add: distinct_map del: power_Suc) from nz' distinct have coprime: "pairwise coprime (?g ` set cs)" by (auto intro!: pairwise_imageI coprime_linear_poly' simp add: eq_key_imp_eq_value simp del: power_Suc) have [simp]: "length zs = n" using assms by (simp add: poly_pfd_simple_def n_def split: if_split_asm) have [simp]: "i < length cs \ length (zs!i) = snd (cs!i)+1" for i using assms by (simp add: poly_pfd_simple_def Let_def case_prod_unfold split: if_split_asm) let ?f = "\(c, n). ([:1,-c:], n)" let ?cs' = "map ?f cs" have "fps_of_poly (fst (poly_pfd_simple p cs)) + (\ij\snd (?cs' ! i). from_decomp (map (map (\c. [:c:])) (snd (poly_pfd_simple p cs)) ! i ! j) (fst (?cs' ! i)) (snd (?cs' ! i)+1 - j)) = fps_of_poly p / fps_of_poly (\(x, n)\?cs'. x ^ Suc n)" (is "?A = ?B") using nz distinct' coprime by (intro partial_fraction_decomposition poly_pfd_simple) (force simp: o_def case_prod_unfold simp del: power_Suc)+ note this [symmetric] also from azs [symmetric] have "?A = fps_of_poly a + (\ij\snd (cs ! i). from_decomp (map (map (\c. [:c:])) zs ! i ! j) [:1,-fst (cs ! i):] (snd (cs ! i)+1 - j))" (is "_ = _ + ?S") by (simp add: case_prod_unfold Let_def n_def) also have "?S = (\ij\snd (cs ! i). fps_const (zs ! i ! j) / ((1 - fps_const (fst (cs!i))*fps_X) ^ (snd (cs!i)+1 - j)))" by (intro sum.cong refl) (auto simp: from_decomp_def map_nth n_def fps_of_poly_linear' fps_of_poly_simps fps_const_neg [symmetric] mult_ac simp del: fps_const_neg) also have "\ = (\ij\snd (cs ! i) . Abs_fps (\k. poly (inverse_irred_power_poly (zs ! i ! j) (snd (cs ! i)+1 - j)) (of_nat k) * (fst (cs ! i)) ^ k))" using nz by (intro sum.cong refl one_minus_const_fps_X_neg_power'') auto also have "fps_of_poly a + \ = ?rhs" by (intro fps_ext) (simp_all add: sum_distrib_right fps_sum_nth poly_sum) finally show ?thesis by (simp add: o_def case_prod_unfold) qed definition solve_factored_ratfps :: - "('a :: {field_char_0,factorial_ring_gcd,normalization_euclidean_semiring}) poly \ ('a \ nat) list \ 'a poly \ ('a poly \ 'a) list" where + "('a :: {field_char_0,field_gcd}) poly \ ('a \ nat) list \ 'a poly \ ('a poly \ 'a) list" where "solve_factored_ratfps p cs = (let n = length cs in case poly_pfd_simple p cs of (a, zs) \ (a, zip_with (\zs (c,n). ((\(z,j) \ zip zs [0.. length (snd (poly_pfd_simple p cs) ! i) = snd (cs!i) + 1" by (auto simp: poly_pfd_simple_def case_prod_unfold Let_def) lemma solve_factored_ratfps_roots: "map snd (snd (solve_factored_ratfps p cs)) = map fst cs" by (rule nth_equalityI) (simp_all add: solve_factored_ratfps_def poly_pfd_simple case_prod_unfold Let_def zip_with_altdef o_def) definition interp_ratfps_solution where "interp_ratfps_solution = (\(p,cs) n. coeff p n + (\(q,c)\cs. poly q (of_nat n) * c ^ n))" lemma solve_factored_ratfps: - fixes p :: "'a :: {field_char_0,factorial_ring_gcd,normalization_euclidean_semiring} poly" and cs :: "('a \ nat) list" + fixes p :: "'a :: {field_char_0,field_gcd} poly" and cs :: "('a \ nat) list" assumes distinct: "distinct (map fst cs)" assumes nz: "0 \ fst ` set cs" shows "fps_of_poly p / fps_of_poly (\(c,n)\cs. [:1,-c:]^Suc n) = Abs_fps (interp_ratfps_solution (solve_factored_ratfps p cs))" (is "?lhs = ?rhs") proof - obtain a zs where azs: "(a, zs) = solve_factored_ratfps p cs" using prod.exhaust by metis from azs have a: "a = fst (poly_pfd_simple p cs)" by (simp add: solve_factored_ratfps_def Let_def case_prod_unfold) define zs' where "zs' = snd (poly_pfd_simple p cs)" with a have azs': "(a, zs') = poly_pfd_simple p cs" by simp from azs have zs: "zs = snd (solve_factored_ratfps p cs)" by (auto simp add: snd_def split: prod.split) have "?lhs = Abs_fps (\k. coeff a k + (\ij\snd (cs ! i). inverse_irred_power_poly (zs' ! i ! j) (snd (cs ! i)+1 - j)) (of_nat k) * (fst (cs ! i)) ^ k))" by (rule solve_rat_fps_aux[OF distinct azs' nz]) also from azs have "\ = ?rhs" unfolding interp_ratfps_solution_def by (auto simp: a zs solve_factored_ratfps_def Let_def case_prod_unfold zip_altdef zip_with_altdef' sum_list_sum_nth atLeast0LessThan zs'_def lessThan_Suc_atMost intro!: fps_ext sum.cong simp del: upt_Suc) finally show ?thesis . qed definition solve_factored_ratfps' where "solve_factored_ratfps' = (\p (a,cs). solve_factored_ratfps (smult (inverse a) p) cs)" lemma solve_factored_ratfps': assumes "is_alt_factorization_of fctrs q" "q \ 0" shows "Abs_fps (interp_ratfps_solution (solve_factored_ratfps' p fctrs)) = fps_of_poly p / fps_of_poly q" proof - from assms have q: "q = interp_alt_factorization fctrs" by (simp add: is_alt_factorization_of_def) from assms(2) have nz: "fst fctrs \ 0" by (subst (asm) q) (auto simp: interp_alt_factorization_def case_prod_unfold) note q also from nz have "coeff (interp_alt_factorization fctrs) 0 \ 0" by (auto simp: interp_alt_factorization_def case_prod_unfold coeff_0_prod_list o_def coeff_0_power prod_list_zero_iff) finally have "coeff q 0 \ 0" . obtain a cs where fctrs: "fctrs = (a, cs)" by (cases fctrs) simp_all obtain b zs where sol: "solve_factored_ratfps' p fctrs = (b, zs)" using prod.exhaust by metis from assms have [simp]: "a \ 0" by (auto simp: is_alt_factorization_of_def interp_alt_factorization_def fctrs) have "fps_of_poly p / fps_of_poly (smult a (\(c, n)\cs. [:1, - c:] ^ Suc n)) = fps_of_poly p / (fps_const a * fps_of_poly (\(c, n)\cs. [:1, - c:] ^ Suc n))" by (simp_all add: fps_of_poly_smult case_prod_unfold del: power_Suc) also have "\ = fps_of_poly p / fps_const a / fps_of_poly (\(c, n)\cs. [:1, - c:] ^ Suc n)" by (subst is_unit_div_mult2_eq) (auto simp: coeff_0_power coeff_0_prod_list prod_list_zero_iff) also have "fps_of_poly p / fps_const a = fps_of_poly (smult (inverse a) p)" by (simp add: fps_const_inverse fps_divide_unit) also from assms have "smult a (\(c, n)\cs. [:1, - c:] ^ Suc n) = q" by (simp add: is_alt_factorization_of_def interp_alt_factorization_def fctrs del: power_Suc) also have "fps_of_poly (smult (inverse a) p) / fps_of_poly (\(c, n)\cs. [:1, - c:] ^ Suc n) = Abs_fps (interp_ratfps_solution (solve_factored_ratfps (smult (inverse a) p) cs))" (is "?lhs = _") using assms by (intro solve_factored_ratfps) (simp_all add: is_alt_factorization_of_def fctrs solve_factored_ratfps'_def) also have "\ = Abs_fps (interp_ratfps_solution (solve_factored_ratfps' p fctrs))" by (simp add: solve_factored_ratfps'_def fctrs) finally show ?thesis .. qed lemma degree_Poly_eq: assumes "xs = [] \ last xs \ 0" shows "degree (Poly xs) = length xs - 1" proof - from assms consider "xs = []" | "xs \ []" "last xs \ 0" by blast thus ?thesis proof cases assume "last xs \ 0" "xs \ []" hence "no_trailing ((=) 0) xs" by (auto simp: no_trailing_unfold) thus ?thesis by (simp add: degree_eq_length_coeffs) qed auto qed lemma degree_Poly': "degree (Poly xs) \ length xs - 1" using length_strip_while_le[of "(=) 0" xs] by (simp add: degree_eq_length_coeffs) lemma degree_inverse_irred_power_poly_le: "degree (inverse_irred_power_poly c n) \ n - 1" by (auto simp: inverse_irred_power_poly_def intro: order.trans[OF degree_Poly']) lemma degree_inverse_irred_power_poly: assumes "c \ 0" shows "degree (inverse_irred_power_poly c n) = n - 1" unfolding inverse_irred_power_poly_def using assms by (subst degree_Poly_eq) (auto simp: last_conv_nth) lemma reflect_poly_0_iff [simp]: "reflect_poly p = 0 \ p = 0" using coeff_0_reflect_poly_0_iff[of p] by fastforce lemma degree_sum_list_le: "(\p. p \ set ps \ degree p \ T) \ degree (sum_list ps) \ T" by (induction ps) (auto intro: degree_add_le) theorem ratfps_closed_form_exists: fixes q :: "complex poly" assumes nz: "coeff q 0 \ 0" defines "q' \ reflect_poly q" obtains r rs where "\n. fps_nth (fps_of_poly p / fps_of_poly q) n = coeff r n + (\c | poly q' c = 0. poly (rs c) (of_nat n) * c ^ n)" and "\z. poly q' z = 0 \ degree (rs z) \ order z q' - 1" proof - from assms have nz': "q \ 0" by auto from complex_alt_factorization_exists [OF nz] obtain fctrs where fctrs: "is_alt_factorization_of fctrs q" .. with nz have fctrs': "is_factorization_of fctrs q'" unfolding q'_def by (rule reflect_factorization') define r where "r = fst (solve_factored_ratfps' p fctrs)" define ts where "ts = snd (solve_factored_ratfps' p fctrs)" define rs where "rs = the \ map_of (map (\(x,y). (y,x)) ts)" from nz' have "q' \ 0" by (simp add: q'_def) hence roots: "{z. poly q' z = 0} = set (map fst (snd fctrs))" using is_factorization_of_roots [of "fst fctrs" "snd fctrs" q'] fctrs' by simp have rs: "rs c = r" if "(r, c) \ set ts" for c r proof - have "map_of (map (\(x,y). (y, x)) (snd (solve_factored_ratfps' p fctrs))) c = Some r" using that fctrs by (intro map_of_is_SomeI) (force simp: o_def case_prod_unfold solve_factored_ratfps'_def ts_def solve_factored_ratfps_roots is_alt_factorization_of_def)+ thus ?thesis by (simp add: rs_def ts_def) qed have [simp]: "length ts = length (snd fctrs)" by (auto simp: ts_def solve_factored_ratfps'_def case_prod_unfold solve_factored_ratfps_def) { fix n :: nat have "fps_of_poly p / fps_of_poly q = Abs_fps (interp_ratfps_solution (solve_factored_ratfps' p fctrs))" using solve_factored_ratfps' [OF fctrs nz'] .. also have "fps_nth \ n = interp_ratfps_solution (solve_factored_ratfps' p fctrs) n" by simp also have "\ = coeff r n + (\p\snd (solve_factored_ratfps' p fctrs). poly (fst p) (of_nat n) * snd p ^ n)" (is "_ = _ + ?A") unfolding interp_ratfps_solution_def case_prod_unfold r_def by simp also have "?A = (\p\ts. poly (rs (snd p)) (of_nat n) * snd p ^ n)" by (intro arg_cong[OF map_cong] refl) (auto simp: rs ts_def) also have "\ = (\c\map snd ts. poly (rs c) (of_nat n) * c ^ n)" by (simp add: o_def) also have "map snd ts = map fst (snd fctrs)" unfolding solve_factored_ratfps'_def case_prod_unfold ts_def by (rule solve_factored_ratfps_roots) also have "(\c\\. poly (rs c) (of_nat n) * c ^ n) = (\c | poly q' c = 0. poly (rs c) (of_nat n) * c ^ n)" unfolding roots using fctrs by (intro sum_list_distinct_conv_sum_set) (auto simp: is_alt_factorization_of_def) finally have "fps_nth (fps_of_poly p / fps_of_poly q) n = coeff r n + (\c\{z. poly q' z = 0}. poly (rs c) (of_nat n) * c ^ n)" . } moreover { fix z assume "poly q' z = 0" hence "z \ set (map fst (snd fctrs))" using roots by blast then obtain i where i: "i < length (snd fctrs)" and [simp]: "z = fst (snd fctrs ! i)" by (auto simp: set_conv_nth) from i have "(fst (ts ! i), snd (ts ! i)) \ set ts" by (auto simp: set_conv_nth) also from i have "snd (ts ! i) = z" by (simp add: ts_def solve_factored_ratfps'_def case_prod_unfold solve_factored_ratfps_def) finally have "rs z = fst (ts ! i)" by (intro rs) auto also have "\ = (\p\zip (snd (poly_pfd_simple (smult (inverse (fst fctrs)) p) (snd fctrs)) ! i) [0.. \ snd (snd fctrs ! i)" by (intro degree_sum_list_le) (auto intro!: order.trans [OF degree_inverse_irred_power_poly_le]) also have "order z q' = Suc \" using nz' fctrs' i by (intro is_factorization_of_order[of q' "fst fctrs" "snd fctrs"]) (auto simp: q'_def) hence "snd (snd fctrs ! i) = order z q' - 1" by simp finally have "degree (rs z) \ \" . } ultimately show ?thesis using that[of r rs] by blast qed end diff --git a/thys/Linear_Recurrences/Solver/Show_RatFPS.thy b/thys/Linear_Recurrences/Solver/Show_RatFPS.thy --- a/thys/Linear_Recurrences/Solver/Show_RatFPS.thy +++ b/thys/Linear_Recurrences/Solver/Show_RatFPS.thy @@ -1,36 +1,36 @@ (* File: Show_RatFPS.thy Author: Manuel Eberl, TU MĂ¼nchen *) section \Pretty-printing for rational formal power series\ theory Show_RatFPS imports Linear_Recurrences.RatFPS Show.Show_Poly begin definition ratfps_parens where "ratfps_parens s = (if list_ex (\c. c \ set ''+-'') (tl s) then ''('' @ s @ '')'' else s)" definition "show_ratfps x = (case quot_of_ratfps x of (p, q) \ if q = 1 then show p else ratfps_parens (show p) @ '' / '' @ ratfps_parens (show q))" -definition showsp_ratfps :: "'a :: {field, euclidean_ring_gcd,show} ratfps showsp" +definition showsp_ratfps :: "'a :: {field_gcd,show} ratfps showsp" where "showsp_ratfps p x y = (show_ratfps x @ y)" -instantiation ratfps:: ("{show,field,euclidean_ring_gcd}") "show" +instantiation ratfps:: ("{show,field_gcd}") "show" begin definition "shows_prec p (x :: 'a ratfps) = showsp_ratfps p x" definition "shows_list (ps :: 'a ratfps list) = showsp_list shows_prec 0 ps" lemma show_law_ratfps [show_law_simps]: "shows_prec p (a :: 'a ratfps) (r @ s) = shows_prec p a r @ s" by (simp add: shows_prec_ratfps_def showsp_ratfps_def show_law_simps) instance by standard (auto simp: shows_list_ratfps_def show_law_simps) end end diff --git a/thys/Mason_Stothers/Mason_Stothers.thy b/thys/Mason_Stothers/Mason_Stothers.thy --- a/thys/Mason_Stothers/Mason_Stothers.thy +++ b/thys/Mason_Stothers/Mason_Stothers.thy @@ -1,384 +1,384 @@ (* File: Mason_Stothers.thy Author: Manuel Eberl, TU MĂ¼nchen A simple proof of the Mason--Stothers theorem, the analogue of the abc conjecture for polynomials. *) section \The Mason--Stother's Theorem\ theory Mason_Stothers imports "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin subsection \Auxiliary material\ (* TODO Rename this to fps_radical or whatever *) hide_const (open) Formal_Power_Series.radical lemma degree_div: assumes "a dvd b" shows "degree (b div a) = degree b - degree a" using assms by (cases "a = 0"; cases "b = 0") (auto elim!: dvdE simp: degree_mult_eq) lemma degree_pderiv_le: shows "degree (pderiv p) \ degree p - 1" by (rule degree_le, cases "degree p = 0") (auto simp: coeff_pderiv coeff_eq_0) lemma degree_pderiv_less: assumes "pderiv p \ 0" shows "degree (pderiv p) < degree p" proof - have "degree (pderiv p) \ degree p - 1" by (rule degree_pderiv_le) also have "degree p \ 0" using assms by (auto intro!: Nat.gr0I elim!: degree_eq_zeroE) hence "degree p - 1 < degree p" by simp finally show ?thesis . qed lemma pderiv_eq_0: assumes "degree p = 0" shows "pderiv p = 0" using assms by (auto elim!: degree_eq_zeroE) subsection \Definition of a radical\ text \ The following definition of a radical is generic for any factorial semiring. \ context factorial_semiring begin definition radical :: "'a \ 'a" where "radical x = (if x = 0 then 0 else \(prime_factors x))" lemma radical_0 [simp]: "radical 0 = 0" by (simp add: radical_def) lemma radical_nonzero: "x \ 0 \ radical x = \(prime_factors x)" by (simp add: radical_def) lemma radical_eq_0_iff [simp]: "radical x = 0 \ x = 0" by (auto simp: radical_def) lemma prime_factorization_radical [simp]: assumes "x \ 0" shows "prime_factorization (radical x) = mset_set (prime_factors x)" proof - have "prime_factorization (radical x) = (\p\prime_factors x. prime_factorization p)" unfolding radical_def using assms by (auto intro!: prime_factorization_prod) also have "\ = (\p\prime_factors x. {#p#})" by (intro Groups_Big.sum.cong) (auto intro!: prime_factorization_prime) also have "\ = mset_set (prime_factors x)" by simp finally show ?thesis . qed lemma prime_factors_radical [simp]: "x \ 0 \ prime_factors (radical x) = prime_factors x" by simp lemma radical_dvd [simp, intro]: "radical x dvd x" by (cases "x = 0") (force intro: prime_factorization_subset_imp_dvd mset_set_set_mset_msubset)+ lemma multiplicity_radical_prime: assumes "prime p" "x \ 0" shows "multiplicity p (radical x) = (if p dvd x then 1 else 0)" proof - have "multiplicity p (radical x) = (\q\prime_factors x. multiplicity p q)" using assms unfolding radical_def by (auto simp: prime_elem_multiplicity_prod_distrib) also have "\ = (\q\prime_factors x. if p = q then 1 else 0)" using assms by (intro Groups_Big.sum.cong) (auto intro!: prime_multiplicity_other) also have "\ = (if p \ prime_factors x then 1 else 0)" by simp also have "\ = (if p dvd x then 1 else 0)" using assms by (auto simp: prime_factors_dvd) finally show ?thesis . qed lemma radical_1 [simp]: "radical 1 = 1" by (simp add: radical_def) lemma radical_unit [simp]: "is_unit x \ radical x = 1" by (auto simp: radical_def prime_factorization_unit) lemma prime_factors_power: assumes "n > 0" shows "prime_factors (x ^ n) = prime_factors x" using assms by (cases "x = 0") (auto simp: prime_factors_dvd zero_power prime_dvd_power_iff) lemma radical_power [simp]: "n > 0 \ radical (x ^ n) = radical x" by (auto simp add: radical_def prime_factors_power) end context factorial_semiring_gcd begin lemma radical_mult_coprime: assumes "coprime a b" shows "radical (a * b) = radical a * radical b" proof (cases "a = 0 \ b = 0") case False with assms have "prime_factors a \ prime_factors b = {}" using not_prime_unit coprime_common_divisor by (auto simp: prime_factors_dvd) hence "\(prime_factors a \ prime_factors b) = \(prime_factors a) * \(prime_factors b)" by (intro prod.union_disjoint) auto with False show ?thesis by (simp add: radical_def prime_factorization_mult) qed auto lemma multiplicity_le_imp_dvd': assumes "x \ 0" "\p. p \ prime_factors x \ multiplicity p x \ multiplicity p y" shows "x dvd y" proof (rule multiplicity_le_imp_dvd) fix p assume "prime p" thus "multiplicity p x \ multiplicity p y" using assms(1) assms(2)[of p] by (cases "p dvd x") (auto simp: prime_factors_dvd not_dvd_imp_multiplicity_0) qed fact+ end subsection \Main result\ text \ The following proofs are basically a one-to-one translation of Franz Lemmermeyer's presentation~\cite{lemmermeyer05} of Snyder's proof of the Mason--Stothers theorem. \ lemma prime_power_dvd_pderiv: - fixes f p :: "'a :: {factorial_ring_gcd, field} poly" + fixes f p :: "'a :: field_gcd poly" assumes "prime_elem p" defines "n \ multiplicity p f - 1" shows "p ^ n dvd pderiv f" proof (cases "p dvd f \ f \ 0") case True hence "multiplicity p f > 0" using assms by (subst prime_multiplicity_gt_zero_iff) auto hence Suc_n: "Suc n = multiplicity p f" by (simp add: n_def) define g where "g = f div p ^ Suc n" have "p ^ Suc n dvd f" unfolding Suc_n by (rule multiplicity_dvd) hence f_eq: "f = p ^ Suc n * g" by (simp add: g_def) also have "pderiv \ = p ^ n * (smult (of_nat (Suc n)) (pderiv p * g) + p * pderiv g)" by (simp only: pderiv_mult pderiv_power_Suc) (simp add: algebra_simps) also have "p ^ n dvd \" by simp finally show ?thesis . qed (auto simp: n_def not_dvd_imp_multiplicity_0) lemma poly_div_radical_dvd_pderiv: - fixes p :: "'a :: {factorial_ring_gcd, field} poly" + fixes p :: "'a :: field_gcd poly" shows "p div radical p dvd pderiv p" proof (cases "pderiv p = 0") case False hence "p \ 0" by auto show ?thesis proof (rule multiplicity_le_imp_dvd') fix q :: "'a poly" assume q: "q \ prime_factors (p div radical p)" hence "q dvd p div radical p" by auto also from \p \ 0\ have "\ dvd p" by (subst div_dvd_iff_mult) auto finally have "q dvd p" . have "p = p div radical p * radical p" by simp also from q and \p \ 0\ have "multiplicity q \ = Suc (multiplicity q (p div radical p))" by (subst prime_elem_multiplicity_mult_distrib) (auto simp: dvd_div_eq_0_iff multiplicity_radical_prime \q dvd p\ prime_factors_dvd) finally have "multiplicity q (p div radical p) \ multiplicity q p - 1" by simp also have "\ \ multiplicity q (pderiv p)" using \pderiv p \ 0\ and q and \p \ 0\ by (intro multiplicity_geI prime_power_dvd_pderiv) (auto simp: prime_factors_dvd dvd_div_eq_0_iff) finally show "multiplicity q (p div radical p) \ multiplicity q (pderiv p)" . qed (insert \p \ 0\, auto simp: dvd_div_eq_0_iff) qed auto lemma degree_pderiv_mult_less: assumes "pderiv C \ 0" shows "degree (pderiv C * B) < degree B + degree C" proof - have "degree (pderiv C * B) \ degree (pderiv C) + degree B" by (rule degree_mult_le) also from assms have "degree (pderiv C) < degree C" by (rule degree_pderiv_less) finally show ?thesis by simp qed lemma Mason_Stothers_aux: - fixes A B C :: "'a :: {factorial_ring_gcd, field} poly" + fixes A B C :: "'a :: field_gcd poly" assumes nz: "A \ 0" "B \ 0" "C \ 0" and sum: "A + B + C = 0" and coprime: "Gcd {A, B, C} = 1" and deg_ge: "degree A \ degree (radical (A * B * C))" shows "pderiv A = 0" "pderiv B = 0" "pderiv C = 0" proof - have C_eq: "C = -A - B" "-C = A + B" using sum by algebra+ from coprime have "gcd A (gcd B (-C)) = 1" by simp also note C_eq(2) finally have "coprime A B" by (simp add: gcd.commute add.commute[of A B] coprime_iff_gcd_eq_1) hence "coprime A (-C)" "coprime B (-C)" unfolding C_eq by (simp_all add: gcd.commute[of B A] gcd.commute[of B "A + B"] add.commute coprime_iff_gcd_eq_1) hence "coprime A C" "coprime B C" by simp_all note coprime = coprime \coprime A B\ this have coprime1: "coprime (A div radical A) (B div radical B)" by (rule coprime_divisors[OF _ _ \coprime A B\]) (insert nz, auto simp: div_dvd_iff_mult) have coprime2: "coprime (A div radical A) (C div radical C)" by (rule coprime_divisors[OF _ _ \coprime A C\]) (insert nz, auto simp: div_dvd_iff_mult) have coprime3: "coprime (B div radical B) (C div radical C)" by (rule coprime_divisors[OF _ _ \coprime B C\]) (insert nz, auto simp: div_dvd_iff_mult) have coprime4: "coprime (A div radical A * (B div radical B)) (C div radical C)" using coprime2 coprime3 by (subst coprime_mult_left_iff) auto have eq: "A * pderiv B - pderiv A * B = pderiv C * B - C * pderiv B" by (simp add: C_eq pderiv_add pderiv_diff pderiv_minus algebra_simps) have "A div radical A dvd (A * pderiv B - pderiv A * B)" using nz by (intro dvd_diff dvd_mult2 poly_div_radical_dvd_pderiv) (auto simp: div_dvd_iff_mult) with eq have "A div radical A dvd (pderiv C * B - C * pderiv B)" by simp moreover have "C div radical C dvd (pderiv C * B - C * pderiv B)" using nz by (intro dvd_diff dvd_mult2 poly_div_radical_dvd_pderiv) (auto simp: div_dvd_iff_mult) moreover have "B div radical B dvd (pderiv C * B - C * pderiv B)" using nz by (intro dvd_diff dvd_mult poly_div_radical_dvd_pderiv) (auto simp: div_dvd_iff_mult) ultimately have "(A div radical A) * (B div radical B) * (C div radical C) dvd (pderiv C * B - C * pderiv B)" using coprime coprime1 coprime4 by (intro divides_mult) auto also have "(A div radical A) * (B div radical B) * (C div radical C) = (A * B * C) div (radical A * radical B * radical C)" by (simp add: div_mult_div_if_dvd mult_dvd_mono) also have "radical A * radical B * radical C = radical (A * B) * radical C" using coprime by (subst radical_mult_coprime) auto also have "\ = radical (A * B * C)" using coprime by (subst radical_mult_coprime [symmetric]) auto finally have dvd: "((A * B * C) div radical (A * B * C)) dvd (pderiv C * B - C * pderiv B)" . have "pderiv B = 0 \ pderiv C = 0" proof (rule ccontr) assume "\(pderiv B = 0 \ pderiv C = 0)" hence *: "pderiv B \ 0 \ pderiv C \ 0" by blast have "degree (pderiv C * B - C * pderiv B) \ max (degree (pderiv C * B)) (degree (C * pderiv B))" by (rule degree_diff_le_max) also have "\ < degree B + degree C" using degree_pderiv_mult_less[of B C] degree_pderiv_mult_less[of C B] * by (cases "pderiv B = 0"; cases "pderiv C = 0") (auto simp add: algebra_simps) also have "degree B + degree C = degree (B * C)" using nz by (subst degree_mult_eq) auto also have "\ = degree (A * (B * C)) - degree A" using nz by (subst (2) degree_mult_eq) auto also have "\ \ degree (A * B * C) - degree (radical (A * B * C))" unfolding mult.assoc using assms by (intro diff_le_mono2) (auto simp: mult_ac) also have "\ = degree ((A * B * C) div radical (A * B * C))" by (intro degree_div [symmetric]) auto finally have less: "degree (pderiv C * B - C * pderiv B) < degree (A * B * C div radical (A * B * C))" by simp have eq': "pderiv C * B - C * pderiv B = 0" proof (rule ccontr) assume "pderiv C * B - C * pderiv B \ 0" hence "degree (A * B * C div radical (A * B * C)) \ degree (pderiv C * B - C * pderiv B)" using dvd by (intro dvd_imp_degree_le) auto with less show False by linarith qed from * show False proof (elim disjE) assume [simp]: "pderiv C \ 0" have "C dvd C * pderiv B" by simp also from eq' have "\ = pderiv C * B" by simp finally have "C dvd pderiv C" using coprime by (subst (asm) coprime_dvd_mult_left_iff) (auto simp: coprime_commute) hence "degree C \ degree (pderiv C)" by (intro dvd_imp_degree_le) auto moreover have "degree (pderiv C) < degree C" by (intro degree_pderiv_less) auto ultimately show False by simp next assume [simp]: "pderiv B \ 0" have "B dvd B * pderiv C" by simp also from eq' have "\ = pderiv B * C" by (simp add: mult_ac) finally have "B dvd pderiv B" using coprime by (subst (asm) coprime_dvd_mult_left_iff) auto hence "degree B \ degree (pderiv B)" by (intro dvd_imp_degree_le) auto moreover have "degree (pderiv B) < degree B" by (intro degree_pderiv_less) auto ultimately show False by simp qed qed with eq and nz show "pderiv A = 0" "pderiv B = 0" "pderiv C = 0" by auto qed theorem Mason_Stothers: - fixes A B C :: "'a :: {factorial_ring_gcd, field} poly" + fixes A B C :: "'a :: field_gcd poly" assumes nz: "A \ 0" "B \ 0" "C \ 0" "\p\{A,B,C}. pderiv p \ 0" and sum: "A + B + C = 0" and coprime: "Gcd {A, B, C} = 1" shows "Max {degree A, degree B, degree C} < degree (radical (A * B * C))" proof - have "degree A < degree (radical (A * B * C))" if "\p\{A,B,C}. p \ 0" "\p\{A,B,C}. pderiv p \ 0" "sum_mset {#A,B,C#} = 0" "Gcd {A, B, C} = 1" for A B C :: "'a poly" proof (rule ccontr) assume "\(degree A < degree (radical (A * B * C)))" hence "degree A \ degree (radical (A * B * C))" by simp with Mason_Stothers_aux[of A B C] that show False by (auto simp: add_ac) qed from this[of A B C] this[of B C A] this[of C A B] assms show ?thesis by (simp only: insert_commute mult_ac add_ac) (auto simp: add_ac mult_ac) qed text \ The result can be simplified a bit more in fields of characteristic 0: \ corollary Mason_Stothers_char_0: - fixes A B C :: "'a :: {factorial_ring_gcd, field_char_0} poly" + fixes A B C :: "'a :: {field_gcd, field_char_0} poly" assumes nz: "A \ 0" "B \ 0" "C \ 0" and deg: "\p\{A,B,C}. degree p \ 0" and sum: "A + B + C = 0" and coprime: "Gcd {A, B, C} = 1" shows "Max {degree A, degree B, degree C} < degree (radical (A * B * C))" proof - from deg have "\p\{A,B,C}. pderiv p \ 0" by (auto simp: pderiv_eq_0_iff) from Mason_Stothers[OF assms(1-3) this assms(5-)] show ?thesis . qed text \ As a nice corollary, we get a kind of analogue of Fermat's last theorem for polynomials: Given non-zero polynomials $A$, $B$, $C$ with $A ^ n + B ^ n + C ^ n = 0$ on lowest terms, we must either have $n \leq 2$ or $(A ^ n)' = (B ^ n)' = (C ^ n)' = 0$. In the case of a field with characteristic 0, this last possibility is equivalent to $A$, $B$, and $C$ all being constant. \ corollary fermat_poly: - fixes A B C :: "'a :: {factorial_ring_gcd,field} poly" + fixes A B C :: "'a :: field_gcd poly" assumes sum: "A ^ n + B ^ n + C ^ n = 0" and cop: "Gcd {A, B, C} = 1" assumes nz: "A \ 0" "B \ 0" "C \ 0" and deg: "\p\{A,B,C}. pderiv (p ^ n) \ 0" shows "n \ 2" proof (rule ccontr) assume "\(n \ 2)" hence "n > 2" by simp have "Max {degree (A ^ n), degree (B ^ n), degree (C ^ n)} < degree (radical (A ^ n * B ^ n * C ^ n))" (is "_ < ?d") using assms by (intro Mason_Stothers) (auto simp: degree_power_eq gcd_exp) hence "Max {degree (A ^ n), degree (B ^ n), degree (C ^ n)} + 1 \ ?d" by linarith hence "n * degree A + 1 \ ?d" "n * degree B + 1 \ ?d" "n * degree C + 1 \ ?d" using assms by (simp_all add: degree_power_eq) hence "n * (degree A + degree B + degree C) + 3 \ 3 * ?d" unfolding ring_distribs by linarith also have "A ^ n * B ^ n * C ^ n = (A * B * C) ^ n" by (simp add: mult_ac power_mult_distrib) also have "radical \ = radical (A * B * C)" using \n > 2\ by simp also have "degree (radical (A * B * C)) \ degree (A * B * C)" using nz by (intro dvd_imp_degree_le) auto also have "\ = degree A + degree B + degree C" using nz by (simp add: degree_mult_eq) finally have "(3 - n) * (degree A + degree B + degree C) \ 3" by (simp add: algebra_simps) hence "3 - n \ 0" by (intro notI) auto hence "n < 3" by simp with \n > 2\ show False by simp qed corollary fermat_poly_char_0: - fixes A B C :: "'a :: {factorial_ring_gcd,field_char_0} poly" + fixes A B C :: "'a :: {field_gcd,field_char_0} poly" assumes sum: "A ^ n + B ^ n + C ^ n = 0" and cop: "Gcd {A, B, C} = 1" assumes nz: "A \ 0" "B \ 0" "C \ 0" and deg: "\p\{A,B,C}. degree p > 0" shows "n \ 2" proof (rule ccontr) assume *: "\(n \ 2)" with nz and deg have "\p\{A,B,C}. pderiv (p ^ n) \ 0" by (auto simp: pderiv_eq_0_iff degree_power_eq) from fermat_poly[OF assms(1-5) this] and * show False by simp qed end \ No newline at end of file diff --git a/thys/Polynomial_Factorization/Missing_Polynomial_Factorial.thy b/thys/Polynomial_Factorization/Missing_Polynomial_Factorial.thy --- a/thys/Polynomial_Factorization/Missing_Polynomial_Factorial.thy +++ b/thys/Polynomial_Factorization/Missing_Polynomial_Factorial.thy @@ -1,329 +1,334 @@ section \More on Polynomials\ text \This theory contains several results on content, gcd, primitive part, etc.. Moreover, there is a slightly improved code-equation for computing the gcd.\ theory Missing_Polynomial_Factorial imports "HOL-Computational_Algebra.Polynomial_Factorial" Polynomial_Interpolation.Missing_Polynomial begin text \Improved code equation for @{const gcd_poly_code} which avoids computing the content twice.\ lemma gcd_poly_code_code[code]: "gcd_poly_code p q = (if p = 0 then normalize q else if q = 0 then normalize p else let c1 = content p; c2 = content q; p' = map_poly (\ x. x div c1) p; q' = map_poly (\ x. x div c2) q in smult (gcd c1 c2) (gcd_poly_code_aux p' q'))" unfolding gcd_poly_code_def Let_def primitive_part_def by simp -lemma gcd_smult: fixes f g :: "'a :: factorial_ring_gcd poly" +lemma gcd_smult: fixes f g :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" defines cf: "cf \ content f" and cg: "cg \ content g" shows "gcd (smult a f) g = (if a = 0 \ f = 0 then normalize g else smult (gcd a (cg div (gcd cf cg))) (gcd f g))" proof (cases "a = 0 \ f = 0") case False let ?c = "content" let ?pp = primitive_part let ?ua = "unit_factor a" let ?na = "normalize a" define H where "H = gcd (?c f) (?c g)" have "H dvd ?c f" unfolding H_def by auto then obtain F where fh: "?c f = H * F" unfolding dvd_def by blast from False have cf0: "?c f \ 0" by auto hence H: "H \ 0" unfolding H_def by auto from arg_cong[OF fh, of "\ f. f div H"] H have F: "F = ?c f div H" by auto have "H dvd ?c g" unfolding H_def by auto then obtain G where gh: "?c g = H * G" unfolding dvd_def by blast from arg_cong[OF gh, of "\ f. f div H"] H have G: "G = ?c g div H" by auto have "coprime F G" using H unfolding F G H_def using cf0 div_gcd_coprime by blast have "is_unit ?ua" using False by simp then have ua: "is_unit [: ?ua :]" by (simp add: is_unit_const_poly_iff) have "gcd (smult a f) g = smult (gcd (?na * ?c f) (?c g)) (gcd (smult ?ua (?pp f)) (?pp g))" unfolding gcd_poly_decompose[of "smult a f"] content_smult primitive_part_smult by simp also have "smult ?ua (?pp f) = ?pp f * [: ?ua :]" by simp also have "gcd \ (?pp g) = gcd (?pp f) (?pp g)" unfolding gcd_mult_unit1[OF ua] .. also have "gcd (?na * ?c f) (?c g) = gcd ((?na * F) * H) (G * H)" unfolding fh gh by (simp add: ac_simps) - also have "\ = gcd (?na * F) G * normalize H" unfolding gcd_mult_right gcd.commute[of G] by simp + also have "\ = gcd (?na * F) G * normalize H" unfolding gcd_mult_right gcd.commute[of G] + by (simp add: normalize_mult) also have "normalize H = H" by (metis H_def normalize_gcd) finally have "gcd (smult a f) g = smult (gcd (?na * F) G) (smult H (gcd (?pp f) (?pp g)))" by simp also have "smult H (gcd (?pp f) (?pp g)) = gcd f g" unfolding H_def by (rule gcd_poly_decompose[symmetric]) also have "gcd (?na * F) G = gcd (F * ?na) G" by (simp add: ac_simps) also have "\ = gcd ?na G" using \coprime F G\ by (simp add: gcd_mult_right_left_cancel ac_simps) finally show ?thesis unfolding G H_def cg cf using False by simp next case True hence "gcd (smult a f) g = normalize g" by (cases "a = 0", auto) thus ?thesis using True by simp qed lemma gcd_smult_ex: assumes "a \ 0" shows "\ b. gcd (smult a f) g = smult b (gcd f g) \ b \ 0" proof (cases "f = 0") case True thus ?thesis by (intro exI[of _ 1], auto) next case False hence id: "(a = 0 \ f = 0) = False" using assms by auto show ?thesis unfolding gcd_smult id if_False by (intro exI conjI, rule refl, insert assms, auto) qed -lemma primitive_part_idemp[simp]: "primitive_part (primitive_part f) = primitive_part f" - by (metis content_primitive_part primitive_part_eq_0_iff primitive_part_prim) +lemma primitive_part_idemp[simp]: + fixes f :: "'a :: {semiring_gcd,normalization_semidom_multiplicative} poly" + shows "primitive_part (primitive_part f) = primitive_part f" + by (metis content_primitive_part[of f] primitive_part_eq_0_iff primitive_part_prim) lemma content_gcd_primitive: "f \ 0 \ content (gcd (primitive_part f) g) = 1" "f \ 0 \ content (gcd (primitive_part f) (primitive_part g)) = 1" by (metis (no_types, lifting) content_dvd_contentI content_primitive_part gcd_dvd1 is_unit_content_iff)+ lemma content_gcd_content: "content (gcd f g) = gcd (content f) (content g)" (is "?l = ?r") proof - let ?c = "content" have "?l = normalize (gcd (?c f) (?c g)) * ?c (gcd (primitive_part f) (primitive_part g))" unfolding gcd_poly_decompose[of f g] content_smult .. also have "\ = gcd (?c f) (?c g) * ?c (gcd (primitive_part f) (primitive_part g))" by simp also have "\ = ?r" using content_gcd_primitive[of f g] by (metis (no_types, lifting) content_dvd_contentI content_eq_zero_iff content_primitive_part gcd_dvd2 gcd_eq_0_iff is_unit_content_iff mult_cancel_left1) finally show ?thesis . qed lemma gcd_primitive_part: "gcd (primitive_part f) (primitive_part g) = normalize (primitive_part (gcd f g))" proof(cases "f = 0") case True show ?thesis unfolding gcd_poly_decompose[of f g] gcd_0_left primitive_part_0 True by (simp add: associatedI primitive_part_dvd_primitive_partI) next case False have "normalize 1 = normalize (unit_factor (gcd (content f) (content g)))" by (simp add: False) then show ?thesis unfolding gcd_poly_decompose[of f g] by (metis (no_types) Polynomial.normalize_smult content_gcd_primitive(1)[OF False] content_times_primitive_part normalize_gcd primitive_part_smult) qed lemma primitive_part_gcd: "primitive_part (gcd f g) = unit_factor (gcd f g) * gcd (primitive_part f) (primitive_part g)" unfolding gcd_primitive_part by (metis (no_types, lifting) content_times_primitive_part gcd.normalize_idem mult_cancel_left2 mult_smult_left normalize_eq_0_iff normalize_mult_unit_factor primitive_part_eq_0_iff smult_content_normalize_primitive_part unit_factor_mult_normalize) -lemma primitive_part_normalize: "primitive_part (normalize f) = normalize (primitive_part f)" +lemma primitive_part_normalize: + fixes f :: "'a :: {semiring_gcd,idom_divide,normalization_semidom_multiplicative} poly" + shows "primitive_part (normalize f) = normalize (primitive_part f)" proof (cases "f = 0") case True thus ?thesis by simp next case False have "normalize (content (normalize (primitive_part f))) = 1" using content_primitive_part[OF False] content_dvd content_const content_dvd_contentI dvd_normalize_iff is_unit_content_iff by (metis (no_types)) then have "content (normalize (primitive_part f)) = 1" by fastforce then have "content (normalize f) = 1 * content f" by (metis (no_types) content_smult mult.commute normalize_content smult_content_normalize_primitive_part) then have "content f = content (normalize f)" by simp then show ?thesis unfolding smult_content_normalize_primitive_part[of f,symmetric] by (metis (no_types) False content_times_primitive_part mult.commute mult_cancel_left mult_smult_right smult_content_normalize_primitive_part) qed lemma length_coeffs_primitive_part[simp]: "length (coeffs (primitive_part f)) = length (coeffs f)" proof (cases "f = 0") case False hence "length (coeffs f) \ 0" "length (coeffs (primitive_part f)) \ 0" by auto thus ?thesis using degree_primitive_part[of f, unfolded degree_eq_length_coeffs] by linarith qed simp lemma degree_unit_factor[simp]: "degree (unit_factor f) = 0" by (simp add: monom_0 unit_factor_poly_def) lemma degree_normalize[simp]: "degree (normalize f) = degree f" proof (cases "f = 0") case False have "degree f = degree (unit_factor f * normalize f)" by simp also have "\ = degree (unit_factor f) + degree (normalize f)" by (rule degree_mult_eq, insert False, auto) finally show ?thesis by simp qed simp lemma content_iff: "x dvd content p \ (\ c \ set (coeffs p). x dvd c)" by (simp add: content_def dvd_gcd_list_iff) lemma is_unit_field_poly[simp]: "(p::'a::field poly) dvd 1 \ p \ 0 \ degree p = 0" proof(intro iffI conjI, unfold conj_imp_eq_imp_imp) assume "is_unit p" then obtain q where *: "p * q = 1" by (elim dvdE, auto) from * show p0: "p \ 0" by auto from * have q0: "q \ 0" by auto from * degree_mult_eq[OF p0 q0] show "degree p = 0" by auto next assume "degree p = 0" from degree0_coeffs[OF this] obtain c where c: "p = [:c:]" by auto assume "p \ 0" with c have "c \ 0" by auto with c have "1 = p * [:1/c:]" by auto from dvdI[OF this] show "is_unit p". qed definition primitive where "primitive f \ (\x. (\y \ set (coeffs f). x dvd y) \ x dvd 1)" lemma primitiveI: assumes "(\x. (\y. y \ set (coeffs f) \ x dvd y) \ x dvd 1)" shows "primitive f" by (insert assms, auto simp: primitive_def) lemma primitiveD: assumes "primitive f" shows "(\y. y \ set (coeffs f) \ x dvd y) \ x dvd 1" by (insert assms, auto simp: primitive_def) lemma not_primitiveE: assumes "\ primitive f" and "\x. (\y. y \ set (coeffs f) \ x dvd y) \ \ x dvd 1 \ thesis" shows thesis by (insert assms, auto simp: primitive_def) lemma primitive_iff_content_eq_1[simp]: fixes f :: "'a :: semiring_gcd poly" shows "primitive f \ content f = 1" proof(intro iffI primitiveI) fix x assume "(\y. y \ set (coeffs f) \ x dvd y)" from gcd_list_greatest[of "coeffs f", OF this] have "x dvd content f" by (simp add: content_def) also assume "content f = 1" finally show "x dvd 1". next assume "primitive f" from primitiveD[OF this list_gcd[of _ "coeffs f"], folded content_def] show "content f = 1" by simp qed lemma primitive_prod_list: - fixes fs :: "'a :: {factorial_semiring,semiring_Gcd} poly list" + fixes fs :: "'a :: {factorial_semiring,semiring_Gcd,normalization_semidom_multiplicative} poly list" assumes "primitive (prod_list fs)" and "f \ set fs" shows "primitive f" proof (insert assms, induct fs arbitrary: f) case (Cons f' fs) from Cons.prems have "is_unit (content f' * content (prod_list fs))" by (auto simp: content_mult) from this[unfolded is_unit_mult_iff] have "content f' = 1" and "content (prod_list fs) = 1" by auto moreover from Cons.prems have "f = f' \ f \ set fs" by auto ultimately show ?case using Cons.hyps[of f] by auto qed auto lemma irreducible_imp_primitive: fixes f :: "'a :: {idom,semiring_gcd} poly" assumes irr: "irreducible f" and deg: "degree f \ 0" shows "primitive f" proof (rule ccontr) assume not: "\ ?thesis" then have "\ [:content f:] dvd 1" by simp moreover have "f = [:content f:] * primitive_part f" by simp note Factorial_Ring.irreducibleD[OF irr this] ultimately have "primitive_part f dvd 1" by auto from this[unfolded poly_dvd_1] have "degree f = 0" by auto with deg show False by auto qed lemma irreducible_primitive_connect: fixes f :: "'a :: {idom,semiring_gcd} poly" assumes cf: "primitive f" shows "irreducible\<^sub>d f \ irreducible f" (is "?l \ ?r") proof assume l: ?l show ?r proof(rule ccontr, elim not_irreducibleE) from l have deg: "degree f > 0" by (auto dest: irreducible\<^sub>dD) from cf have f0: "f \ 0" by auto then show "f = 0 \ False" by auto show "f dvd 1 \ False" using deg by (auto simp:poly_dvd_1) fix a b assume fab: "f = a * b" and a1: "\ a dvd 1" and b1: "\ b dvd 1" then have af: "a dvd f" and bf: "b dvd f" by auto with f0 have a0: "a \ 0" and b0: "b \ 0" by auto from irreducible\<^sub>dD(2)[OF l, of a] af dvd_imp_degree_le[OF af f0] have "degree a = 0 \ degree a = degree f" by (metis degree_smult_le irreducible\<^sub>d_dvd_smult l le_antisym Nat.neq0_conv) then show False proof(elim disjE) assume "degree a = 0" then obtain c where ac: "a = [:c:]" by (auto dest: degree0_coeffs) from fab[unfolded ac] have "c dvd content f" by (simp add: content_iff coeffs_smult) with cf have "c dvd 1" by simp then have "a dvd 1" by (auto simp: ac) with a1 show False by auto next assume dega: "degree a = degree f" with f0 degree_mult_eq[OF a0 b0] fab have "degree b = 0" by (auto simp: ac_simps) then obtain c where bc: "b = [:c:]" by (auto dest: degree0_coeffs) from fab[unfolded bc] have "c dvd content f" by (simp add: content_iff coeffs_smult) with cf have "c dvd 1" by simp then have "b dvd 1" by (auto simp: bc) with b1 show False by auto qed qed next assume r: ?r show ?l proof(intro irreducible\<^sub>dI) show "degree f > 0" proof (rule ccontr) assume "\degree f > 0" then obtain f0 where f: "f = [:f0:]" by (auto dest: degree0_coeffs) from cf[unfolded this] have "normalize f0 = 1" by auto then have "f0 dvd 1" by (unfold normalize_1_iff) with r[unfolded f irreducible_const_poly_iff] show False by auto qed next fix g h assume deg_g: "degree g > 0" and deg_gf: "degree g < degree f" and fgh: "f = g * h" with r have "g dvd 1 \ h dvd 1" by auto with deg_g have "degree h = 0" by (auto simp: poly_dvd_1) with deg_gf[unfolded fgh] degree_mult_eq[of g h] show False by (cases "g = 0 \ h = 0", auto) qed qed lemma deg_not_zero_imp_not_unit: fixes f:: "'a::{idom_divide,semidom_divide_unit_factor} poly" assumes deg_f: "degree f > 0" shows "\ is_unit f" proof - have "degree (normalize f) > 0" using deg_f degree_normalize by auto hence "normalize f \ 1" by fastforce thus "\ is_unit f" using normalize_1_iff by auto qed lemma content_pCons[simp]: "content (pCons a p) = gcd a (content p)" proof(induct p arbitrary: a) case 0 show ?case by simp next case (pCons c p) then show ?case by (cases "p = 0", auto simp: content_def cCons_def) qed lemma content_field_poly: fixes f :: "'a :: {field,semiring_gcd} poly" shows "content f = (if f = 0 then 0 else 1)" by(induct f, auto simp: dvd_field_iff is_unit_normalize) end diff --git a/thys/Polynomial_Factorization/Prime_Factorization.thy b/thys/Polynomial_Factorization/Prime_Factorization.thy --- a/thys/Polynomial_Factorization/Prime_Factorization.thy +++ b/thys/Polynomial_Factorization/Prime_Factorization.thy @@ -1,800 +1,801 @@ (* Author: RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Prime Factorization\ text \This theory contains not-completely naive algorithms to test primality and to perform prime factorization. More precisely, it corresponds to prime factorization algorithm A in Knuth's textbook \cite{Knuth}.\ theory Prime_Factorization imports "HOL-Computational_Algebra.Primes" Missing_List Missing_Multiset begin subsection \Definitions\ definition primes_1000 :: "nat list" where "primes_1000 = [2, 3, 5, 7, 11, 13, 17, 19, 23, 29, 31, 37, 41, 43, 47, 53, 59, 61, 67, 71, 73, 79, 83, 89, 97, 101, 103, 107, 109, 113, 127, 131, 137, 139, 149, 151, 157, 163, 167, 173, 179, 181, 191, 193, 197, 199, 211, 223, 227, 229, 233, 239, 241, 251, 257, 263, 269, 271, 277, 281, 283, 293, 307, 311, 313, 317, 331, 337, 347, 349, 353, 359, 367, 373, 379, 383, 389, 397, 401, 409, 419, 421, 431, 433, 439, 443, 449, 457, 461, 463, 467, 479, 487, 491, 499, 503, 509, 521, 523, 541, 547, 557, 563, 569, 571, 577, 587, 593, 599, 601, 607, 613, 617, 619, 631, 641, 643, 647, 653, 659, 661, 673, 677, 683, 691, 701, 709, 719, 727, 733, 739, 743, 751, 757, 761, 769, 773, 787, 797, 809, 811, 821, 823, 827, 829, 839, 853, 857, 859, 863, 877, 881, 883, 887, 907, 911, 919, 929, 937, 941, 947, 953, 967, 971, 977, 983, 991, 997]" lemma primes_1000: "primes_1000 = filter prime [0..<1001]" by eval definition next_candidates :: "nat \ nat \ nat list" where "next_candidates n = (if n = 0 then (1001,primes_1000) else (n + 30, [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]))" definition "candidate_invariant n = (n = 0 \ n mod 30 = (11 :: nat))" partial_function (tailrec) remove_prime_factor :: "nat \ nat \ nat list \ nat \ nat list " where [code]: "remove_prime_factor p n ps = (case Divides.divmod_nat n p of (n',m) \ if m = 0 then remove_prime_factor p n' (p # ps) else (n,ps))" partial_function (tailrec) prime_factorization_nat_main :: "nat \ nat \ nat list \ nat list \ nat list" where [code]: "prime_factorization_nat_main n j is ps = (case is of [] \ (case next_candidates j of (j,is) \ prime_factorization_nat_main n j is ps) | (i # is) \ (case Divides.divmod_nat n i of (n',m) \ if m = 0 then case remove_prime_factor i n' (i # ps) of (n',ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' j is ps' else if i * i \ n then prime_factorization_nat_main n j is ps else (n # ps)))" partial_function (tailrec) prime_nat_main :: "nat \ nat \ nat list \ bool" where [code]: "prime_nat_main n j is = (case is of [] \ (case next_candidates j of (j,is) \ prime_nat_main n j is) | (i # is) \ (if i dvd n then i \ n else if i * i \ n then prime_nat_main n j is else True))" definition prime_nat :: "nat \ bool" where "prime_nat n \ if n < 2 then False else \ \TODO: integrate precomputed map\ case next_candidates 0 of (j,is) \ prime_nat_main n j is" definition prime_factorization_nat :: "nat \ nat list" where "prime_factorization_nat n \ rev (if n < 2 then [] else case next_candidates 0 of (j,is) \ prime_factorization_nat_main n j is [])" definition divisors_nat :: "nat \ nat list" where "divisors_nat n \ if n = 0 then [] else remdups_adj (sort (map prod_list (subseqs (prime_factorization_nat n))))" definition divisors_int_pos :: "int \ int list" where "divisors_int_pos x \ map int (divisors_nat (nat (abs x)))" definition divisors_int :: "int \ int list" where "divisors_int x \ let xs = divisors_int_pos x in xs @ (map uminus xs)" subsection \Proofs\ lemma remove_prime_factor: assumes res: "remove_prime_factor i n ps = (m,qs)" and i: "i > 1" and n: "n \ 0" shows "\ rs. qs = rs @ ps \ n = m * prod_list rs \ \ i dvd m \ set rs \ {i}" using res n proof (induct n arbitrary: ps rule: less_induct) case (less n ps) obtain n' mo where dm: "Divides.divmod_nat n i = (n',mo)" by force hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: divmod_nat_def) from less(2)[unfolded remove_prime_factor.simps[of i n] dm] have res: "(if mo = 0 then remove_prime_factor i n' (i # ps) else (n, ps)) = (m, qs)" by auto from less(3) have n: "n \ 0" by auto with n' i have "n' < n" by auto note IH = less(1)[OF this] show ?case proof (cases "mo = 0") case True with mo n' have n: "n = n' * i" by auto with \n \ 0\ have n': "n' \ 0" by auto from True res have "remove_prime_factor i n' (i # ps) = (m,qs)" by auto from IH[OF this n'] obtain rs where "qs = rs @ i # ps" and "n' = m * prod_list rs \ \ i dvd m \ set rs \ {i}" by auto thus ?thesis by (intro exI[of _ "rs @ [i]"], unfold n, auto) next case False with mo have i_n: "\ i dvd n" by auto from False res have id: "m = n" "qs = ps" by auto show ?thesis unfolding id using i_n by auto qed qed lemma prime_sqrtI: assumes n: "n \ 2" and small: "\ j. 2 \ j \ j < i \ \ j dvd n" and i: "\ i * i \ n" shows "prime (n::nat)" unfolding prime_nat_iff proof (intro conjI impI allI) show "1 < n" using n by auto fix j assume jn: "j dvd n" from jn obtain k where njk: "n = j * k" unfolding dvd_def by auto with \1 < n\ have jn: "j \ n" by (metis dvd_imp_le jn neq0_conv not_less0) show "j = 1 \ j = n" proof (rule ccontr) assume "\ ?thesis" with njk n have j1: "j > 1 \ j \ n" by simp have "\ j k. 1 < j \ j \ k \ n = j * k" proof (cases "j \ k") case True thus ?thesis unfolding njk using j1 by blast next case False show ?thesis by (rule exI[of _ k], rule exI[of _ j], insert \1 < n\ j1 njk False, auto) (metis Suc_lessI mult_0_right neq0_conv) qed then obtain j k where j1: "1 < j" and jk: "j \ k" and njk: "n = j * k" by auto with small[of j] have ji: "j \ i" unfolding dvd_def by force from mult_mono[OF ji ji] have "i * i \ j * j" by auto with i have "j * j > n" by auto from this[unfolded njk] have "k < j" by auto with jk show False by auto qed qed lemma candidate_invariant_0: "candidate_invariant 0" unfolding candidate_invariant_def by auto lemma next_candidates: assumes res: "next_candidates n = (m,ps)" and n: "candidate_invariant n" shows "candidate_invariant m" "sorted ps" "{i. prime i \ n \ i \ i < m} \ set ps" "set ps \ {2..} \ {n.. []" "n < m" unfolding candidate_invariant_def proof - note res = res[unfolded next_candidates_def] note n = n[unfolded candidate_invariant_def] show "m = 0 \ m mod 30 = 11" using res n by (auto split: if_splits) show "sorted ps" using res n by (auto split: if_splits simp: primes_1000_def sorted2_simps simp del: sorted.simps(2)) show "set ps \ {2..} \ {n.. []" using res n by (auto split: if_splits simp: primes_1000_def) show "n < m" using res by (auto split: if_splits) show "{i. prime i \ n \ i \ i < m} \ set ps" proof (cases "n = 0") case True hence *: "m = 1001" "ps = primes_1000" using res by auto show ?thesis unfolding * True primes_1000 by auto next case False hence n: "n mod 30 = 11" and m: "m = n + 30" and ps: "ps = [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]" using res n by auto { fix i assume *: "prime i" "n \ i" "i < n + 30" "i \ set ps" from n * have i11: "i \ 11" by auto define j where "j = i - n" have i: "i = n + j" using \n \ i\ j_def by auto have "i mod 30 = (j + n) mod 30" using \n \ i\ unfolding j_def by simp also have "\ = (j mod 30 + n mod 30) mod 30" by (simp add: mod_simps) also have "\ = (j mod 30 + 11) mod 30" unfolding n by simp finally have i30: "i mod 30 = (j mod 30 + 11) mod 30" by simp have 2: "2 dvd (30 :: nat)" and 112: "11 mod (2 :: nat) = 1" by simp_all have "(j + 11) mod 2 = (j + 1) mod 2" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\j. j mod 2"] have 2: "i mod 2 = (j mod 2 + 1) mod 2" by (simp add: mod_simps mod_mod_cancel [OF 2]) have 3: "3 dvd (30 :: nat)" and 113: "11 mod (3 :: nat) = 2" by simp_all have "(j + 11) mod 3 = (j + 2) mod 3" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 3"] have 3: "i mod 3 = (j mod 3 + 2) mod 3" by (simp add: mod_simps mod_mod_cancel [OF 3]) have 5: "5 dvd (30 :: nat)" and 115: "11 mod (5 :: nat) = 1" by simp_all have "(j + 11) mod 5 = (j + 1) mod 5" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 5"] have 5: "i mod 5 = (j mod 5 + 1) mod 5" by (simp add: mod_simps mod_mod_cancel [OF 5]) from n *(2-)[unfolded ps i, simplified] have "j \ {1,3,5,7,9,11,13,15,17,19,21,23,25,27,29} \ j \ {4,10,16,22,28} \ j \ {14,24}" (is "j \ ?j2 \ j \ ?j3 \ j \ ?j5") by simp presburger moreover { assume "j \ ?j2" hence "j mod 2 = 1" by auto with 2 have "i mod 2 = 0" by auto with i11 have "2 dvd i" "i \ 2" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j3" hence "j mod 3 = 1" by auto with 3 have "i mod 3 = 0" by auto with i11 have "3 dvd i" "i \ 3" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j5" hence "j mod 5 = 4" by auto with 5 have "i mod 5 = 0" by auto with i11 have "5 dvd i" "i \ 5" by auto with *(1) have False unfolding prime_nat_iff by auto } ultimately have False by blast } thus ?thesis unfolding m ps by auto qed qed lemma prime_test_iterate2: assumes small: "\ j. 2 \ j \ j < (i :: nat) \ \ j dvd n" and odd: "odd n" and n: "n \ 3" and i: "i \ 3" "odd i" and mod: "\ i dvd n" and j: "2 \ j" "j < i + 2" shows "\ j dvd n" proof assume dvd: "j dvd n" with small[OF j(1)] have "j \ i" by linarith with dvd mod have "j > i" by (cases "i = j", auto) with j have "j = Suc i" by simp with i have "even j" by auto with dvd j(1) have "2 dvd n" by (metis dvd_trans) with odd show False by auto qed lemma prime_divisor: assumes "j \ 2" and "j dvd n" shows "\ p :: nat. prime p \ p dvd j \ p dvd n" proof - let ?pf = "prime_factors j" from assms have "j > 0" by auto from prime_factorization_nat[OF this] have "j = (\p\?pf. p ^ multiplicity p j)" by auto with \j \ 2\ have "?pf \ {}" by auto then obtain p where p: "p \ ?pf" by auto hence pr: "prime p" by auto define rem where "rem = (\p\?pf - {p}. p ^ multiplicity p j)" from p have mult: "multiplicity p j \ 0" by (auto simp: prime_factors_multiplicity) have "finite ?pf" by simp have "j = (\p\?pf. p ^ multiplicity p j)" by fact also have "?pf = (insert p (?pf - {p}))" using p by auto also have "(\p\insert p (?pf - {p}). p ^ multiplicity p j) = p ^ multiplicity p j * rem" unfolding rem_def by (subst prod.insert, auto) also have "\ = p * (p ^ (multiplicity p j - 1) * rem)" using mult by (cases "multiplicity p j", auto) finally have pj: "p dvd j" unfolding dvd_def by blast with \j dvd n\ have "p dvd n" by (metis dvd_trans) with pj pr show ?thesis by blast qed lemma prime_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_nat_main n jj is \ res = prime n" proof (induct ni arbitrary: n i "is" jj res rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_nat_main n jjj iis" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) with res[unfolded simps] have res: "res = (if i' dvd n then n \ i' else if i' * i' \ n then prime_nat_main n jj iis else True)" by simp from 1(11) Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) have "set iis \ {i'..}" by(auto simp: Cons) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by(auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 n _ _ i'(3) sd_iis 1(10) iis] show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_nat_main n jj iis" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis by (rule IH[OF _ dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = True" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True have "i' \ 2" using i i' by auto from \i' dvd n\ obtain k where "n = i' * k" .. with n have "k \ 0" by (cases "k = 0", auto) with \n = i' * k\ have *: "i' < n \ i' = n" by auto with True res have "res \ i' = n" by auto also have "\ = prime n" using * proof assume "i' < n" with \i' \ 2\ \i' dvd n\ have "\ prime n" by (auto simp add: prime_nat_iff) with \i' < n\ show ?thesis by auto next assume "i' = n" with dvd n have "prime n" by (simp add: prime_nat_iff') with \i' = n\ show ?thesis by auto qed finally show ?thesis . qed qed qed lemma prime_factorization_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_factorization_nat_main n jj is ps \ \ qs. res = qs @ ps \ Ball (set qs) prime \ n = prod_list qs" proof (induct ni arbitrary: n i "is" jj res ps rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res ps) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_factorization_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_factorization_nat_main n jjj iis ps" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) obtain n' m where dm: "Divides.divmod_nat n i' = (n',m)" by force hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: divmod_nat_def) have m: "(m = 0) = (i' dvd n)" unfolding m by auto from Cons res[unfolded simps] dm m n' have res: "res = ( if i' dvd n then case remove_prime_factor i' (n div i') (i' # ps) of (n', ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' jj iis ps' else if i' * i' \ n then prime_factorization_nat_main n jj iis ps else n # ps)" by simp from 1(11) i Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" "i' > 1" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) Cons have "set iis \ {i'..}" by(auto) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by (auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 _ _ _ i'(3) sd_iis 1(10) iis] { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_factorization_nat_main n jj iis ps" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast show ?thesis by (rule IH[OF _ n dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = n # ps" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True note i_n = this obtain n'' qs where rp: "remove_prime_factor i' (n div i') (i' # ps) = (n'',qs)" by force with res True have res: "res = (if n'' = 1 then qs else prime_factorization_nat_main n'' jj iis qs)" by auto have pi: "prime i'" unfolding prime_nat_iff proof (intro conjI allI impI) show "1 < i'" using i' i by auto fix j assume ji: "j dvd i'" with i' i have j0: "j \ 0" by (cases "j = 0", auto) from ji i_n have jn: "j dvd n" by (metis dvd_trans) with dvd[of j] have j: "2 > j \ j \ i'" by linarith from ji \1 < i'\ have "j \ i'" unfolding dvd_def by (simp add: dvd_imp_le ji) with j j0 show "j = 1 \ j = i'" by linarith qed from True n' have id: "n = n' * i'" by auto from n id have "n' \ 0" by (cases "n = 0", auto) with id have "i' \ n" by auto from remove_prime_factor[OF rp[folded n'] \1 < i'\ \n' \ 0\] obtain rs where qs: "qs = rs @ i' # ps" and n': "n' = n'' * prod_list rs" and i_n'': "\ i' dvd n''" and rs: "set rs \ {i'}" by auto { fix j assume "j dvd n''" hence "j dvd n" unfolding id n' by auto } note dvd' = this show ?thesis proof (cases "n'' = 1") case False with res have res: "res = prime_factorization_nat_main n'' jj iis qs" by simp from i i' have "i' \ 2" by simp from False n' \n' \ 0\ have n2: "n'' \ 2" by (cases "n'' = 0"; auto) have lrs: "prod_list rs \ 0" using n' \n' \ 0\ by (cases "prod_list rs = 0", auto) with \i' \ 2\ have "prod_list rs * i' \ 2" by (cases "prod_list rs", auto) hence nn'': "n > n''" unfolding id n' using n2 by simp have "i' \ n" unfolding id n' using pi False by fastforce with \i' \ n\ i' have "n > i" by auto with nn'' i i' have less: "n - i > n'' - Suc i'" by simp { fix j assume 2: "2 \ j" and ji: "j < Suc i'" have "\ j dvd n''" proof (cases "j = i'") case False with ji have "j < i'" by auto from dvd' dvd[OF 2 this] show ?thesis by blast qed (insert i_n'', auto) } from IH[OF _ n2 this iis res] less obtain ss where res: "res = ss @ qs \ Ball (set ss) prime \ n'' = prod_list ss" by auto thus ?thesis unfolding id n' qs using pi rs by auto next case True with res have res: "res = qs" by auto show ?thesis unfolding id n' res qs True using rs \prime i'\ by (intro exI[of _ "rs @ [i']"], auto) qed qed qed qed lemma prime_nat[simp]: "prime_nat n = prime n" proof (cases "n < 2") case True thus ?thesis unfolding prime_nat_def prime_nat_iff by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) from n can have res: "prime_nat n = prime_nat_main n jj is" unfolding prime_nat_def by auto show ?thesis using prime_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub res] cann(4) by auto qed lemma prime_factorization_nat: fixes n :: nat defines "pf \ prime_factorization_nat n" shows "Ball (set pf) prime" and "n \ 0 \ prod_list pf = n" and "n = 0 \ pf = []" proof - note pf = pf_def[unfolded prime_factorization_nat_def] have "Ball (set pf) prime \ (n \ 0 \ prod_list pf = n) \ (n = 0 \ pf = [])" proof (cases "n < 2") case True thus ?thesis using pf by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) let ?pfm = "prime_factorization_nat_main n jj is []" from pf[unfolded can] False have res: "pf = rev ?pfm" by simp from prime_factorization_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub refl, of Nil] cann(4) have "Ball (set ?pfm) prime" "n = prod_list ?pfm" by auto thus ?thesis unfolding res using n by auto qed thus "Ball (set pf) prime" "n \ 0 \ prod_list pf = n" "n = 0 \ pf = []" by auto qed lemma prod_mset_multiset_prime_factorization_nat [simp]: "(x::nat) \ 0 \ prod_mset (prime_factorization x) = x" by simp (* TODO Move *) lemma prime_factorization_unique'': + fixes A :: "'a :: {factorial_semiring_multiplicative} multiset" assumes "\p. p \# A \ prime p" assumes "prod_mset A = normalize x" shows "prime_factorization x = A" proof - have "prod_mset A \ 0" by (auto dest: assms(1)) with assms(2) have "x \ 0" by simp hence "prod_mset (prime_factorization x) = prod_mset A" by (simp add: assms prod_mset_prime_factorization) with assms show ?thesis by (intro prime_factorization_unique') auto qed lemma multiset_prime_factorization_nat_correct: "prime_factorization n = mset (prime_factorization_nat n)" proof - note pf = prime_factorization_nat[of n] show ?thesis proof (cases "n = 0") case True thus ?thesis using pf(3) by simp next case False note pf = pf(1) pf(2)[OF False] show ?thesis proof (rule prime_factorization_unique'') show "prime p" if "p \# mset (prime_factorization_nat n)" for p using pf(1) that by simp let ?l = "\i\#prime_factorization n. i" let ?r = "\i\#mset (prime_factorization_nat n). i" show "prod_mset (mset (prime_factorization_nat n)) = normalize n" by (simp add: pf(2) prod_mset_prod_list) qed qed qed lemma multiset_prime_factorization_code[code_unfold]: "prime_factorization = (\n. mset (prime_factorization_nat n))" by (intro ext multiset_prime_factorization_nat_correct) lemma divisors_nat: "n \ 0 \ set (divisors_nat n) = {p. p dvd n}" "distinct (divisors_nat n)" "divisors_nat 0 = []" proof - show "distinct (divisors_nat n)" "divisors_nat 0 = []" unfolding divisors_nat_def by auto assume n: "n \ 0" from n have "n > 0" by auto { fix x have "(x dvd n) = (x \ 0 \ (\p. multiplicity p x \ multiplicity p n))" proof (cases "x = 0") case False with \n > 0\ show ?thesis by (auto simp: dvd_multiplicity_eq) next case True with n show ?thesis by auto qed } note dvd = this let ?dn = "set (divisors_nat n)" let ?mf = "\ (n :: nat). prime_factorization n" have "?dn = prod_list ` set (subseqs (prime_factorization_nat n))" unfolding divisors_nat_def using n by auto also have "\ = prod_mset ` mset ` set (subseqs (prime_factorization_nat n))" by (force simp: prod_mset_prod_list) also have "mset ` set (subseqs (prime_factorization_nat n)) = { ps. ps \# mset (prime_factorization_nat n)}" unfolding multiset_of_subseqs by simp also have "\ = { ps. ps \# ?mf n}" thm multiset_prime_factorization_code[symmetric] unfolding multiset_prime_factorization_nat_correct[symmetric] by auto also have "prod_mset ` \ = {p. p dvd n}" (is "?l = ?r") proof - { fix x assume "x dvd n" from this[unfolded dvd] have x: "x \ 0" by auto from \x dvd n\ \x \ 0\ \n \ 0\ have sub: "?mf x \# ?mf n" by (subst prime_factorization_subset_iff_dvd) auto have "prod_mset (?mf x) = x" using x by (simp add: prime_factorization_nat) hence "x \ ?l" using sub by force } moreover { fix x assume "x \ ?l" then obtain ps where x: "x = prod_mset ps" and sub: "ps \# ?mf n" by auto have "x dvd n" using prod_mset_subset_imp_dvd[OF sub] n x by simp } ultimately show ?thesis by blast qed finally show "set (divisors_nat n) = {p. p dvd n}" . qed lemma divisors_int_pos: "x \ 0 \ set (divisors_int_pos x) = {i. i dvd x \ i > 0}" "distinct (divisors_int_pos x)" "divisors_int_pos 0 = []" proof - show "divisors_int_pos 0 = []" by code_simp show "distinct (divisors_int_pos x)" unfolding divisors_int_pos_def using divisors_nat(2)[of "nat (abs x)"] by (simp add: distinct_map inj_on_def) assume x: "x \ 0" let ?x = "nat (abs x)" from x have xx: "?x \ 0" by auto from x have 0: "\ y. y dvd x \ y \ 0" by auto have id: "int ` {p. int p dvd x} = {i. i dvd x \ 0 < i}" (is "?l = ?r") proof - { fix y assume "y \ ?l" then obtain p where y: "y = int p" and dvd: "int p dvd x" by auto have "y \ ?r" unfolding y using dvd 0[OF dvd] by auto } moreover { fix y assume "y \ ?r" hence dvd: "y dvd x" and y0: "y > 0" by auto define n where "n = nat y" from y0 have y: "y = int n" unfolding n_def by auto with dvd have "y \ ?l" by auto } ultimately show ?thesis by blast qed from xx show "set (divisors_int_pos x) = {i. i dvd x \ i > 0}" by (simp add: divisors_int_pos_def divisors_nat id) qed lemma divisors_int: "x \ 0 \ set (divisors_int x) = {i. i dvd x}" "distinct (divisors_int x)" "divisors_int 0 = []" proof - show "divisors_int 0 = []" by code_simp show "distinct (divisors_int x)" proof (cases "x = 0") case True show ?thesis unfolding True by code_simp next case False from divisors_int_pos(1)[OF False] divisors_int_pos(2) show ?thesis unfolding divisors_int_def Let_def distinct_append distinct_map inj_on_def by auto qed assume x: "x \ 0" show "set (divisors_int x) = {i. i dvd x}" unfolding divisors_int_def Let_def set_append set_map divisors_int_pos(1)[OF x] using x by auto (metis (no_types, lifting) dvd_mult_div_cancel image_eqI linorder_neqE_linordered_idom mem_Collect_eq minus_dvd_iff minus_minus mult_zero_left neg_less_0_iff_less) qed definition divisors_fun :: "('a \ ('a :: {comm_monoid_mult,zero}) list) \ bool" where "divisors_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x }) \ (\ x. distinct (df x))" lemma divisors_funD: "divisors_fun df \ x \ 0 \ d dvd x \ d \ set (df x)" unfolding divisors_fun_def by auto definition divisors_pos_fun :: "('a \ ('a :: {comm_monoid_mult,zero,ord}) list) \ bool" where "divisors_pos_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x \ d > 0}) \ (\ x. distinct (df x))" lemma divisors_pos_funD: "divisors_pos_fun df \ x \ 0 \ d dvd x \ d > 0 \ d \ set (df x)" unfolding divisors_pos_fun_def by auto lemma divisors_fun_nat: "divisors_fun divisors_nat" unfolding divisors_fun_def using divisors_nat by auto lemma divisors_fun_int: "divisors_fun divisors_int" unfolding divisors_fun_def using divisors_int by auto lemma divisors_pos_fun_int: "divisors_pos_fun divisors_int_pos" unfolding divisors_pos_fun_def using divisors_int_pos by auto end diff --git a/thys/Polynomial_Factorization/Square_Free_Factorization.thy b/thys/Polynomial_Factorization/Square_Free_Factorization.thy --- a/thys/Polynomial_Factorization/Square_Free_Factorization.thy +++ b/thys/Polynomial_Factorization/Square_Free_Factorization.thy @@ -1,1446 +1,1451 @@ (* Author: RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Square Free Factorization\ text \We implemented Yun's algorithm to perform a square-free factorization of a polynomial. We further show properties of a square-free factorization, namely that the exponents in the square-free factorization are exactly the orders of the roots. We also show that factorizing the result of square-free factorization further will again result in a square-free factorization, and that square-free factorizations can be lifted homomorphically.\ theory Square_Free_Factorization imports Matrix.Utility Polynomial_Divisibility Order_Polynomial Fundamental_Theorem_Algebra_Factorized Polynomial_Interpolation.Ring_Hom_Poly begin definition square_free :: "'a :: comm_semiring_1 poly \ bool" where "square_free p = (p \ 0 \ (\ q. degree q > 0 \ \ (q * q dvd p)))" lemma square_freeI: assumes "\ q. degree q > 0 \ q \ 0 \ q * q dvd p \ False" and p: "p \ 0" shows "square_free p" unfolding square_free_def proof (intro allI conjI[OF p] impI notI, goal_cases) case (1 q) from assms(1)[OF 1(1) _ 1(2)] 1(1) show False by (cases "q = 0", auto) qed lemma square_free_multD: assumes sf: "square_free (f * g)" shows "h dvd f \ h dvd g \ degree h = 0" "square_free f" "square_free g" proof - from sf[unfolded square_free_def] have 0: "f \ 0" "g \ 0" and dvd: "\ q. q * q dvd f * g \ degree q = 0" by auto then show "square_free f" "square_free g" by (auto simp: square_free_def) assume "h dvd f" "h dvd g" then have "h * h dvd f * g" by (rule mult_dvd_mono) from dvd[OF this] show "degree h = 0". qed lemma irreducible\<^sub>d_square_free: fixes p :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} poly" shows "irreducible\<^sub>d p \ square_free p" by (metis degree_0 degree_mult_eq degree_mult_eq_0 irreducible\<^sub>dD(1) irreducible\<^sub>dD(2) irreducible\<^sub>d_dvd_smult irreducible\<^sub>d_smultI less_add_same_cancel2 not_gr_zero square_free_def) lemma square_free_factor: assumes dvd: "a dvd p" and sf: "square_free p" shows "square_free a" proof (intro square_freeI) fix q assume q: "degree q > 0" and "q * q dvd a" hence "q * q dvd p" using dvd dvd_trans sf square_free_def by blast with sf[unfolded square_free_def] q show False by auto qed (insert dvd sf, auto simp: square_free_def) lemma square_free_prod_list_distinct: assumes sf: "square_free (prod_list us :: 'a :: idom poly)" and us: "\ u. u \ set us \ degree u > 0" shows "distinct us" proof (rule ccontr) assume "\ distinct us" from not_distinct_decomp[OF this] obtain xs ys zs u where "us = xs @ u # ys @ u # zs" by auto hence dvd: "u * u dvd prod_list us" and u: "u \ set us" by auto from dvd us[OF u] sf have "prod_list us = 0" unfolding square_free_def by auto hence "0 \ set us" by (simp add: prod_list_zero_iff) from us[OF this] show False by auto qed definition separable where "separable f = coprime f (pderiv f)" lemma separable_imp_square_free: - assumes sep: "separable (f :: 'a::{field, factorial_ring_gcd} poly)" + assumes sep: "separable (f :: 'a::{field, factorial_ring_gcd, semiring_gcd_mult_normalize} poly)" shows "square_free f" proof (rule ccontr) note sep = sep[unfolded separable_def] from sep have f0: "f \ 0" by (cases f, auto) assume "\ square_free f" then obtain g where g: "degree g \ 0" and "g * g dvd f" using f0 unfolding square_free_def by auto then obtain h where f: "f = g * (g * h)" unfolding dvd_def by (auto simp: ac_simps) have "pderiv f = g * ((g * pderiv h + h * pderiv g) + h * pderiv g)" unfolding f pderiv_mult[of g] by (simp add: field_simps) hence "g dvd pderiv f" unfolding dvd_def by blast moreover have "g dvd f" unfolding f dvd_def by blast ultimately have dvd: "g dvd (gcd f (pderiv f))" by simp have "gcd f (pderiv f) \ 0" using f0 by simp with g dvd have "degree (gcd f (pderiv f)) \ 0" by (simp add: sep poly_dvd_1) hence "\ coprime f (pderiv f)" by auto with sep show False by simp qed lemma square_free_rsquarefree: assumes f: "square_free f" shows "rsquarefree f" unfolding rsquarefree_def proof (intro conjI allI) fix x show "order x f = 0 \ order x f = 1" proof (rule ccontr) assume "\ ?thesis" then obtain n where ord: "order x f = Suc (Suc n)" by (cases "order x f"; cases "order x f - 1"; auto) define p where "p = [:-x,1:]" from order_divides[of x "Suc (Suc 0)" f, unfolded ord] have "p * p dvd f" "degree p \ 0" unfolding p_def by auto hence "\ square_free f" using f(1) unfolding square_free_def by auto with assms show False by auto qed qed (insert f, auto simp: square_free_def) lemma square_free_prodD: - fixes fs :: "'a :: {field,euclidean_ring_gcd} poly set" + fixes fs :: "'a :: {field,euclidean_ring_gcd,semiring_gcd_mult_normalize} poly set" assumes sf: "square_free (\ fs)" and fin: "finite fs" and f: "f \ fs" and g: "g \ fs" and fg: "f \ g" shows "coprime f g" proof - have "(\ fs) = f * (\ (fs - {f}))" by (rule prod.remove[OF fin f]) also have "(\ (fs - {f})) = g * (\ (fs - {f} - {g}))" by (rule prod.remove, insert fin g fg, auto) finally obtain k where sf: "square_free (f * g * k)" using sf by (simp add: ac_simps) from sf[unfolded square_free_def] have 0: "f \ 0" "g \ 0" and dvd: "\ q. q * q dvd f * g * k \ degree q = 0" by auto have "gcd f g * gcd f g dvd f * g * k" by (simp add: mult_dvd_mono) from dvd[OF this] have "degree (gcd f g) = 0" . moreover have "gcd f g \ 0" using 0 by auto ultimately show "coprime f g" using is_unit_gcd[of f g] is_unit_iff_degree[of "gcd f g"] by simp qed lemma rsquarefree_square_free_complex: assumes "rsquarefree (p :: complex poly)" shows "square_free p" proof (rule square_freeI) fix q assume d: "degree q > 0" and dvd: "q * q dvd p" from d have "\ constant (poly q)" by (simp add: constant_degree) from fundamental_theorem_of_algebra[OF this] obtain x where "poly q x = 0" by auto hence "[:-x,1:] dvd q" by (simp add: poly_eq_0_iff_dvd) then obtain k where q: "q = [:-x,1:] * k" unfolding dvd_def by auto from dvd obtain l where p: "p = q * q * l" unfolding dvd_def by auto from p[unfolded q] have "p = [:-x,1:]^2 * (k * k * l)" by algebra hence "[:-x,1:]^2 dvd p" unfolding dvd_def by blast from this[unfolded order_divides] have "p = 0 \ \ order x p \ 1" by auto thus False using assms unfolding rsquarefree_def' by auto qed (insert assms, auto simp: rsquarefree_def) lemma square_free_separable_main: - fixes f :: "'a :: {field,factorial_ring_gcd} poly" + fixes f :: "'a :: {field,factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes "square_free f" and sep: "\ separable f" shows "\ g k. f = g * k \ degree g \ 0 \ pderiv g = 0" proof - note cop = sep[unfolded separable_def] from assms have f: "f \ 0" unfolding square_free_def by auto let ?g = "gcd f (pderiv f)" define G where "G = ?g" from poly_gcd_monic[of f "pderiv f"] f have mon: "monic ?g" by auto have deg: "degree G > 0" proof (cases "degree G") case 0 from degree0_coeffs[OF this] cop mon show ?thesis by (auto simp: G_def coprime_iff_gcd_eq_1) qed auto have gf: "G dvd f" unfolding G_def by auto have gf': "G dvd pderiv f" unfolding G_def by auto from irreducible\<^sub>d_factor[OF deg] obtain g r where g: "irreducible g" and G: "G = g * r" by auto from gf have gf: "g dvd f" unfolding G by (rule dvd_mult_left) from gf' have gf': "g dvd pderiv f" unfolding G by (rule dvd_mult_left) have g0: "degree g \ 0" using g unfolding irreducible\<^sub>d_def by auto from gf obtain k where fgk: "f = g * k" unfolding dvd_def by auto have id1: "pderiv f = g * pderiv k + k * pderiv g" unfolding fgk pderiv_mult by simp from gf' obtain h where "pderiv f = g * h" unfolding dvd_def by auto from id1[unfolded this] have "k * pderiv g = g * (h - pderiv k)" by (simp add: field_simps) hence dvd: "g dvd k * pderiv g" unfolding dvd_def by auto { assume "g dvd k" then obtain h where k: "k = g * h" unfolding dvd_def by auto with fgk have "g * g dvd f" by auto with g0 have "\ square_free f" unfolding square_free_def using f by auto with assms have False by simp } with g dvd have "g dvd pderiv g" by auto from divides_degree[OF this] degree_pderiv_le[of g] g0 have "pderiv g = 0" by linarith with fgk g0 show ?thesis by auto qed -lemma square_free_imp_separable: fixes f :: "'a :: {field_char_0,factorial_ring_gcd} poly" +lemma square_free_imp_separable: fixes f :: "'a :: {field_char_0,factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes "square_free f" shows "separable f" proof (rule ccontr) assume "\ separable f" from square_free_separable_main[OF assms this] obtain g k where *: "f = g * k" "degree g \ 0" "pderiv g = 0" by auto hence "g dvd pderiv g" by auto thus False unfolding dvd_pderiv_iff using * by auto qed lemma square_free_iff_separable: - "square_free (f :: 'a :: {field_char_0,factorial_ring_gcd} poly) = separable f" + "square_free (f :: 'a :: {field_char_0,factorial_ring_gcd,semiring_gcd_mult_normalize} poly) = separable f" using separable_imp_square_free[of f] square_free_imp_separable[of f] by auto context assumes "SORT_CONSTRAINT('a::{field,factorial_ring_gcd})" begin lemma square_free_smult: "c \ 0 \ square_free (f :: 'a poly) \ square_free (smult c f)" by (unfold square_free_def, insert dvd_smult_cancel[of _ c], auto) lemma square_free_smult_iff[simp]: "c \ 0 \ square_free (smult c f) = square_free (f :: 'a poly)" using square_free_smult[of c f] square_free_smult[of "inverse c" "smult c f"] by auto end context assumes "SORT_CONSTRAINT('a::factorial_ring_gcd)" begin definition square_free_factorization :: "'a poly \ 'a \ ('a poly \ nat) list \ bool" where "square_free_factorization p cbs \ case cbs of (c,bs) \ (p = smult c (\(a, i)\ set bs. a ^ Suc i)) \ (p = 0 \ c = 0 \ bs = []) \ (\ a i. (a,i) \ set bs \ square_free a \ degree a > 0) \ (\ a i b j. (a,i) \ set bs \ (b,j) \ set bs \ (a,i) \ (b,j) \ coprime a b) \ distinct bs" lemma square_free_factorizationD: assumes "square_free_factorization p (c,bs)" shows "p = smult c (\(a, i)\ set bs. a ^ Suc i)" "(a,i) \ set bs \ square_free a \ degree a \ 0" "(a,i) \ set bs \ (b,j) \ set bs \ (a,i) \ (b,j) \ coprime a b" "p = 0 \ c = 0 \ bs = []" "distinct bs" using assms unfolding square_free_factorization_def split by blast+ lemma square_free_factorization_prod_list: assumes "square_free_factorization p (c,bs)" shows "p = smult c (prod_list (map (\ (a,i). a ^ Suc i) bs))" proof - note sff = square_free_factorizationD[OF assms] show ?thesis unfolding sff(1) by (simp add: prod.distinct_set_conv_list[OF sff(5)]) qed end subsection \Yun's factorization algorithm\ locale yun_gcd = fixes Gcd :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" begin partial_function (tailrec) yun_factorization_main :: "'a poly \ 'a poly \ nat \ ('a poly \ nat)list \ ('a poly \ nat)list" where [code]: "yun_factorization_main bn cn i sqr = ( if bn = 1 then sqr else ( let dn = cn - pderiv bn; an = Gcd bn dn in yun_factorization_main (bn div an) (dn div an) (Suc i) ((an,i) # sqr)))" definition yun_monic_factorization :: "'a poly \ ('a poly \ nat)list" where "yun_monic_factorization p = (let pp = pderiv p; u = Gcd p pp; b0 = p div u; c0 = pp div u in (filter (\ (a,i). a \ 1) (yun_factorization_main b0 c0 0 [])))" definition square_free_monic_poly :: "'a poly \ 'a poly" where "square_free_monic_poly p = (p div (Gcd p (pderiv p)))" end declare yun_gcd.yun_monic_factorization_def [code] declare yun_gcd.yun_factorization_main.simps [code] declare yun_gcd.square_free_monic_poly_def [code] context fixes Gcd :: "'a :: {field_char_0,euclidean_ring_gcd} poly \ 'a poly \ 'a poly" begin interpretation yun_gcd Gcd . definition square_free_poly :: "'a poly \ 'a poly" where "square_free_poly p = (if p = 0 then 0 else square_free_monic_poly (smult (inverse (coeff p (degree p))) p))" definition yun_factorization :: "'a poly \ 'a \ ('a poly \ nat)list" where "yun_factorization p = (if p = 0 then (0,[]) else (let c = coeff p (degree p); q = smult (inverse c) p in (c, yun_monic_factorization q)))" lemma yun_factorization_0[simp]: "yun_factorization 0 = (0,[])" unfolding yun_factorization_def by simp end locale monic_factorization = - fixes as :: "('a :: {field_char_0,euclidean_ring_gcd} poly \ nat) set" + fixes as :: "('a :: {field_char_0,euclidean_ring_gcd,semiring_gcd_mult_normalize} poly \ nat) set" and p :: "'a poly" assumes p: "p = prod (\ (a,i). a ^ Suc i) as" and fin: "finite as" assumes as_distinct: "\ a i b j. (a,i) \ as \ (b,j) \ as \ (a,i) \ (b,j) \ a \ b" and as_irred: "\ a i. (a,i) \ as \ irreducible\<^sub>d a" and as_monic: "\ a i. (a,i) \ as \ monic a" begin lemma poly_exp_expand: "p = (prod (\ (a,i). a ^ i) as) * prod (\ (a,i). a) as" unfolding p prod.distrib[symmetric] by (rule prod.cong, auto) lemma pderiv_exp_prod: "pderiv p = (prod (\ (a,i). a ^ i) as * sum (\ (a,i). prod (\ (b,j). b) (as - {(a,i)}) * smult (of_nat (Suc i)) (pderiv a)) as)" unfolding p pderiv_prod sum_distrib_left proof (rule sum.cong[OF refl]) fix x assume "x \ as" then obtain a i where x: "x = (a,i)" and mem: "(a,i) \ as" by (cases x, auto) let ?si = "smult (of_nat (Suc i)) :: 'a poly \ 'a poly" show "(\(a, i)\as - {x}. a ^ Suc i) * pderiv (case x of (a, i) \ a ^ Suc i) = (\(a, i)\as. a ^ i) * (case x of (a, i) \ (\(a, i)\as - {(a, i)}. a) * smult (of_nat (Suc i)) (pderiv a))" unfolding x split pderiv_power_Suc proof - let ?prod = "\(a, i)\as - {(a, i)}. a ^ Suc i" let ?l = "?prod * (?si (a ^ i) * pderiv a)" let ?r = "(\(a, i)\as. a ^ i) * ((\(a, i)\as - {(a, i)}. a) * ?si (pderiv a))" have "?r = a ^ i * ((\(a, i)\as - {(a, i)}. a ^ i) * (\(a, i)\as - {(a, i)}. a) * ?si (pderiv a))" unfolding prod.remove[OF fin mem] by (simp add: ac_simps) also have "(\(a, i)\as - {(a, i)}. a ^ i) * (\(a, i)\as - {(a, i)}. a) = ?prod" unfolding prod.distrib[symmetric] by (rule prod.cong[OF refl], auto) finally show "?l = ?r" by (simp add: ac_simps) qed qed lemma monic_gen: assumes "bs \ as" shows "monic (\ (a, i) \ bs. a)" by (rule monic_prod, insert assms as_monic, auto) lemma nonzero_gen: assumes "bs \ as" shows "(\ (a, i) \ bs. a) \ 0" using monic_gen[OF assms] by auto lemma monic_Prod: "monic ((\(a, i)\as. a ^ i))" by (rule monic_prod, insert as_monic, auto intro: monic_power) -lemma coprime_generic: +lemma coprime_generic: assumes bs: "bs \ as" and f: "\ a i. (a,i) \ bs \ f i > 0" shows "coprime (\(a, i) \ bs. a) (\(a, i)\ bs. (\(b, j)\ bs - {(a, i)} . b) * smult (of_nat (f i)) (pderiv a))" (is "coprime ?single ?onederiv") proof - have single: "?single \ 0" by (rule nonzero_gen[OF bs]) show ?thesis proof (rule gcd_eq_1_imp_coprime, rule gcdI [symmetric]) fix k assume dvd: "k dvd ?single" "k dvd ?onederiv" note bs_monic = as_monic[OF subsetD[OF bs]] from dvd(1) single have k: "k \ 0" by auto show "k dvd 1" proof (cases "degree k > 0") case False with k obtain c where "k = [:c:]" by (auto dest: degree0_coeffs) with k have "c \ 0" by auto with \k = [:c:]\ show "is_unit k" using dvdI [of 1 "[:c:]" "[:1 / c:]"] by auto next case True from irreducible\<^sub>d_factor[OF this] obtain p q where k: "k = p * q" and p: "irreducible p" by auto from k dvd have dvd: "p dvd ?single" "p dvd ?onederiv" unfolding dvd_def by auto from irreducible_dvd_prod[OF p dvd(1)] obtain a i where ai: "(a,i) \ bs" and pa: "p dvd a" by force then obtain q where a: "a = p * q" unfolding dvd_def by auto from p[unfolded irreducible\<^sub>d_def] have p0: "degree p > 0" by auto from irreducible\<^sub>d_dvd_smult[OF p0 as_irred pa] ai bs obtain c where c: "c \ 0" and ap: "a = smult c p" by auto hence ap': "p = smult (1/c) a" by auto let ?prod = "\ a i. (\(b, j)\bs - {(a, i)}. b) * smult (of_nat (f i)) (pderiv a)" let ?prod' = "\ aa ii a i. (\(b, j)\bs - {(a, i),(aa,ii)}. b) * smult (of_nat (f i)) (pderiv a)" define factor where "factor = sum (\ (b,j). ?prod' a i b j ) (bs - {(a,i)})" define fac where "fac = q * factor" from fin finite_subset[OF bs] have fin: "finite bs" by auto have "?onederiv = ?prod a i + sum (\ (b,j). ?prod b j) (bs - {(a,i)})" by (subst sum.remove[OF fin ai], auto) also have "sum (\ (b,j). ?prod b j) (bs - {(a,i)}) = a * factor" unfolding factor_def sum_distrib_left proof (rule sum.cong[OF refl]) fix bj assume mem: "bj \ bs - {(a,i)}" obtain b j where bj: "bj = (b,j)" by force from mem bj ai have ai: "(a,i) \ bs - {(b,j)}" by auto have id: "bs - {(b, j)} - {(a, i)} = bs - {(b,j),(a,i)}" by auto show "(\ (b,j). ?prod b j) bj = a * (\ (b,j). ?prod' a i b j) bj" unfolding bj split by (subst prod.remove[OF _ ai], insert fin, auto simp: id ac_simps) qed finally have "?onederiv = ?prod a i + p * fac" unfolding fac_def a by simp from dvd(2)[unfolded this] have "p dvd ?prod a i" by algebra from this[unfolded field_poly_irreducible_dvd_mult[OF p]] have False proof assume "p dvd (\(b, j)\bs - {(a, i)}. b)" from irreducible_dvd_prod[OF p this] obtain b j where bj': "(b,j) \ bs - {(a,i)}" and pb: "p dvd b" by auto hence bj: "(b,j) \ bs" by auto from as_irred bj bs have "irreducible\<^sub>d b" by auto from irreducible\<^sub>d_dvd_smult[OF p0 this pb] obtain d where d: "d \ 0" and b: "b = smult d p" by auto with ap c have id: "smult (c/d) b = a" and deg: "degree a = degree b" by auto from coeff_smult[of "c/d" b "degree b", unfolded id] deg bs_monic[OF ai] bs_monic[OF bj] have "c / d = 1" by simp from id[unfolded this] have "a = b" by simp with as_distinct[OF subsetD[OF bs ai] subsetD[OF bs bj]] bj' show False by auto next from f[OF ai] obtain k where fi: "f i = Suc k" by (cases "f i", auto) assume "p dvd smult (of_nat (f i)) (pderiv a)" hence "p dvd (pderiv a)" unfolding fi using dvd_smult_cancel of_nat_eq_0_iff by blast from this[unfolded ap] have "p dvd pderiv p" using c by (metis \p dvd pderiv a\ ap' dvd_trans dvd_triv_right mult.left_neutral pderiv_smult smult_dvd_cancel) with not_dvd_pderiv p0 show False by auto qed thus "k dvd 1" by simp qed qed (insert \?single \ 0\, auto) qed lemma pderiv_exp_gcd: "gcd p (pderiv p) = (\(a, i)\as. a ^ i)" (is "_ = ?prod") proof - let ?sum = "(\(a, i)\as. (\(b, j)\as - {(a, i)}. b) * smult (of_nat (Suc i)) (pderiv a))" let ?single = "(\(a, i)\as. a)" let ?prd = "\ a i. (\(b, j)\as - {(a, i)}. b) * smult (of_nat (Suc i)) (pderiv a)" let ?onederiv = "\(a, i)\as. ?prd a i" have pp: "pderiv p = ?prod * ?sum" by (rule pderiv_exp_prod) have p: "p = ?prod * ?single" by (rule poly_exp_expand) have monic: "monic ?prod" by (rule monic_Prod) have gcd: "coprime ?single ?onederiv" by (rule coprime_generic, auto) then have gcd: "gcd ?single ?onederiv = 1" by simp show ?thesis unfolding pp unfolding p poly_gcd_monic_factor [OF monic] gcd by simp qed lemma p_div_gcd_p_pderiv: "p div (gcd p (pderiv p)) = (\(a, i)\as. a)" unfolding pderiv_exp_gcd unfolding poly_exp_expand by (rule nonzero_mult_div_cancel_left, insert monic_Prod, auto) fun A B C D :: "nat \ 'a poly" where "A n = gcd (B n) (D n)" | "B 0 = p div (gcd p (pderiv p))" | "B (Suc n) = B n div A n" | "C 0 = pderiv p div (gcd p (pderiv p))" | "C (Suc n) = D n div A n" | "D n = C n - pderiv (B n)" lemma A_B_C_D: "A n = (\ (a, i) \ as \ UNIV \ {n}. a)" "B n = (\ (a, i) \ as - UNIV \ {0 ..< n}. a)" "C n = (\(a, i)\as - UNIV \ {0 ..< n}. (\(b, j)\as - UNIV \ {0 ..< n} - {(a, i)}. b) * smult (of_nat (Suc i - n)) (pderiv a))" "D n = (\ (a, i) \ as \ UNIV \ {n}. a) * (\ (a,i)\as - UNIV \ {0 ..< Suc n}. (\(b, j)\ as - UNIV \ {0 ..< Suc n} - {(a, i)}. b) * (smult (of_nat (i - n)) (pderiv a)))" proof (induct n and n and n and n rule: A_B_C_D.induct) case (1 n) (* A *) note Bn = 1(1) note Dn = 1(2) have "(\(a, i)\as - UNIV \ {0..< n}. a) = (\(a, i)\as \ UNIV \ {n}. a) * (\(a, i)\as - UNIV \ {0.. 0" by auto let ?dn = "(\ (a,i)\as - UNIV \ {0 ..< Suc n}. (\(b, j)\ as - UNIV \ {0 ..< Suc n} - {(a, i)}. b) * (smult (of_nat (i - n)) (pderiv a)))" have Dn: "D n = ?an * ?dn" unfolding Dn by auto show dvd1: "?an dvd B n" unfolding Bn' dvd_def by blast show dvd2: "?an dvd D n" unfolding Dn dvd_def by blast { fix k assume "k dvd B n" "k dvd D n" from dvd_gcd_mult[OF this[unfolded Bn' Dn]] have "k dvd ?an * (gcd ?bn ?dn)" . moreover have "coprime ?bn ?dn" by (rule coprime_generic, auto) ultimately show "k dvd ?an" by simp } qed auto next case 2 (* B 0 *) have as: "as - UNIV \ {0..<0} = as" by auto show ?case unfolding B.simps as p_div_gcd_p_pderiv by auto next case (3 n) (* B n *) have id: "(\(a, i)\as - UNIV \ {0..< n}. a) = (\(a, i)\as - UNIV \ {0..(a, i)\as \ UNIV \ {n}. a)" by (subst prod.union_disjoint[symmetric], auto, insert fin, auto intro: prod.cong) show ?case unfolding B.simps 3 id by (subst nonzero_mult_div_cancel_right[OF nonzero_gen], auto) next case 4 (* C 0 *) have as: "as - UNIV \ {0..<0} = as" "\ i. Suc i - 0 = Suc i" by auto show ?case unfolding C.simps pderiv_exp_gcd unfolding pderiv_exp_prod as by (rule nonzero_mult_div_cancel_left, insert monic_Prod, auto) next case (5 n) (* C n *) show ?case unfolding C.simps 5 by (subst nonzero_mult_div_cancel_left, rule nonzero_gen, auto) next case (6 n) (* D n *) let ?f = "\ (a,i). (\(b, j)\as - UNIV \ {0 ..< n} - {(a, i)}. b) * (smult (of_nat (i - n)) (pderiv a))" have "D n = (\ (a,i)\as - UNIV \ {0 ..< n}. (\(b, j)\as - UNIV \ {0 ..< n} - {(a, i)}. b) * (smult (of_nat (Suc i - n)) (pderiv a) - pderiv a))" unfolding D.simps 6 pderiv_prod sum_subtractf[symmetric] right_diff_distrib by (rule sum.cong, auto) also have "\ = sum ?f (as - UNIV \ {0 ..< n})" proof (rule sum.cong[OF refl]) fix x assume "x \ as - UNIV \ {0 ..< n}" then obtain a i where x: "x = (a,i)" and i: "Suc i > n" by (cases x, auto) hence id: "Suc i - n = Suc (i - n)" by arith have id: "of_nat (Suc i - n) = of_nat (i - n) + (1 :: 'a)" unfolding id by simp have id: "smult (of_nat (Suc i - n)) (pderiv a) - pderiv a = smult (of_nat (i - n)) (pderiv a)" unfolding id smult_add_left by auto have cong: "\ x y z :: 'a poly. x = y \ x * z = y * z" by auto show "(case x of (a, i) \ (\(b, j)\as - UNIV \ {0.. (\(b, j)\as - UNIV \ {0.. = sum ?f (as - UNIV \ {0 ..< Suc n}) + sum ?f (as \ UNIV \ {n})" by (subst sum.union_disjoint[symmetric], insert fin, auto intro: sum.cong) also have "sum ?f (as \ UNIV \ {n}) = 0" by (rule sum.neutral, auto) finally have id: "D n = sum ?f (as - UNIV \ {0 ..< Suc n})" by simp show ?case unfolding id sum_distrib_left proof (rule sum.cong[OF refl]) fix x assume mem: "x \ as - UNIV \ {0 ..< Suc n}" obtain a i where x: "x = (a,i)" by force with mem have i: "i > n" by auto have cong: "\ x y z v :: 'a poly. x = y * v \ x * z = y * (v * z)" by auto show "(case x of (a, i) \ (\(b, j)\as - UNIV \ {0..(a, i)\as \ UNIV \ {n}. a) * (case x of (a, i) \ (\(b, j)\as - UNIV \ {0..i = 0..< n. A i ^ Suc i) = (\(a, i)\ as \ UNIV \ {0 ..< n}. a ^ Suc i)" proof (induct n) case (Suc n) have id: "{0 ..< Suc n} = insert n {0 ..< n}" by auto have id2: "as \ UNIV \ {0 ..< Suc n} = as \ UNIV \ {n} \ as \ UNIV \ {0 ..< n}" by auto have cong: "\ x y z. x = y \ x * z = y * z" by auto show ?case unfolding id2 unfolding id proof (subst prod.insert; (subst prod.union_disjoint)?; (unfold Suc)?; (unfold A, rule cong)?) show "(\(a, i)\as \ UNIV \ {n}. a) ^ Suc n = (\(a, i)\as \ UNIV \ {n}. a ^ Suc i)" unfolding prod_power_distrib by (rule prod.cong, auto) qed (insert fin, auto) qed simp lemma prod_A_is_p_unknown: assumes "\ a i. (a,i) \ as \ i < n" shows "p = (\i = 0..< n. A i ^ Suc i)" proof - have "p = (\(a, i)\as. a ^ Suc i)" by (rule p) also have "\ = (\i = 0..< n. A i ^ Suc i)" unfolding prod_A by (rule prod.cong, insert assms, auto) finally show ?thesis . qed definition bound :: nat where "bound = Suc (Max (snd ` as))" lemma bound: assumes m: "m \ bound" shows "B m = 1" proof - let ?set = "as - UNIV \ {0.. ?set" hence "i \ snd ` as" by force from Max_ge[OF _ this] fin have "i \ Max (snd ` as)" by auto with ai m[unfolded bound_def] have False by auto } hence id: "?set = {}" by force show "B m = 1" unfolding B id by simp qed lemma coprime_A_A: assumes "i \ j" shows "coprime (A i) (A j)" proof (rule coprimeI) fix k assume dvd: "k dvd A i" "k dvd A j" have Ai: "A i \ 0" unfolding A by (rule nonzero_gen, auto) with dvd have k: "k \ 0" by auto show "is_unit k" proof (cases "degree k > 0") case False then obtain c where kc: "k = [: c :]" by (auto dest: degree0_coeffs) with k have "1 = k * [:1 / c:]" by simp thus ?thesis unfolding dvd_def by blast next case True from irreducible_monic_factor[OF this] obtain q r where k: "k = q * r" and q: "irreducible q" and mq: "monic q" by auto with dvd have dvd: "q dvd A i" "q dvd A j" unfolding dvd_def by auto from q have q0: "degree q > 0" unfolding irreducible\<^sub>d_def by auto from irreducible_dvd_prod[OF q dvd(1)[unfolded A]] obtain a where ai: "(a,i) \ as" and qa: "q dvd a" by auto from irreducible_dvd_prod[OF q dvd(2)[unfolded A]] obtain b where bj: "(b,j) \ as" and qb: "q dvd b" by auto from as_distinct[OF ai bj] assms have neq: "a \ b" by auto from irreducible\<^sub>d_dvd_smult[OF q0 as_irred[OF ai] qa] irreducible\<^sub>d_dvd_smult[OF q0 as_irred[OF bj] qb] obtain c d where "c \ 0" "d \ 0" "a = smult c q" "b = smult d q" by auto hence ab: "a = smult (c / d) b" and "c / d \ 0" by auto with as_monic[OF bj] as_monic[OF ai] arg_cong[OF ab, of "\ p. coeff p (degree p)"] have "a = b" unfolding coeff_smult degree_smult_eq by auto with neq show ?thesis by auto qed qed lemma A_monic: "monic (A i)" unfolding A by (rule monic_gen, auto) lemma A_square_free: "square_free (A i)" proof (rule square_freeI) fix q k have mon: "monic (A i)" by (rule A_monic) hence Ai: "A i \ 0" by auto assume q: "degree q > 0" and dvd: "q * q dvd A i" from irreducible_monic_factor[OF q] obtain r s where q: "q = r * s" and irr: "irreducible r" and mr: "monic r" by auto from dvd[unfolded q] have dvd2: "r * r dvd A i" and dvd1: "r dvd A i" unfolding dvd_def by auto from irreducible_dvd_prod[OF irr dvd1[unfolded A]] obtain a where ai: "(a,i) \ as" and ra: "r dvd a" by auto let ?rem = "(\(a, i)\as \ UNIV \ {i} - {(a,i)}. a)" have a: "irreducible\<^sub>d a" by (rule as_irred[OF ai]) from irreducible\<^sub>d_dvd_smult[OF _ a ra] irr obtain c where ar: "a = smult c r" and "c \ 0" by force with mr as_monic[OF ai] arg_cong[OF ar, of "\ p. coeff p (degree p)"] have "a = r" unfolding coeff_smult degree_smult_eq by auto with dvd2 have dvd: "a * a dvd A i" by simp have id: "A i = a * ?rem" unfolding A by (subst prod.remove[of _ "(a,i)"], insert ai fin, auto) with dvd have "a dvd ?rem" using a id Ai by auto from irreducible_dvd_prod[OF _ this] a obtain b where bi: "(b,i) \ as" and neq: "b \ a" and ab: "a dvd b" by auto from as_irred[OF bi] have b: "irreducible\<^sub>d b" . from irreducible\<^sub>d_dvd_smult[OF _ b ab] a[unfolded irreducible\<^sub>d_def] obtain c where "c \ 0" and ba: "b = smult c a" by auto with as_monic[OF bi] as_monic[OF ai] arg_cong[OF ba, of "\ p. coeff p (degree p)"] have "a = b" unfolding coeff_smult degree_smult_eq by auto with neq show False by auto qed (insert A_monic[of i], auto) lemma prod_A_is_p_B_bound: assumes "B n = 1" shows "p = (\i = 0..< n. A i ^ Suc i)" proof (rule prod_A_is_p_unknown) fix a i assume ai: "(a,i) \ as" let ?rem = "(\(a, i)\as - UNIV \ {0.. 0" by (rule nonzero_gen, auto) have "irreducible\<^sub>d a" using as_irred[OF ai] . hence a: "a \ 0" "degree a \ 0" unfolding irreducible\<^sub>d_def by auto show "i < n" proof (rule ccontr) assume "\ ?thesis" hence "i \ n" by auto with ai have mem: "(a,i) \ as - UNIV \ {0 ..< n}" by auto have "0 = degree (\(a, i)\as - UNIV \ {0.. = degree (a * ?rem)" by (subst prod.remove[OF _ mem], insert fin, auto) also have "\ = degree a + degree ?rem" by (rule degree_mult_eq[OF a(1) rem]) finally show False using a(2) by auto qed qed interpretation yun_gcd gcd . lemma square_free_monic_poly: "(poly (square_free_monic_poly p) x = 0) = (poly p x = 0)" proof - show ?thesis unfolding square_free_monic_poly_def unfolding p_div_gcd_p_pderiv unfolding p poly_prod prod_zero_iff[OF fin] by force qed lemma yun_factorization_induct: assumes base: "\ bn cn. bn = 1 \ P bn cn" and step: "\ bn cn. bn \ 1 \ P (bn div (gcd bn (cn - pderiv bn))) ((cn - pderiv bn) div (gcd bn (cn - pderiv bn))) \ P bn cn" and id: "bn = p div gcd p (pderiv p)" "cn = pderiv p div gcd p (pderiv p)" shows "P bn cn" proof - define n where "n = (0 :: nat)" let ?m = "\ n. bound - n" have "P (B n) (C n)" proof (induct n rule: wf_induct[OF wf_measure[of ?m]]) case (1 n) note IH = 1(1)[rule_format] show ?case proof (cases "B n = 1") case True with base show ?thesis by auto next case False note Bn = this with bound[of n] have "\ bound \ n" by auto hence "(Suc n, n) \ measure ?m" by auto note IH = IH[OF this] show ?thesis by (rule step[OF Bn], insert IH, simp add: D.simps C.simps B.simps A.simps) qed qed thus ?thesis unfolding id n_def B.simps C.simps . qed lemma yun_factorization_main: assumes "yun_factorization_main (B n) (C n) n bs = cs" "set bs = {(A i, i) | i. i < n}" "distinct (map snd bs)" shows "\ m. set cs = {(A i, i) | i. i < m} \ B m = 1 \ distinct (map snd cs)" using assms proof - let ?m = "\ n. bound - n" show ?thesis using assms proof (induct n arbitrary: bs rule: wf_induct[OF wf_measure[of ?m]]) case (1 n) note IH = 1(1)[rule_format] have res: "yun_factorization_main (B n) (C n) n bs = cs" by fact note res = res[unfolded yun_factorization_main.simps[of "B n"]] have bs: "set bs = {(A i, i) |i. i < n}" "distinct (map snd bs)" by fact+ show ?case proof (cases "B n = 1") case True with res have "bs = cs" by auto with True bs show ?thesis by auto next case False note Bn = this with bound[of n] have "\ bound \ n" by auto hence "(Suc n, n) \ measure ?m" by auto note IH = IH[OF this] from Bn res[unfolded Let_def, folded D.simps C.simps B.simps A.simps] have res: "yun_factorization_main (B (Suc n)) (C (Suc n)) (Suc n) ((A n, n) # bs) = cs" by simp note IH = IH[OF this] { fix i assume "i < Suc n" "\ i < n" hence "n = i" by arith } note missing = this have "set ((A n, n) # bs) = {(A i, i) |i. i < Suc n}" unfolding list.simps bs by (auto, subst missing, auto) note IH = IH[OF this] from bs have "distinct (map snd ((A n, n) # bs))" by auto note IH = IH[OF this] show ?thesis by (rule IH) qed qed qed lemma yun_monic_factorization_res: assumes res: "yun_monic_factorization p = bs" shows "\ m. set bs = {(A i, i) | i. i < m \ A i \ 1} \ B m = 1 \ distinct (map snd bs)" proof - from res[unfolded yun_monic_factorization_def Let_def, folded B.simps C.simps] obtain cs where yun: "yun_factorization_main (B 0) (C 0) 0 [] = cs" and bs: "bs = filter (\ (a,i). a \ 1) cs" by auto from yun_factorization_main[OF yun] show ?thesis unfolding bs by (auto simp: distinct_map_filter) qed lemma yun_monic_factorization: assumes yun: "yun_monic_factorization p = bs" shows "square_free_factorization p (1,bs)" "(b,i) \ set bs \ monic b" "distinct (map snd bs)" proof - from yun_monic_factorization_res[OF yun] obtain m where bs: "set bs = {(A i, i) | i. i < m \ A i \ 1}" and B: "B m = 1" and dist: "distinct (map snd bs)" by auto have id: "{0 ..< m} = {i. i < m \ A i = 1} \ {i. i < m \ A i \ 1}" (is "_ = ?ignore \ _") by auto have "p = (\i = 0.. = prod (\ i. A i ^ Suc i) {i. i < m \ A i \ 1}" unfolding id by (subst prod.union_disjoint, (force+)[3], subst prod.neutral[of ?ignore], auto) also have "\ = (\(a, i)\ set bs. a ^ Suc i)" unfolding bs by (rule prod.reindex_cong[of snd], auto simp: inj_on_def, force) finally have 1: "p = (\(a, i)\ set bs. a ^ Suc i)" . { fix a i assume "(a,i) \ set bs" hence A: "a = A i" "A i \ 1" unfolding bs by auto with A_square_free[of i] A_monic[of i] have "square_free a \ degree a \ 0" "monic a" by (auto simp: monic_degree_0) } note 2 = this { fix a i b j assume ai: "(a,i) \ set bs" and bj: "(b,j) \ set bs" and neq: "(a,i) \ (b,j)" hence a: "a = A i" and b: "b = A j" unfolding bs by auto from neq dist ai bj have neq: "i \ j" using a b by blast from coprime_A_A [OF neq] have "coprime a b" unfolding a b . } note 3 = this have "monic p" unfolding p by (rule monic_prod, insert as_monic, auto intro: monic_power monic_mult) hence 4: "p \ 0" by auto from dist have 5: "distinct bs" unfolding distinct_map .. show "square_free_factorization p (1,bs)" unfolding square_free_factorization_def using 1 2 3 4 5 by auto show "(b,i) \ set bs \ monic b" using 2 by auto show "distinct (map snd bs)" by fact qed end lemma monic_factorization: assumes "monic p" shows "\ as. monic_factorization as p" proof - from monic_irreducible_factorization[OF assms] obtain as f where fin: "finite as" and p: "p = (\a\as. a ^ Suc (f a))" and as: "as \ {q. irreducible\<^sub>d q \ monic q}" by auto define cs where "cs = {(a, f a) | a. a \ as}" show ?thesis proof (rule exI, standard) show "finite cs" unfolding cs_def using fin by auto { fix a i assume "(a,i) \ cs" thus "irreducible\<^sub>d a" "monic a" unfolding cs_def using as by auto } note irr = this show "\a i b j. (a, i) \ cs \ (b, j) \ cs \ (a, i) \ (b, j) \ a \ b" unfolding cs_def by auto show "p = (\(a, i)\cs. a ^ Suc i)" unfolding p cs_def by (rule prod.reindex_cong, auto, auto simp: inj_on_def) qed qed -lemma square_free_monic_poly: assumes "monic (p :: 'a :: {field_char_0, euclidean_ring_gcd} poly)" +lemma square_free_monic_poly: + assumes "monic (p :: 'a :: {field_char_0, euclidean_ring_gcd,semiring_gcd_mult_normalize} poly)" shows "(poly (yun_gcd.square_free_monic_poly gcd p) x = 0) = (poly p x = 0)" proof - from monic_factorization[OF assms] obtain as where "monic_factorization as p" .. from monic_factorization.square_free_monic_poly[OF this] show ?thesis . qed -lemma yun_factorization_induct: assumes base: "\ bn cn. bn = 1 \ P bn cn" +lemma yun_factorization_induct: + assumes base: "\ bn cn. bn = 1 \ P bn cn" and step: "\ bn cn. bn \ 1 \ P (bn div (gcd bn (cn - pderiv bn))) ((cn - pderiv bn) div (gcd bn (cn - pderiv bn))) \ P bn cn" and id: "bn = p div gcd p (pderiv p)" "cn = pderiv p div gcd p (pderiv p)" - and monic: "monic (p :: 'a :: {field_char_0,euclidean_ring_gcd} poly)" + and monic: "monic (p :: 'a :: {field_char_0,euclidean_ring_gcd,semiring_gcd_mult_normalize} poly)" shows "P bn cn" proof - from monic_factorization[OF monic] obtain as where "monic_factorization as p" .. from monic_factorization.yun_factorization_induct[OF this base step id] show ?thesis . qed lemma square_free_poly: "(poly (square_free_poly gcd p) x = 0) = (poly p x = 0)" proof (cases "p = 0") case True thus ?thesis unfolding square_free_poly_def by auto next case False let ?c = "coeff p (degree p)" let ?ic = "inverse ?c" have id: "square_free_poly gcd p = yun_gcd.square_free_monic_poly gcd (smult ?ic p)" unfolding square_free_poly_def using False by auto from False have mon: "monic (smult ?ic p)" and ic: "?ic \ 0" by auto show ?thesis unfolding id square_free_monic_poly[OF mon] using ic by simp qed -lemma yun_monic_factorization: fixes p :: "'a :: {field_char_0,euclidean_ring_gcd} poly" +lemma yun_monic_factorization: + fixes p :: "'a :: {field_char_0,euclidean_ring_gcd,semiring_gcd_mult_normalize} poly" assumes res: "yun_gcd.yun_monic_factorization gcd p = bs" and monic: "monic p" shows "square_free_factorization p (1,bs)" "(b,i) \ set bs \ monic b" "distinct (map snd bs)" proof - from monic_factorization[OF monic] obtain as where "monic_factorization as p" .. from "monic_factorization.yun_monic_factorization"[OF this res] show "square_free_factorization p (1,bs)" "(b,i) \ set bs \ monic b" "distinct (map snd bs)" by auto qed lemma square_free_factorization_smult: assumes c: "c \ 0" and sf: "square_free_factorization p (d,bs)" shows "square_free_factorization (smult c p) (c * d, bs)" proof - from sf[unfolded square_free_factorization_def split] have p: "p = smult d (\(a, i)\set bs. a ^ Suc i)" and eq: "p = 0 \ d = 0 \ bs = []" by blast+ from eq c have eq: "smult c p = 0 \ c * d = 0 \ bs = []" by auto from p have p: "smult c p = smult (c * d) (\(a, i)\set bs. a ^ Suc i)" by auto from eq p sf show ?thesis unfolding square_free_factorization_def by blast qed lemma yun_factorization: assumes res: "yun_factorization gcd p = c_bs" shows "square_free_factorization p c_bs" "(b,i) \ set (snd c_bs) \ monic b" proof - interpret yun_gcd gcd . note res = res[unfolded yun_factorization_def Let_def] have "square_free_factorization p c_bs \ ((b,i) \ set (snd c_bs) \ monic b)" proof (cases "p = 0") case True with res have "c_bs = (0, [])" by auto thus ?thesis unfolding True by (auto simp: square_free_factorization_def) next case False let ?c = "coeff p (degree p)" let ?ic = "inverse ?c" obtain c bs where cbs: "c_bs = (c,bs)" by force with False res have c: "c = ?c" "?c \ 0" and fact: "yun_monic_factorization (smult ?ic p) = bs" by auto from False have mon: "monic (smult ?ic p)" by auto from yun_monic_factorization[OF fact mon] have sff: "square_free_factorization (smult ?ic p) (1, bs)" "(b, i) \ set bs \ monic b" by auto have id: "smult ?c (smult ?ic p) = p" using False by auto from square_free_factorization_smult[OF c(2) sff(1), unfolded id] sff show ?thesis unfolding cbs c by simp qed thus "square_free_factorization p c_bs" "(b,i) \ set (snd c_bs) \ monic b" by blast+ qed lemma prod_list_pow_suc: "(\x\bs. (x :: 'a :: comm_monoid_mult) * x ^ i) = prod_list bs * prod_list bs ^ i" by (induct bs, auto simp: field_simps) declare irreducible_linear_field_poly[intro!] context - assumes "SORT_CONSTRAINT('a :: {field, factorial_ring_gcd})" + assumes "SORT_CONSTRAINT('a :: {field, factorial_ring_gcd,semiring_gcd_mult_normalize})" begin lemma square_free_factorization_order_root_mem: assumes sff: "square_free_factorization p (c,bs)" and p: "p \ (0 :: 'a poly)" and ai: "(a,i) \ set bs" and rt: "poly a x = 0" shows "order x p = Suc i" proof - note sff = square_free_factorizationD[OF sff] let ?prod = "(\(a, i)\set bs. a ^ Suc i)" from sff have pf: "p = smult c ?prod" by blast with p have c: "c \ 0" by auto have ord: "order x p = order x ?prod" unfolding pf using order_smult[OF c] by auto define q where "q = [: -x, 1 :]" have q0: "q \ 0" unfolding q_def by auto have iq: "irreducible q" by (auto simp: q_def) from rt have qa: "q dvd a" unfolding q_def poly_eq_0_iff_dvd . then obtain b where aqb: "a = q * b" unfolding dvd_def by auto from sff(2)[OF ai] have sq: "square_free a" and mon: "degree a \ 0" by auto let ?rem = "(\(a, i)\set bs - {(a,i)}. a ^ Suc i)" have p0: "?prod \ 0" using p pf by auto have "?prod = a ^ Suc i * ?rem" by (subst prod.remove[OF _ ai], auto) also have "a ^ Suc i = q ^ Suc i * b ^ Suc i" unfolding aqb by (simp add: field_simps) finally have id: "?prod = q ^ Suc i * (b ^ Suc i * ?rem)" by simp hence dvd: "q ^ Suc i dvd ?prod" by auto { assume "q ^ Suc (Suc i) dvd ?prod" hence "q dvd ?prod div q ^ Suc i" by (metis dvd dvd_0_left_iff dvd_div_iff_mult p0 power_Suc) also have "?prod div q ^ Suc i = b ^ Suc i * ?rem" unfolding id by (rule nonzero_mult_div_cancel_left, insert q0, auto) finally have "q dvd b \ q dvd ?rem" using iq irreducible_dvd_pow[OF iq] by auto hence False proof assume "q dvd b" with aqb have "q * q dvd a" by auto with sq[unfolded square_free_def] mon iq show False unfolding irreducible\<^sub>d_def by auto next assume "q dvd ?rem" from irreducible_dvd_prod[OF iq this] obtain b j where bj: "(b,j) \ set bs" and neq: "(a,i) \ (b,j)" and dvd: "q dvd b ^ Suc j" by auto from irreducible_dvd_pow[OF iq dvd] have qb: "q dvd b" . from sff(3)[OF ai bj neq] have gcd: "coprime a b" . from qb qa have "q dvd gcd a b" by simp from dvd_imp_degree_le[OF this[unfolded gcd]] iq q0 show False using gcd by auto qed } hence ndvd: "\ q ^ Suc (Suc i) dvd ?prod" by blast with dvd have "order x ?prod = Suc i" unfolding q_def by (rule order_unique_lemma[symmetric]) thus ?thesis unfolding ord . qed lemma square_free_factorization_order_root_no_mem: assumes sff: "square_free_factorization p (c,bs)" and p: "p \ (0 :: 'a poly)" and no_root: "\ a i. (a,i) \ set bs \ poly a x \ 0" shows "order x p = 0" proof (rule ccontr) assume o0: "order x p \ 0" with order_root[of p x] p have 0: "poly p x = 0" by auto note sff = square_free_factorizationD[OF sff] let ?prod = "(\(a, i)\set bs. a ^ Suc i)" from sff have pf: "p = smult c ?prod" by blast with p have c: "c \ 0" by auto with 0 have 0: "poly ?prod x = 0" unfolding pf by auto define q where "q = [: -x, 1 :]" from 0 have dvd: "q dvd ?prod" unfolding poly_eq_0_iff_dvd by (simp add: q_def) have q0: "q \ 0" unfolding q_def by auto have iq: "irreducible q" by (unfold q_def, auto intro:) from irreducible_dvd_prod[OF iq dvd] obtain a i where ai: "(a,i) \ set bs" and dvd: "q dvd a ^ Suc i" by auto from irreducible_dvd_pow[OF iq dvd] have dvd: "q dvd a" . hence "poly a x = 0" unfolding q_def by (simp add: poly_eq_0_iff_dvd q_def) with no_root[OF ai] show False by simp qed lemma square_free_factorization_order_root: assumes sff: "square_free_factorization p (c,bs)" and p: "p \ (0 :: 'a poly)" shows "order x p = i \ (i = 0 \ (\ a j. (a,j) \ set bs \ poly a x \ 0) \ (\ a j. (a,j) \ set bs \ poly a x = 0 \ i = Suc j))" (is "?l = (?r1 \ ?r2)") proof - note mem = square_free_factorization_order_root_mem[OF sff p] note no_mem = square_free_factorization_order_root_no_mem[OF sff p] show ?thesis proof assume "?r1 \ ?r2" thus ?l proof assume ?r2 then obtain a j where aj: "(a,j) \ set bs" "poly a x = 0" and i: "i = Suc j" by auto from mem[OF aj] i show ?l by simp next assume ?r1 with no_mem[of x] show ?l by auto qed next assume ?l show "?r1 \ ?r2" proof (cases "\a j. (a, j) \ set bs \ poly a x = 0") case True then obtain a j where "(a, j) \ set bs" "poly a x = 0" by auto with mem[OF this] \?l\ have ?r2 by auto thus ?thesis .. next case False with no_mem[of x] \?l\ have ?r1 by auto thus ?thesis .. qed qed qed lemma square_free_factorization_root: assumes sff: "square_free_factorization p (c,bs)" and p: "p \ (0 :: 'a poly)" shows "{x. poly p x = 0} = {x. \ a i. (a,i) \ set bs \ poly a x = 0}" using square_free_factorization_order_root[OF sff p] p unfolding order_root by auto lemma square_free_factorizationD': fixes p :: "'a poly" assumes sf: "square_free_factorization p (c, bs)" shows "p = smult c (\(a, i) \ bs. a ^ Suc i)" and "square_free (prod_list (map fst bs))" and "\ b i. (b,i) \ set bs \ degree b \ 0" and "p = 0 \ c = 0 \ bs = []" proof - note sf = square_free_factorizationD[OF sf] show "p = smult c (\(a, i) \ bs. a ^ Suc i)" unfolding sf(1) using sf(5) by (simp add: prod.distinct_set_conv_list) show bs: "\ b i. (b,i) \ set bs \ degree b \ 0" using sf(2) by auto show "p = 0 \ c = 0 \ bs = []" using sf(4) . show "square_free (prod_list (map fst bs))" proof (rule square_freeI) from bs have "\ b. b \ set (map fst bs) \ b \ 0" by fastforce thus "prod_list (map fst bs) \ 0" unfolding prod_list_zero_iff by auto fix q assume "degree q > 0" "q * q dvd prod_list (map fst bs)" from irreducible\<^sub>d_factor[OF this(1)] this(2) obtain q where irr: "irreducible q" and dvd: "q * q dvd prod_list (map fst bs)" unfolding dvd_def by auto hence dvd': "q dvd prod_list (map fst bs)" unfolding dvd_def by auto from irreducible_dvd_prod_list[OF irr dvd'] obtain b i where mem: "(b,i) \ set bs" and dvd1: "q dvd b" by auto from dvd1 obtain k where b: "b = q * k" unfolding dvd_def by auto from split_list[OF mem] b obtain bs1 bs2 where bs: "bs = bs1 @ (b, i) # bs2" by auto from irr have q0: "q \ 0" and dq: "degree q > 0" unfolding irreducible\<^sub>d_def by auto from sf(2)[OF mem, unfolded b] have "square_free (q * k)" by auto from this[unfolded square_free_def, THEN conjunct2, rule_format, OF dq] have qk: "\ q dvd k" by simp from dvd[unfolded bs b] have "q * q dvd q * (k * prod_list (map fst (bs1 @ bs2)))" by (auto simp: ac_simps) with q0 have "q dvd k * prod_list (map fst (bs1 @ bs2))" by auto with irr qk have "q dvd prod_list (map fst (bs1 @ bs2))" by auto from irreducible_dvd_prod_list[OF irr this] obtain b' i' where mem': "(b',i') \ set (bs1 @ bs2)" and dvd2: "q dvd b'" by fastforce from dvd1 dvd2 have "q dvd gcd b b'" by auto with dq is_unit_iff_degree[OF q0] have cop: "\ coprime b b'" by force from mem' have "(b',i') \ set bs" unfolding bs by auto from sf(3)[OF mem this] cop have b': "(b',i') = (b,i)" by (auto simp add: coprime_iff_gcd_eq_1) with mem' sf(5)[unfolded bs] show False by auto qed qed lemma square_free_factorizationI': fixes p :: "'a poly" assumes prod: "p = smult c (\(a, i) \ bs. a ^ Suc i)" and sf: "square_free (prod_list (map fst bs))" and deg: "\ b i. (b,i) \ set bs \ degree b > 0" and 0: "p = 0 \ c = 0 \ bs = []" shows "square_free_factorization p (c, bs)" unfolding square_free_factorization_def split proof (intro conjI impI allI) show "p = 0 \ c = 0" "p = 0 \ bs = []" using 0 by auto { fix b i assume bi: "(b,i) \ set bs" from deg[OF this] show "degree b > 0" . have "b dvd prod_list (map fst bs)" by (intro prod_list_dvd, insert bi, force) from square_free_factor[OF this sf] show "square_free b" . } show dist: "distinct bs" proof (rule ccontr) assume "\ ?thesis" from not_distinct_decomp[OF this] obtain bs1 bs2 bs3 b i where bs: "bs = bs1 @ [(b,i)] @ bs2 @ [(b,i)] @ bs3" by force hence "b * b dvd prod_list (map fst bs)" by auto with sf[unfolded square_free_def, THEN conjunct2, rule_format, of b] have db: "degree b = 0" by auto from bs have "(b,i) \ set bs" by auto from deg[OF this] db show False by auto qed show "p = smult c (\(a, i)\set bs. a ^ Suc i)" unfolding prod using dist by (simp add: prod.distinct_set_conv_list) { fix a i b j assume ai: "(a, i) \ set bs" and bj: "(b, j) \ set bs" and diff: "(a, i) \ (b, j)" from split_list[OF ai] obtain bs1 bs2 where bs: "bs = bs1 @ (a,i) # bs2" by auto with bj diff have "(b,j) \ set (bs1 @ bs2)" by auto from split_list[OF this] obtain cs1 cs2 where cs: "bs1 @ bs2 = cs1 @ (b,j) # cs2" by auto have "prod_list (map fst bs) = a * prod_list (map fst (bs1 @ bs2))" unfolding bs by simp also have "\ = a * b * prod_list (map fst (cs1 @ cs2))" unfolding cs by simp finally obtain c where lp: "prod_list (map fst bs) = a * b * c" by auto from deg[OF ai] have 0: "gcd a b \ 0" by auto have gcd: "gcd a b * gcd a b dvd prod_list (map fst bs)" unfolding lp by (simp add: mult_dvd_mono) { assume "degree (gcd a b) > 0" from sf[unfolded square_free_def, THEN conjunct2, rule_format, OF this] gcd have False by simp } hence "degree (gcd a b) = 0" by auto with 0 show "coprime a b" using is_unit_gcd is_unit_iff_degree by blast } qed lemma square_free_factorization_def': fixes p :: "'a poly" shows "square_free_factorization p (c,bs) \ (p = smult c (\(a, i) \ bs. a ^ Suc i)) \ (square_free (prod_list (map fst bs))) \ (\ b i. (b,i) \ set bs \ degree b > 0) \ (p = 0 \ c = 0 \ bs = [])" using square_free_factorizationD'[of p c bs] square_free_factorizationI'[of p c bs] by blast lemma square_free_factorization_smult_prod_listI: fixes p :: "'a poly" assumes sff: "square_free_factorization p (c, bs1 @ (smult b (prod_list bs),i) # bs2)" and bs: "\ b. b \ set bs \ degree b > 0" shows "square_free_factorization p (c * b^(Suc i), bs1 @ map (\ b. (b,i)) bs @ bs2)" proof - from square_free_factorizationD'(3)[OF sff, of "smult b (prod_list bs)" i] have b: "b \ 0" by auto note sff = square_free_factorizationD'[OF sff] show ?thesis proof (intro square_free_factorizationI', goal_cases) case 1 thus ?case unfolding sff(1) by (simp add: o_def field_simps smult_power prod_list_pow_suc) next case 2 show ?case using sff(2) by (simp add: ac_simps o_def square_free_smult_iff[OF b]) next case 3 with sff(3) bs show ?case by auto next case 4 from sff(4)[OF this] show ?case by simp qed qed lemma square_free_factorization_further_factorization: fixes p :: "'a poly" assumes sff: "square_free_factorization p (c, bs)" and bs: "\ b i d fs. (b,i) \ set bs \ f b = (d,fs) \ b = smult d (prod_list fs) \ (\ f \ set fs. degree f > 0)" and h: "h = (\ (b,i). case f b of (d,fs) \ (d^Suc i,map (\ f. (f,i)) fs))" and gs: "gs = map h bs" and d: "d = c * prod_list (map fst gs)" and es: "es = concat (map snd gs)" shows "square_free_factorization p (d, es)" proof - note sff = square_free_factorizationD'[OF sff] show ?thesis proof (rule square_free_factorizationI') assume "p = 0" from sff(4)[OF this] show "d = 0 \ es = []" unfolding d es gs by auto next have id: "(\(a, i)\bs. a * a ^ i) = smult (prod_list (map fst gs)) (\(a, i)\es. a * a ^ i)" unfolding es gs h map_map o_def using bs proof (induct bs) case (Cons bi bs) obtain b i where bi: "bi = (b,i)" by force obtain d fs where f: "f b = (d,fs)" by force from Cons(2)[OF _ f, of i] have b: "b = smult d (prod_list fs)" unfolding bi by auto note IH = Cons(1)[OF Cons(2), of "\ _ i _ _ . i"] show ?case unfolding bi by (simp add: f o_def, simp add: b ac_simps, subst IH, auto simp: smult_power prod_list_pow_suc ac_simps) qed simp show "p = smult d (\(a, i)\es. a ^ Suc i)" unfolding sff(1) using id by (simp add: d) next fix fi i assume fi: "(fi, i) \ set es" from this[unfolded es] obtain G where G: "G \ snd ` set gs" and fi: "(fi,i) \ set G" by auto from G[unfolded gs] obtain b i where bi: "(b,i) \ set bs" and G: "G = snd (h (b,i))" by auto obtain d fs where f: "f b = (d,fs)" by force show "degree fi > 0" by (rule bs[THEN conjunct2, rule_format, OF bi f], insert fi G f, unfold h, auto) next have id: "\ c. prod_list (map fst bs) = smult c (prod_list (map fst es))" unfolding es gs map_map o_def using bs proof (induct bs) case (Cons bi bs) obtain b i where bi: "bi = (b,i)" by force obtain d fs where f: "f b = (d,fs)" by force from Cons(2)[OF _ f, of i] have b: "b = smult d (prod_list fs)" unfolding bi by auto have "\c. prod_list (map fst bs) = smult c (prod_list (map fst (concat (map (\x. snd (h x)) bs))))" by (rule Cons(1), rule Cons(2), auto) then obtain c where IH: "prod_list (map fst bs) = smult c (prod_list (map fst (concat (map (\x. snd (h x)) bs))))" by auto show ?case unfolding bi by (intro exI[of _ "c * d"], auto simp: b IH, auto simp: h f[unfolded b] o_def) qed (intro exI[of _ 1], auto) then obtain c where "prod_list (map fst bs) = smult c (prod_list (map fst es))" by blast from sff(2)[unfolded this] show "square_free (prod_list (map fst es))" by (metis smult_eq_0_iff square_free_def square_free_smult_iff) qed qed lemma square_free_factorization_prod_listI: fixes p :: "'a poly" assumes sff: "square_free_factorization p (c, bs1 @ ((prod_list bs),i) # bs2)" and bs: "\ b. b \ set bs \ degree b > 0" shows "square_free_factorization p (c, bs1 @ map (\ b. (b,i)) bs @ bs2)" using square_free_factorization_smult_prod_listI[of p c bs1 1 bs i bs2] sff bs by auto lemma square_free_factorization_factorI: fixes p :: "'a poly" assumes sff: "square_free_factorization p (c, bs1 @ (a,i) # bs2)" and r: "degree r \ 0" and s: "degree s \ 0" and a: "a = r * s" shows "square_free_factorization p (c, bs1 @ ((r,i) # (s,i) # bs2))" using square_free_factorization_prod_listI[of p c bs1 "[r,s]" i bs2] sff r s a by auto end lemma monic_square_free_irreducible_factorization: assumes mon: "monic (f :: 'b :: field poly)" and sf: "square_free f" shows "\ P. finite P \ f = \P \ P \ {q. irreducible q \ monic q}" proof - from mon have f0: "f \ 0" by auto from monic_irreducible_factorization[OF assms(1)] obtain P n where P: "finite P" "P \ {q. irreducible\<^sub>d q \ monic q}" and f: "f = (\a\P. a ^ Suc (n a))" by auto have *: "\ a \ P. n a = 0" proof (rule ccontr) assume "\ ?thesis" then obtain a where a: "a \ P" and n: "n a \ 0" by auto have "f = a ^ (Suc (n a)) * (\b\P - {a}. b ^ Suc (n b))" unfolding f by (rule prod.remove[OF P(1) a]) with n have "a * a dvd f" by (cases "n a", auto) with sf[unfolded square_free_def] f0 have "degree a = 0" by auto with a P(2)[unfolded irreducible\<^sub>d_def] show False by auto qed have "f = \ P" unfolding f by (rule prod.cong[OF refl], insert *, auto) with P show ?thesis by auto qed context assumes "SORT_CONSTRAINT('a :: {field, factorial_ring_gcd})" begin lemma monic_factorization_uniqueness: fixes P::"'a poly set" assumes finite_P: "finite P" and PQ: "\P = \Q" and P: "P \ {q. irreducible\<^sub>d q \ monic q}" and finite_Q: "finite Q" and Q: "Q \ {q. irreducible\<^sub>d q \ monic q}" shows "P = Q" proof (rule; rule subsetI) fix x assume x: "x \ P" have irr_x: "irreducible x" using x P by auto then have "\a\Q. x dvd id a" proof (rule irreducible_dvd_prod) show "x dvd prod id Q" using PQ x by (metis dvd_refl dvd_prod finite_P id_apply prod.cong) qed from this obtain a where a: "a\Q" and x_dvd_a: "x dvd a" unfolding id_def by blast have "x=a" using x P a Q irreducible\<^sub>d_dvd_eq[OF _ _ x_dvd_a] by fast thus "x \ Q" using a by simp next fix x assume x: "x \ Q" have irr_x: "irreducible x" using x Q by auto then have "\a\P. x dvd id a" proof (rule irreducible_dvd_prod) show "x dvd prod id P" using PQ x by (metis dvd_refl dvd_prod finite_Q id_apply prod.cong) qed from this obtain a where a: "a\P" and x_dvd_a: "x dvd a" unfolding id_def by blast have "x=a" using x P a Q irreducible\<^sub>d_dvd_eq[OF _ _ x_dvd_a] by fast thus "x \ P" using a by simp qed end subsection \Yun factorization and homomorphisms\ locale field_hom_0' = field_hom hom - for hom :: "'a :: {field_char_0,euclidean_ring_gcd} \ 'b :: {field_char_0,euclidean_ring_gcd}" + for hom :: "'a :: {field_char_0,field_gcd} \ + 'b :: {field_char_0,field_gcd}" begin sublocale field_hom' .. end lemma (in field_hom_0') yun_factorization_main_hom: defines hp: "hp \ map_poly hom" defines hpi: "hpi \ map (\ (f,i). (hp f, i :: nat))" assumes monic: "monic p" and f: "f = p div gcd p (pderiv p)" and g: "g = pderiv p div gcd p (pderiv p)" shows "yun_gcd.yun_factorization_main gcd (hp f) (hp g) i (hpi as) = hpi (yun_gcd.yun_factorization_main gcd f g i as)" proof - let ?P = "\ f g. \ i as. yun_gcd.yun_factorization_main gcd (hp f) (hp g) i (hpi as) = hpi (yun_gcd.yun_factorization_main gcd f g i as)" note ind = yun_factorization_induct[OF _ _ f g monic, of ?P, rule_format] interpret map_poly_hom: map_poly_inj_comm_ring_hom.. interpret p: inj_comm_ring_hom hp unfolding hp.. note homs = map_poly_gcd[folded hp] map_poly_pderiv[folded hp] p.hom_minus map_poly_div[folded hp] show ?thesis proof (induct rule: ind) case (1 f g i as) show ?case unfolding yun_gcd.yun_factorization_main.simps[of _ "hp f"] yun_gcd.yun_factorization_main.simps[of _ f] unfolding 1 by simp next case (2 f g i as) have id: "\ f i fis. hpi ((f,i) # fis) = (hp f, i) # hpi fis" unfolding hpi by auto show ?case unfolding yun_gcd.yun_factorization_main.simps[of _ "hp f"] yun_gcd.yun_factorization_main.simps[of _ f] unfolding "p.hom_1_iff" unfolding Let_def unfolding homs[symmetric] id[symmetric] unfolding 2(2) by simp qed qed lemma square_free_square_free_factorization: - "square_free (p :: 'a :: {field,factorial_ring_gcd} poly) \ degree p \ 0 \ square_free_factorization p (1,[(p,0)])" + "square_free (p :: 'a :: {field,factorial_ring_gcd,semiring_gcd_mult_normalize} poly) \ + degree p \ 0 \ square_free_factorization p (1,[(p,0)])" by (intro square_free_factorizationI', auto) lemma constant_square_free_factorization: "degree p = 0 \ square_free_factorization p (coeff p 0,[])" by (drule degree0_coeffs [of p]) (auto simp: square_free_factorization_def) lemma (in field_hom_0') yun_monic_factorization: defines hp: "hp \ map_poly hom" defines hpi: "hpi \ map (\ (f,i). (hp f, i :: nat))" assumes monic: "monic f" shows "yun_gcd.yun_monic_factorization gcd (hp f) = hpi (yun_gcd.yun_monic_factorization gcd f)" proof - interpret map_poly_hom: map_poly_inj_comm_ring_hom.. interpret p: inj_ring_hom hp unfolding hp.. have hpiN: "hpi [] = []" unfolding hpi by simp obtain res where "res = yun_gcd.yun_factorization_main gcd (f div gcd f (pderiv f)) (pderiv f div gcd f (pderiv f)) 0 []" by auto note homs = map_poly_gcd[folded hp] map_poly_pderiv[folded hp] p.hom_minus map_poly_div[folded hp] yun_factorization_main_hom[folded hp, folded hpi, symmetric, OF monic refl refl, of _ Nil, unfolded hpiN] this show ?thesis unfolding yun_gcd.yun_monic_factorization_def Let_def unfolding homs[symmetric] unfolding hpi by (induct res, auto) qed lemma (in field_hom_0') yun_factorization_hom: defines hp: "hp \ map_poly hom" defines hpi: "hpi \ map (\ (f,i). (hp f, i :: nat))" shows "yun_factorization gcd (hp f) = map_prod hom hpi (yun_factorization gcd f)" using yun_monic_factorization[of "smult (inverse (coeff f (degree f))) f"] unfolding yun_factorization_def Let_def hp hpi by (auto simp: hom_distribs) lemma (in field_hom_0') square_free_map_poly: "square_free (map_poly hom f) = square_free f" proof - interpret map_poly_hom: map_poly_inj_comm_ring_hom.. show ?thesis unfolding square_free_iff_separable separable_def by (simp only: hom_distribs [symmetric] (*fold doesn't work!*)) (simp add: coprime_iff_gcd_eq_1 map_poly_gcd [symmetric]) qed end diff --git a/thys/Polynomial_Interpolation/Missing_Polynomial.thy b/thys/Polynomial_Interpolation/Missing_Polynomial.thy --- a/thys/Polynomial_Interpolation/Missing_Polynomial.thy +++ b/thys/Polynomial_Interpolation/Missing_Polynomial.thy @@ -1,1401 +1,1403 @@ (* Author: RenĂ© Thiemann Akihisa Yamada Jose Divason License: BSD *) section \Missing Polynomial\ text \The theory contains some basic results on polynomials which have not been detected in the distribution, especially on linear factors and degrees.\ theory Missing_Polynomial imports "HOL-Computational_Algebra.Polynomial_Factorial" Missing_Unsorted begin subsection \Basic Properties\ lemma degree_0_id: assumes "degree p = 0" shows "[: coeff p 0 :] = p" proof - have "\ x. 0 \ Suc x" by auto thus ?thesis using assms by (metis coeff_pCons_0 degree_pCons_eq_if pCons_cases) qed lemma degree0_coeffs: "degree p = 0 \ \ a. p = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) lemma degree1_coeffs: "degree p = 1 \ \ a b. p = [: b, a :] \ a \ 0" by (metis One_nat_def degree_pCons_eq_if nat.inject old.nat.distinct(2) pCons_0_0 pCons_cases) lemma degree2_coeffs: "degree p = 2 \ \ a b c. p = [: c, b, a :] \ a \ 0" by (metis Suc_1 Suc_neq_Zero degree1_coeffs degree_pCons_eq_if nat.inject pCons_cases) lemma poly_zero: fixes p :: "'a :: comm_ring_1 poly" assumes x: "poly p x = 0" shows "p = 0 \ degree p = 0" proof assume degp: "degree p = 0" hence "poly p x = coeff p (degree p)" by(subst degree_0_id[OF degp,symmetric], simp) hence "coeff p (degree p) = 0" using x by auto thus "p = 0" by auto qed auto lemma coeff_monom_Suc: "coeff (monom a (Suc d) * p) (Suc i) = coeff (monom a d * p) i" by (simp add: monom_Suc) lemma coeff_sum_monom: assumes n: "n \ d" shows "coeff (\i\d. monom (f i) i) n = f n" (is "?l = _") proof - have "?l = (\i\d. coeff (monom (f i) i) n)" (is "_ = sum ?cmf _") using coeff_sum. also have "{..d} = insert n ({..d}-{n})" using n by auto hence "sum ?cmf {..d} = sum ?cmf ..." by auto also have "... = sum ?cmf ({..d}-{n}) + ?cmf n" by (subst sum.insert,auto) also have "sum ?cmf ({..d}-{n}) = 0" by (subst sum.neutral, auto) finally show ?thesis by simp qed lemma linear_poly_root: "(a :: 'a :: comm_ring_1) \ set as \ poly (\ a \ as. [: - a, 1:]) a = 0" proof (induct as) case (Cons b as) show ?case proof (cases "a = b") case False with Cons have "a \ set as" by auto from Cons(1)[OF this] show ?thesis by simp qed simp qed simp lemma degree_lcoeff_sum: assumes deg: "degree (f q) = n" and fin: "finite S" and q: "q \ S" and degle: "\ p . p \ S - {q} \ degree (f p) < n" and cong: "coeff (f q) n = c" shows "degree (sum f S) = n \ coeff (sum f S) n = c" proof (cases "S = {q}") case True thus ?thesis using deg cong by simp next case False with q obtain p where "p \ S - {q}" by auto from degle[OF this] have n: "n > 0" by auto have "degree (sum f S) = degree (f q + sum f (S - {q}))" unfolding sum.remove[OF fin q] .. also have "\ = degree (f q)" proof (rule degree_add_eq_left) have "degree (sum f (S - {q})) \ n - 1" proof (rule degree_sum_le) fix p show "p \ S - {q} \ degree (f p) \ n - 1" using degle[of p] by auto qed (insert fin, auto) also have "\ < n" using n by simp finally show "degree (sum f (S - {q})) < degree (f q)" unfolding deg . qed finally show ?thesis unfolding deg[symmetric] cong[symmetric] proof (rule conjI) have id: "(\x\S - {q}. coeff (f x) (degree (f q))) = 0" by (rule sum.neutral, rule ballI, rule coeff_eq_0[OF degle[folded deg]]) show "coeff (sum f S) (degree (f q)) = coeff (f q) (degree (f q))" unfolding coeff_sum by (subst sum.remove[OF _ q], unfold id, insert fin, auto) qed qed lemma degree_sum_list_le: "(\ p . p \ set ps \ degree p \ n) \ degree (sum_list ps) \ n" proof (induct ps) case (Cons p ps) hence "degree (sum_list ps) \ n" "degree p \ n" by auto thus ?case unfolding sum_list.Cons by (metis degree_add_le) qed simp lemma degree_prod_list_le: "degree (prod_list ps) \ sum_list (map degree ps)" proof (induct ps) case (Cons p ps) show ?case unfolding prod_list.Cons by (rule order.trans[OF degree_mult_le], insert Cons, auto) qed simp lemma smult_sum: "smult (\i \ S. f i) p = (\i \ S. smult (f i) p)" by (induct S rule: infinite_finite_induct, auto simp: smult_add_left) lemma range_coeff: "range (coeff p) = insert 0 (set (coeffs p))" by (metis nth_default_coeffs_eq range_nth_default) lemma smult_power: "(smult a p) ^ n = smult (a ^ n) (p ^ n)" by (induct n, auto simp: field_simps) lemma poly_sum_list: "poly (sum_list ps) x = sum_list (map (\ p. poly p x) ps)" by (induct ps, auto) lemma poly_prod_list: "poly (prod_list ps) x = prod_list (map (\ p. poly p x) ps)" by (induct ps, auto) lemma sum_list_neutral: "(\ x. x \ set xs \ x = 0) \ sum_list xs = 0" by (induct xs, auto) lemma prod_list_neutral: "(\ x. x \ set xs \ x = 1) \ prod_list xs = 1" by (induct xs, auto) lemma (in comm_monoid_mult) prod_list_map_remove1: "x \ set xs \ prod_list (map f xs) = f x * prod_list (map f (remove1 x xs))" by (induct xs) (auto simp add: ac_simps) lemma poly_as_sum: fixes p :: "'a::comm_semiring_1 poly" shows "poly p x = (\i\degree p. x ^ i * coeff p i)" unfolding poly_altdef by (simp add: ac_simps) lemma poly_prod_0: "finite ps \ poly (prod f ps) x = (0 :: 'a :: field) \ (\ p \ ps. poly (f p) x = 0)" by (induct ps rule: finite_induct, auto) lemma coeff_monom_mult: shows "coeff (monom a d * p) i = (if d \ i then a * coeff p (i-d) else 0)" (is "?l = ?r") proof (cases "d \ i") case False thus ?thesis unfolding coeff_mult by simp next case True let ?f = "\j. coeff (monom a d) j * coeff p (i - j)" have "\j. j \ {0..i} - {d} \ ?f j = 0" by auto hence "0 = (\j \ {0..i} - {d}. ?f j)" by auto also have "... + ?f d = (\j \ insert d ({0..i} - {d}). ?f j)" by(subst sum.insert, auto) also have "... = (\j \ {0..i}. ?f j)" by (subst insert_Diff, insert True, auto) also have "... = (\j\i. ?f j)" by (rule sum.cong, auto) also have "... = ?l" unfolding coeff_mult .. finally show ?thesis using True by auto qed lemma poly_eqI2: assumes "degree p = degree q" and "\i. i \ degree p \ coeff p i = coeff q i" shows "p = q" apply(rule poly_eqI) by (metis assms le_degree) text \A nice extension rule for polynomials.\ lemma poly_ext[intro]: fixes p q :: "'a :: {ring_char_0, idom} poly" assumes "\x. poly p x = poly q x" shows "p = q" unfolding poly_eq_poly_eq_iff[symmetric] using assms by (rule ext) text \Copied from non-negative variants.\ lemma coeff_linear_power_neg[simp]: fixes a :: "'a::comm_ring_1" shows "coeff ([:a, -1:] ^ n) n = (-1)^n" apply (induct n, simp_all) apply (subst coeff_eq_0) apply (auto intro: le_less_trans degree_power_le) done lemma degree_linear_power_neg[simp]: fixes a :: "'a::{idom,comm_ring_1}" shows "degree ([:a, -1:] ^ n) = n" apply (rule order_antisym) apply (rule ord_le_eq_trans [OF degree_power_le], simp) apply (rule le_degree) unfolding coeff_linear_power_neg apply (auto) done subsection \Polynomial Composition\ lemmas [simp] = pcompose_pCons lemma pcompose_eq_0: fixes q :: "'a :: idom poly" assumes q: "degree q \ 0" shows "p \\<^sub>p q = 0 \ p = 0" proof (induct p) case 0 show ?case by auto next case (pCons a p) have id: "(pCons a p) \\<^sub>p q = [:a:] + q * (p \\<^sub>p q)" by simp show ?case proof (cases "p = 0") case True show ?thesis unfolding id unfolding True by simp next case False with pCons(2) have "p \\<^sub>p q \ 0" by auto from degree_mult_eq[OF _ this, of q] q have "degree (q * (p \\<^sub>p q)) \ 0" by force hence deg: "degree ([:a:] + q * (p \\<^sub>p q)) \ 0" by (subst degree_add_eq_right, auto) show ?thesis unfolding id using False deg by auto qed qed declare degree_pcompose[simp] subsection \Monic Polynomials\ abbreviation monic where "monic p \ coeff p (degree p) = 1" lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {field,normalization_semidom}) = x" by (cases "is_unit x") (auto simp: is_unit_unit_factor dvd_field_iff) lemma poly_gcd_monic: - fixes p :: "'a :: {field,factorial_ring_gcd} poly" + fixes p :: "'a :: {field,factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes "p \ 0 \ q \ 0" shows "monic (gcd p q)" proof - from assms have "1 = unit_factor (gcd p q)" by (auto simp: unit_factor_gcd) also have "\ = [:lead_coeff (gcd p q):]" unfolding unit_factor_poly_def by (simp add: monom_0) finally show ?thesis by (metis coeff_pCons_0 degree_1 lead_coeff_1) qed lemma normalize_monic: "monic p \ normalize p = p" by (simp add: normalize_poly_eq_map_poly is_unit_unit_factor) lemma lcoeff_monic_mult: assumes monic: "monic (p :: 'a :: comm_semiring_1 poly)" shows "coeff (p * q) (degree p + degree q) = coeff q (degree q)" proof - let ?pqi = "\ i. coeff p i * coeff q (degree p + degree q - i)" have "coeff (p * q) (degree p + degree q) = (\i\degree p + degree q. ?pqi i)" unfolding coeff_mult by simp also have "\ = ?pqi (degree p) + (sum ?pqi ({.. degree p + degree q} - {degree p}))" by (subst sum.remove[of _ "degree p"], auto) also have "?pqi (degree p) = coeff q (degree q)" unfolding monic by simp also have "(sum ?pqi ({.. degree p + degree q} - {degree p})) = 0" proof (rule sum.neutral, intro ballI) fix d assume d: "d \ {.. degree p + degree q} - {degree p}" show "?pqi d = 0" proof (cases "d < degree p") case True hence "degree p + degree q - d > degree q" by auto hence "coeff q (degree p + degree q - d) = 0" by (rule coeff_eq_0) thus ?thesis by simp next case False with d have "d > degree p" by auto hence "coeff p d = 0" by (rule coeff_eq_0) thus ?thesis by simp qed qed finally show ?thesis by simp qed lemma degree_monic_mult: assumes monic: "monic (p :: 'a :: comm_semiring_1 poly)" and q: "q \ 0" shows "degree (p * q) = degree p + degree q" proof - have "degree p + degree q \ degree (p * q)" by (rule degree_mult_le) also have "degree p + degree q \ degree (p * q)" proof - from q have cq: "coeff q (degree q) \ 0" by auto hence "coeff (p * q) (degree p + degree q) \ 0" unfolding lcoeff_monic_mult[OF monic] . thus "degree (p * q) \ degree p + degree q" by (rule le_degree) qed finally show ?thesis . qed lemma degree_prod_sum_monic: assumes S: "finite S" and nzd: "0 \ (degree o f) ` S" and monic: "(\ a . a \ S \ monic (f a))" shows "degree (prod f S) = (sum (degree o f) S) \ coeff (prod f S) (sum (degree o f) S) = 1" proof - from S nzd monic have "degree (prod f S) = sum (degree \ f) S \ (S \ {} \ degree (prod f S) \ 0 \ prod f S \ 0) \ coeff (prod f S) (sum (degree o f) S) = 1" proof (induct S rule: finite_induct) case (insert a S) have IH1: "degree (prod f S) = sum (degree o f) S" using insert by auto have IH2: "coeff (prod f S) (degree (prod f S)) = 1" using insert by auto have id: "degree (prod f (insert a S)) = sum (degree \ f) (insert a S) \ coeff (prod f (insert a S)) (sum (degree o f) (insert a S)) = 1" proof (cases "S = {}") case False with insert have nz: "prod f S \ 0" by auto from insert have monic: "coeff (f a) (degree (f a)) = 1" by auto have id: "(degree \ f) a = degree (f a)" by simp show ?thesis unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] id unfolding degree_monic_mult[OF monic nz] unfolding IH1[symmetric] unfolding lcoeff_monic_mult[OF monic] IH2 by simp qed (insert insert, auto) show ?case using id unfolding sum.insert[OF insert(1-2)] using insert by auto qed simp thus ?thesis by auto qed lemma degree_prod_monic: assumes "\ i. i < n \ degree (f i :: 'a :: comm_semiring_1 poly) = 1" and "\ i. i < n \ coeff (f i) 1 = 1" shows "degree (prod f {0 ..< n}) = n \ coeff (prod f {0 ..< n}) n = 1" proof - from degree_prod_sum_monic[of "{0 ..< n}" f] show ?thesis using assms by force qed lemma degree_prod_sum_lt_n: assumes "\ i. i < n \ degree (f i :: 'a :: comm_semiring_1 poly) \ 1" and i: "i < n" and fi: "degree (f i) = 0" shows "degree (prod f {0 ..< n}) < n" proof - have "degree (prod f {0 ..< n}) \ sum (degree o f) {0 ..< n}" by (rule degree_prod_sum_le, auto) also have "sum (degree o f) {0 ..< n} = (degree o f) i + sum (degree o f) ({0 ..< n} - {i})" by (rule sum.remove, insert i, auto) also have "(degree o f) i = 0" using fi by simp also have "sum (degree o f) ({0 ..< n} - {i}) \ sum (\ _. 1) ({0 ..< n} - {i})" by (rule sum_mono, insert assms, auto) also have "\ = n - 1" using i by simp also have "\ < n" using i by simp finally show ?thesis by simp qed lemma degree_linear_factors: "degree (\ a \ as. [: f a, 1:]) = length as" proof (induct as) case (Cons b as) note IH = this have id: "(\a\b # as. [:f a, 1:]) = [:f b,1 :] * (\a\as. [:f a, 1:])" by simp show ?case unfolding id by (subst degree_monic_mult, insert IH, auto) qed simp lemma monic_mult: fixes p q :: "'a :: idom poly" assumes "monic p" "monic q" shows "monic (p * q)" proof - from assms have nz: "p \ 0" "q \ 0" by auto show ?thesis unfolding degree_mult_eq[OF nz] coeff_mult_degree_sum using assms by simp qed lemma monic_factor: fixes p q :: "'a :: idom poly" assumes "monic (p * q)" "monic p" shows "monic q" proof - from assms have nz: "p \ 0" "q \ 0" by auto from assms[unfolded degree_mult_eq[OF nz] coeff_mult_degree_sum \monic p\] show ?thesis by simp qed lemma monic_prod: fixes f :: "'a \ 'b :: idom poly" assumes "\ a. a \ as \ monic (f a)" shows "monic (prod f as)" using assms proof (induct as rule: infinite_finite_induct) case (insert a as) hence id: "prod f (insert a as) = f a * prod f as" and *: "monic (f a)" "monic (prod f as)" by auto show ?case unfolding id by (rule monic_mult[OF *]) qed auto lemma monic_prod_list: fixes as :: "'a :: idom poly list" assumes "\ a. a \ set as \ monic a" shows "monic (prod_list as)" using assms by (induct as, auto intro: monic_mult) lemma monic_power: assumes "monic (p :: 'a :: idom poly)" shows "monic (p ^ n)" by (induct n, insert assms, auto intro: monic_mult) lemma monic_prod_list_pow: "monic (\(x::'a::idom, i)\xis. [:- x, 1:] ^ Suc i)" proof (rule monic_prod_list, goal_cases) case (1 a) then obtain x i where a: "a = [:-x, 1:]^Suc i" by force show "monic a" unfolding a by (rule monic_power, auto) qed lemma monic_degree_0: "monic p \ (degree p = 0) = (p = 1)" using le_degree poly_eq_iff by force subsection \Roots\ text \The following proof structure is completely similar to the one of @{thm poly_roots_finite}.\ lemma poly_roots_degree: fixes p :: "'a::idom poly" shows "p \ 0 \ card {x. poly p x = 0} \ degree p" 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 ?case by simp next case (Suc n p) show ?case proof (cases "\x. poly p x = 0") case True then obtain a where a: "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 Suc.hyps(1)[OF this \k \ 0\] have le: "card {x. poly k x = 0} \ degree k" . have "card {x. poly p x = 0} = card {x. poly ([:-a, 1:] * k) x = 0}" unfolding k .. also have "{x. poly ([:-a, 1:] * k) x = 0} = insert a {x. poly k x = 0}" by auto also have "card \ \ Suc (card {x. poly k x = 0})" unfolding card_insert_if[OF poly_roots_finite[OF \k \ 0\]] by simp also have "\ \ Suc (degree k)" using le by auto finally show ?thesis using \degree p = Suc (degree k)\ by simp qed simp qed lemma poly_root_factor: "(poly ([: r, 1:] * q) (k :: 'a :: idom) = 0) = (k = -r \ poly q k = 0)" (is ?one) "(poly (q * [: r, 1:]) k = 0) = (k = -r \ poly q k = 0)" (is ?two) "(poly [: r, 1 :] k = 0) = (k = -r)" (is ?three) proof - have [simp]: "r + k = 0 \ k = - r" by (simp add: minus_unique) show ?one unfolding poly_mult by auto show ?two unfolding poly_mult by auto show ?three by auto qed lemma poly_root_constant: "c \ 0 \ (poly (p * [:c:]) (k :: 'a :: idom) = 0) = (poly p k = 0)" unfolding poly_mult by auto lemma poly_linear_exp_linear_factors_rev: "([:b,1:])^(length (filter ((=) b) as)) dvd (\ (a :: 'a :: comm_ring_1) \ as. [: a, 1:])" proof (induct as) case (Cons a as) let ?ls = "length (filter ((=) b) (a # as))" let ?l = "length (filter ((=) b) as)" have prod: "(\ a \ Cons a as. [: a, 1:]) = [: a, 1 :] * (\ a \ as. [: a, 1:])" by simp show ?case proof (cases "a = b") case False hence len: "?ls = ?l" by simp show ?thesis unfolding prod len using Cons by (rule dvd_mult) next case True hence len: "[: b, 1 :] ^ ?ls = [: a, 1 :] * [: b, 1 :] ^ ?l" by simp show ?thesis unfolding prod len using Cons using dvd_refl mult_dvd_mono by blast qed qed simp lemma order_max: assumes dvd: "[: -a, 1 :] ^ k dvd p" and p: "p \ 0" shows "k \ order a p" proof (rule ccontr) assume "\ ?thesis" hence "\ j. k = Suc (order a p + j)" by arith then obtain j where k: "k = Suc (order a p + j)" by auto have "[: -a, 1 :] ^ Suc (order a p) dvd p" by (rule power_le_dvd[OF dvd[unfolded k]], simp) with order_2[OF p, of a] show False by blast qed subsection \Divisibility\ context assumes "SORT_CONSTRAINT('a :: idom)" begin lemma poly_linear_linear_factor: assumes dvd: "[:b,1:] dvd (\ (a :: 'a) \ as. [: a, 1:])" shows "b \ set as" proof - let ?p = "\ as. (\ a \ as. [: a, 1:])" let ?b = "[:b,1:]" from assms[unfolded dvd_def] obtain p where id: "?p as = ?b * p" .. from arg_cong[OF id, of "\ p. poly p (-b)"] have "poly (?p as) (-b) = 0" by simp thus ?thesis proof (induct as) case (Cons a as) have "?p (a # as) = [:a,1:] * ?p as" by simp from Cons(2)[unfolded this] have "poly (?p as) (-b) = 0 \ (a - b) = 0" by simp with Cons(1) show ?case by auto qed simp qed lemma poly_linear_exp_linear_factors: assumes dvd: "([:b,1:])^n dvd (\ (a :: 'a) \ as. [: a, 1:])" shows "length (filter ((=) b) as) \ n" proof - let ?p = "\ as. (\ a \ as. [: a, 1:])" let ?b = "[:b,1:]" from dvd show ?thesis proof (induct n arbitrary: as) case (Suc n as) have bs: "?b ^ Suc n = ?b * ?b ^ n" by simp from poly_linear_linear_factor[OF dvd_mult_left[OF Suc(2)[unfolded bs]], unfolded in_set_conv_decomp] obtain as1 as2 where as: "as = as1 @ b # as2" by auto have "?p as = [:b,1:] * ?p (as1 @ as2)" unfolding as proof (induct as1) case (Cons a as1) have "?p (a # as1 @ b # as2) = [:a,1:] * ?p (as1 @ b # as2)" by simp also have "?p (as1 @ b # as2) = [:b,1:] * ?p (as1 @ as2)" unfolding Cons by simp also have "[:a,1:] * \ = [:b,1:] * ([:a,1:] * ?p (as1 @ as2))" by (metis (no_types, lifting) mult.left_commute) finally show ?case by simp qed simp from Suc(2)[unfolded bs this dvd_mult_cancel_left] have "?b ^ n dvd ?p (as1 @ as2)" by simp from Suc(1)[OF this] show ?case unfolding as by simp qed simp qed end lemma const_poly_dvd: "([:a:] dvd [:b:]) = (a dvd b)" proof assume "a dvd b" then obtain c where "b = a * c" unfolding dvd_def by auto hence "[:b:] = [:a:] * [: c:]" by (auto simp: ac_simps) thus "[:a:] dvd [:b:]" unfolding dvd_def by blast next assume "[:a:] dvd [:b:]" then obtain pc where "[:b:] = [:a:] * pc" unfolding dvd_def by blast from arg_cong[OF this, of "\ p. coeff p 0", unfolded coeff_mult] have "b = a * coeff pc 0" by auto thus "a dvd b" unfolding dvd_def by blast qed lemma const_poly_dvd_1 [simp]: "[:a:] dvd 1 \ a dvd 1" by (metis const_poly_dvd one_poly_eq_simps(2)) lemma poly_dvd_1: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" shows "p dvd 1 \ degree p = 0 \ coeff p 0 dvd 1" proof (cases "degree p = 0") case False with divides_degree[of p 1] show ?thesis by auto next case True from degree0_coeffs[OF this] obtain a where p: "p = [:a:]" by auto show ?thesis unfolding p by auto qed text \Degree based version of irreducibility.\ definition irreducible\<^sub>d :: "'a :: comm_semiring_1 poly \ bool" where "irreducible\<^sub>d p = (degree p > 0 \ (\ q r. degree q < degree p \ degree r < degree p \ p \ q * r))" lemma irreducible\<^sub>dI [intro]: assumes 1: "degree p > 0" and 2: "\q r. degree q > 0 \ degree q < degree p \ degree r > 0 \ degree r < degree p \ p = q * r \ False" shows "irreducible\<^sub>d p" proof (unfold irreducible\<^sub>d_def, intro conjI allI impI notI 1) fix q r assume "degree q < degree p" and "degree r < degree p" and "p = q * r" with degree_mult_le[of q r] show False by (intro 2, auto) qed lemma irreducible\<^sub>dI2: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes deg: "degree p > 0" and ndvd: "\ q. degree q > 0 \ degree q \ degree p div 2 \ \ q dvd p" shows "irreducible\<^sub>d p" proof (rule ccontr) assume "\ ?thesis" from this[unfolded irreducible\<^sub>d_def] deg obtain q r where dq: "degree q < degree p" and dr: "degree r < degree p" and p: "p = q * r" by auto from deg have p0: "p \ 0" by auto with p have "q \ 0" "r \ 0" by auto from degree_mult_eq[OF this] p have dp: "degree p = degree q + degree r" by simp show False proof (cases "degree q \ degree p div 2") case True from ndvd[OF _ True] dq dr dp p show False by auto next case False with dp have dr: "degree r \ degree p div 2" by auto from p have dvd: "r dvd p" by auto from ndvd[OF _ dr] dvd dp dq show False by auto qed qed lemma reducible\<^sub>dI: assumes "degree p > 0 \ \q r. degree q < degree p \ degree r < degree p \ p = q * r" shows "\ irreducible\<^sub>d p" using assms by (auto simp: irreducible\<^sub>d_def) lemma irreducible\<^sub>dE [elim]: assumes "irreducible\<^sub>d p" and "degree p > 0 \ (\q r. degree q < degree p \ degree r < degree p \ p \ q * r) \ thesis" shows thesis using assms by (auto simp: irreducible\<^sub>d_def) lemma reducible\<^sub>dE [elim]: assumes red: "\ irreducible\<^sub>d p" and 1: "degree p = 0 \ thesis" and 2: "\q r. degree q > 0 \ degree q < degree p \ degree r > 0 \ degree r < degree p \ p = q * r \ thesis" shows thesis using red[unfolded irreducible\<^sub>d_def de_Morgan_conj not_not not_all not_imp] proof (elim disjE exE conjE) show "\degree p > 0 \ thesis" using 1 by auto next fix q r assume "degree q < degree p" and "degree r < degree p" and "p = q * r" with degree_mult_le[of q r] show thesis by (intro 2, auto) qed lemma irreducible\<^sub>dD: assumes "irreducible\<^sub>d p" shows "degree p > 0" "\q r. degree q < degree p \ degree r < degree p \ p \ q * r" using assms unfolding irreducible\<^sub>d_def by auto theorem irreducible\<^sub>d_factorization_exists: assumes "degree p > 0" shows "\fs. fs \ [] \ (\f \ set fs. irreducible\<^sub>d f \ degree f \ degree p) \ p = prod_list fs" and "\irreducible\<^sub>d p \ \fs. length fs > 1 \ (\f \ set fs. irreducible\<^sub>d f \ degree f < degree p) \ p = prod_list fs" proof (atomize(full), insert assms, induct "degree p" arbitrary:p rule: less_induct) case less then have deg_f: "degree p > 0" by auto show ?case proof (cases "irreducible\<^sub>d p") case True then have "set [p] \ Collect irreducible\<^sub>d" "p = prod_list [p]" by auto with True show ?thesis by (auto intro: exI[of _ "[p]"]) next case False with deg_f obtain g h where deg_g: "degree g < degree p" "degree g > 0" and deg_h: "degree h < degree p" "degree h > 0" and f_gh: "p = g * h" by auto from less.hyps[OF deg_g] less.hyps[OF deg_h] obtain gs hs where emp: "length gs > 0" "length hs > 0" and "\f \ set gs. irreducible\<^sub>d f \ degree f \ degree g" "g = prod_list gs" and "\f \ set hs. irreducible\<^sub>d f \ degree f \ degree h" "h = prod_list hs" by auto with f_gh deg_g deg_h have len: "length (gs@hs) > 1" and mem: "\f \ set (gs@hs). irreducible\<^sub>d f \ degree f < degree p" and p: "p = prod_list (gs@hs)" by (auto simp del: length_greater_0_conv) with False show ?thesis by (auto intro!: exI[of _ "gs@hs"] simp: less_imp_le) qed qed lemma irreducible\<^sub>d_factor: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree p > 0" shows "\ q r. irreducible\<^sub>d q \ p = q * r \ degree r < degree p" using assms proof (induct "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "irreducible\<^sub>d p") case False with less(2) obtain q r where q: "degree q < degree p" "degree q > 0" and r: "degree r < degree p" "degree r > 0" and p: "p = q * r" by auto from less(1)[OF q] obtain s t where IH: "irreducible\<^sub>d s" "q = s * t" by auto from p have p: "p = s * (t * r)" unfolding IH by (simp add: ac_simps) from less(2) have "p \ 0" by auto hence "degree p = degree s + (degree (t * r))" unfolding p by (subst degree_mult_eq, insert p, auto) with irreducible\<^sub>dD[OF IH(1)] have "degree p > degree (t * r)" by auto with p IH show ?thesis by auto next case True show ?thesis by (rule exI[of _ p], rule exI[of _ 1], insert True less(2), auto) qed qed context mult_zero begin (* least class with times and zero *) definition zero_divisor where "zero_divisor a \ \b. b \ 0 \ a * b = 0" lemma zero_divisorI[intro]: assumes "b \ 0" and "a * b = 0" shows "zero_divisor a" using assms by (auto simp: zero_divisor_def) lemma zero_divisorE[elim]: assumes "zero_divisor a" and "\b. b \ 0 \ a * b = 0 \ thesis" shows thesis using assms by (auto simp: zero_divisor_def) end lemma zero_divisor_0[simp]: "zero_divisor (0::'a::{mult_zero,zero_neq_one})" (* No need for one! *) by (auto intro!: zero_divisorI[of 1]) lemma not_zero_divisor_1: "\ zero_divisor (1 :: 'a :: {monoid_mult,mult_zero})" (* No need for associativity! *) by auto lemma zero_divisor_iff_eq_0[simp]: fixes a :: "'a :: {semiring_no_zero_divisors, zero_neq_one}" shows "zero_divisor a \ a = 0" by auto lemma mult_eq_0_not_zero_divisor_left[simp]: fixes a b :: "'a :: mult_zero" assumes "\ zero_divisor a" shows "a * b = 0 \ b = 0" using assms unfolding zero_divisor_def by force lemma mult_eq_0_not_zero_divisor_right[simp]: fixes a b :: "'a :: {ab_semigroup_mult,mult_zero}" (* No need for associativity! *) assumes "\ zero_divisor b" shows "a * b = 0 \ a = 0" using assms unfolding zero_divisor_def by (force simp: ac_simps) lemma degree_smult_not_zero_divisor_left[simp]: assumes "\ zero_divisor c" shows "degree (smult c p) = degree p" proof(cases "p = 0") case False then have "coeff (smult c p) (degree p) \ 0" using assms by auto from le_degree[OF this] degree_smult_le[of c p] show ?thesis by auto qed auto lemma degree_smult_not_zero_divisor_right[simp]: assumes "\ zero_divisor (lead_coeff p)" shows "degree (smult c p) = (if c = 0 then 0 else degree p)" proof(cases "c = 0") case False then have "coeff (smult c p) (degree p) \ 0" using assms by auto from le_degree[OF this] degree_smult_le[of c p] show ?thesis by auto qed auto lemma irreducible\<^sub>d_smult_not_zero_divisor_left: assumes c0: "\ zero_divisor c" assumes L: "irreducible\<^sub>d (smult c p)" shows "irreducible\<^sub>d p" proof (intro irreducible\<^sub>dI) from L have "degree (smult c p) > 0" by auto also note degree_smult_le finally show "degree p > 0" by auto fix q r assume deg_q: "degree q < degree p" and deg_r: "degree r < degree p" and p_qr: "p = q * r" then have 1: "smult c p = smult c q * r" by auto note degree_smult_le[of c q] also note deg_q finally have 2: "degree (smult c q) < degree (smult c p)" using c0 by auto from deg_r have 3: "degree r < \" using c0 by auto from irreducible\<^sub>dD(2)[OF L 2 3] 1 show False by auto qed lemmas irreducible\<^sub>d_smultI = irreducible\<^sub>d_smult_not_zero_divisor_left [where 'a = "'a :: {comm_semiring_1,semiring_no_zero_divisors}", simplified] lemma irreducible\<^sub>d_smult_not_zero_divisor_right: assumes p0: "\ zero_divisor (lead_coeff p)" and L: "irreducible\<^sub>d (smult c p)" shows "irreducible\<^sub>d p" proof- from L have "c \ 0" by auto with p0 have [simp]: "degree (smult c p) = degree p" by simp show "irreducible\<^sub>d p" proof (intro iffI irreducible\<^sub>dI conjI) from L show "degree p > 0" by auto fix q r assume deg_q: "degree q < degree p" and deg_r: "degree r < degree p" and p_qr: "p = q * r" then have 1: "smult c p = smult c q * r" by auto note degree_smult_le[of c q] also note deg_q finally have 2: "degree (smult c q) < degree (smult c p)" by simp from deg_r have 3: "degree r < \" by simp from irreducible\<^sub>dD(2)[OF L 2 3] 1 show False by auto qed qed lemma zero_divisor_mult_left: fixes a b :: "'a :: {ab_semigroup_mult, mult_zero}" assumes "zero_divisor a" shows "zero_divisor (a * b)" proof- from assms obtain c where c0: "c \ 0" and [simp]: "a * c = 0" by auto have "a * b * c = a * c * b" by (simp only: ac_simps) with c0 show ?thesis by auto qed lemma zero_divisor_mult_right: fixes a b :: "'a :: {semigroup_mult, mult_zero}" assumes "zero_divisor b" shows "zero_divisor (a * b)" proof- from assms obtain c where c0: "c \ 0" and [simp]: "b * c = 0" by auto have "a * b * c = a * (b * c)" by (simp only: ac_simps) with c0 show ?thesis by auto qed lemma not_zero_divisor_mult: fixes a b :: "'a :: {ab_semigroup_mult, mult_zero}" assumes "\ zero_divisor (a * b)" shows "\ zero_divisor a" and "\ zero_divisor b" using assms by (auto dest: zero_divisor_mult_right zero_divisor_mult_left) lemma zero_divisor_smult_left: assumes "zero_divisor a" shows "zero_divisor (smult a f)" proof- from assms obtain b where b0: "b \ 0" and "a * b = 0" by auto then have "smult a f * [:b:] = 0" by (simp add: ac_simps) with b0 show ?thesis by (auto intro!: zero_divisorI[of "[:b:]"]) qed lemma unit_not_zero_divisor: fixes a :: "'a :: {comm_monoid_mult, mult_zero}" assumes "a dvd 1" shows "\zero_divisor a" proof from assms obtain b where ab: "1 = a * b" by (elim dvdE) assume "zero_divisor a" then have "zero_divisor (1::'a)" by (unfold ab, intro zero_divisor_mult_left) then show False by auto qed lemma linear_irreducible\<^sub>d: assumes "degree p = 1" shows "irreducible\<^sub>d p" by (rule irreducible\<^sub>dI, insert assms, auto) lemma irreducible\<^sub>d_dvd_smult: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree p > 0" "irreducible\<^sub>d q" "p dvd q" shows "\ c. c \ 0 \ q = smult c p" proof - from assms obtain r where q: "q = p * r" by (elim dvdE, auto) from degree_mult_eq[of p r] assms(1) q have "\ degree p < degree q" and nz: "p \ 0" "q \ 0" apply (metis assms(2) degree_mult_eq_0 gr_implies_not_zero irreducible\<^sub>dD(2) less_add_same_cancel2) using assms by auto hence deg: "degree p \ degree q" by auto from \p dvd q\ obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with nz have "k \ 0" by auto from deg[unfolded q degree_mult_eq[OF \k \ 0\ \p \ 0\ ]] have "degree k = 0" unfolding q by auto then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed subsection \Map over Polynomial Coefficients\ lemma map_poly_simps: shows "map_poly f (pCons c p) = (if c = 0 \ p = 0 then 0 else pCons (f c) (map_poly f p))" proof (cases "c = 0") case True note c0 = this show ?thesis proof (cases "p = 0") case True thus ?thesis using c0 unfolding map_poly_def by simp next case False thus ?thesis unfolding map_poly_def by auto qed next case False thus ?thesis unfolding map_poly_def by auto qed lemma map_poly_pCons[simp]: assumes "c \ 0 \ p \ 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" unfolding map_poly_simps using assms by auto lemma map_poly_map_poly: assumes f0: "f 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" proof (induct p) case (pCons a p) show ?case proof(cases "g a \ 0 \ map_poly g p \ 0") case True show ?thesis unfolding map_poly_pCons[OF pCons(1)] unfolding map_poly_pCons[OF True] unfolding pCons(2) by simp next case False then show ?thesis unfolding map_poly_pCons[OF pCons(1)] unfolding pCons(2)[symmetric] by (simp add: f0) qed qed simp lemma map_poly_zero: assumes f: "\c. f c = 0 \ c = 0" shows [simp]: "map_poly f p = 0 \ p = 0" by (induct p; auto simp: map_poly_simps f) lemma map_poly_add: assumes h0: "h 0 = 0" and h_add: "\p q. h (p + q) = h p + h q" shows "map_poly h (p + q) = map_poly h p + map_poly h q" proof (induct p arbitrary: q) case (pCons a p) note pIH = this show ?case proof(induct "q") case (pCons b q) note qIH = this show ?case unfolding map_poly_pCons[OF qIH(1)] unfolding map_poly_pCons[OF pIH(1)] unfolding add_pCons unfolding pIH(2)[symmetric] unfolding h_add[rule_format,symmetric] unfolding map_poly_simps using h0 by auto qed auto qed auto subsection \Morphismic properties of @{term "pCons 0"}\ lemma monom_pCons_0_monom: "monom (pCons 0 (monom a n)) d = map_poly (pCons 0) (monom (monom a n) d)" apply (induct d) unfolding monom_0 unfolding map_poly_simps apply simp unfolding monom_Suc map_poly_simps by auto lemma pCons_0_add: "pCons 0 (p + q) = pCons 0 p + pCons 0 q" by auto lemma sum_pCons_0_commute: "sum (\i. pCons 0 (f i)) S = pCons 0 (sum f S)" by(induct S rule: infinite_finite_induct;simp) lemma pCons_0_as_mult: fixes p:: "'a :: comm_semiring_1 poly" shows "pCons 0 p = [:0,1:] * p" by auto subsection \Misc\ fun expand_powers :: "(nat \ 'a)list \ 'a list" where "expand_powers [] = []" | "expand_powers ((Suc n, a) # ps) = a # expand_powers ((n,a) # ps)" | "expand_powers ((0,a) # ps) = expand_powers ps" lemma expand_powers: fixes f :: "'a \ 'b :: comm_ring_1" shows "(\ (n,a) \ n_as. f a ^ n) = (\ a \ expand_powers n_as. f a)" by (rule sym, induct n_as rule: expand_powers.induct, auto) lemma poly_smult_zero_iff: fixes x :: "'a :: idom" shows "(poly (smult a p) x = 0) = (a = 0 \ poly p x = 0)" by simp lemma poly_prod_list_zero_iff: fixes x :: "'a :: idom" shows "(poly (prod_list ps) x = 0) = (\ p \ set ps. poly p x = 0)" by (induct ps, auto) lemma poly_mult_zero_iff: fixes x :: "'a :: idom" shows "(poly (p * q) x = 0) = (poly p x = 0 \ poly q x = 0)" by simp lemma poly_power_zero_iff: fixes x :: "'a :: idom" shows "(poly (p^n) x = 0) = (n \ 0 \ poly p x = 0)" by (cases n, auto) lemma sum_monom_0_iff: assumes fin: "finite S" and g: "\ i j. g i = g j \ i = j" shows "sum (\ i. monom (f i) (g i)) S = 0 \ (\ i \ S. f i = 0)" (is "?l = ?r") proof - { assume "\ ?r" then obtain i where i: "i \ S" and fi: "f i \ 0" by auto let ?g = "\ i. monom (f i) (g i)" have "coeff (sum ?g S) (g i) = f i + sum (\ j. coeff (?g j) (g i)) (S - {i})" by (unfold sum.remove[OF fin i], simp add: coeff_sum) also have "sum (\ j. coeff (?g j) (g i)) (S - {i}) = 0" by (rule sum.neutral, insert g, auto) finally have "coeff (sum ?g S) (g i) \ 0" using fi by auto hence "\ ?l" by auto } thus ?thesis by auto qed lemma degree_prod_list_eq: assumes "\ p. p \ set ps \ (p :: 'a :: idom poly) \ 0" shows "degree (prod_list ps) = sum_list (map degree ps)" using assms proof (induct ps) case (Cons p ps) show ?case unfolding prod_list.Cons by (subst degree_mult_eq, insert Cons, auto simp: prod_list_zero_iff) qed simp lemma degree_power_eq: assumes p: "p \ 0" shows "degree (p ^ n) = degree (p :: 'a :: idom poly) * n" proof (induct n) case (Suc n) from p have pn: "p ^ n \ 0" by auto show ?case using degree_mult_eq[OF p pn] Suc by auto qed simp lemma coeff_Poly: "coeff (Poly xs) i = (nth_default 0 xs i)" unfolding nth_default_coeffs_eq[of "Poly xs", symmetric] coeffs_Poly by simp lemma rsquarefree_def': "rsquarefree p = (p \ 0 \ (\a. order a p \ 1))" proof - have "\ a. order a p \ 1 \ order a p = 0 \ order a p = 1" by linarith thus ?thesis unfolding rsquarefree_def by auto qed lemma order_prod_list: "(\ p. p \ set ps \ p \ 0) \ order x (prod_list ps) = sum_list (map (order x) ps)" by (induct ps, auto, subst order_mult, auto simp: prod_list_zero_iff) lemma irreducible\<^sub>d_dvd_eq: fixes a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "irreducible\<^sub>d a" and "irreducible\<^sub>d b" and "a dvd b" and "monic a" and "monic b" shows "a = b" using assms by (metis (no_types, lifting) coeff_smult degree_smult_eq irreducible\<^sub>dD(1) irreducible\<^sub>d_dvd_smult mult.right_neutral smult_1_left) lemma monic_gcd_dvd: assumes fg: "f dvd g" and mon: "monic f" and gcd: "gcd g h \ {1, g}" shows "gcd f h \ {1, f}" proof (cases "coprime g h") case True with dvd_refl have "coprime f h" using fg by (blast intro: coprime_divisors) then show ?thesis by simp next case False with gcd have gcd: "gcd g h = g" by (simp add: coprime_iff_gcd_eq_1) with fg have "f dvd gcd g h" by simp then have "f dvd h" by simp then have "gcd f h = normalize f" by (simp add: gcd_proj1_iff) also have "normalize f = f" using mon by (rule normalize_monic) finally show ?thesis by simp qed lemma monom_power: "(monom a b)^n = monom (a^n) (b*n)" by (induct n, auto simp add: mult_monom) lemma poly_const_pow: "[:a:]^b = [:a^b:]" by (metis Groups.mult_ac(2) monom_0 monom_power mult_zero_right) lemma degree_pderiv_le: "degree (pderiv f) \ degree f - 1" proof (rule ccontr) assume "\ ?thesis" hence ge: "degree (pderiv f) \ Suc (degree f - 1)" by auto hence "pderiv f \ 0" by auto hence "coeff (pderiv f) (degree (pderiv f)) \ 0" by auto from this[unfolded coeff_pderiv] have "coeff f (Suc (degree (pderiv f))) \ 0" by auto moreover have "Suc (degree (pderiv f)) > degree f" using ge by auto ultimately show False by (simp add: coeff_eq_0) qed lemma map_div_is_smult_inverse: "map_poly (\x. x / (a :: 'a :: field)) p = smult (inverse a) p" unfolding smult_conv_map_poly by (simp add: divide_inverse_commute) lemma normalize_poly_old_def: "normalize (f :: 'a :: {normalization_semidom,field} poly) = smult (inverse (unit_factor (lead_coeff f))) f" by (simp add: normalize_poly_eq_map_poly map_div_is_smult_inverse) (* was in Euclidean_Algorithm in Number_Theory before, but has been removed *) lemma poly_dvd_antisym: fixes p q :: "'b::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 coeff_f_0_code[code_unfold]: "coeff f 0 = (case coeffs f of [] \ 0 | x # _ \ x)" by (cases f, auto simp: cCons_def) lemma poly_compare_0_code[code_unfold]: "(f = 0) = (case coeffs f of [] \ True | _ \ False)" using coeffs_eq_Nil list.disc_eq_case(1) by blast text \Getting more efficient code for abbreviation @{term lead_coeff}"\ definition leading_coeff where [code_abbrev, simp]: "leading_coeff = lead_coeff" lemma leading_coeff_code [code]: "leading_coeff f = (let xs = coeffs f in if xs = [] then 0 else last xs)" by (simp add: last_coeffs_eq_coeff_degree) lemma nth_coeffs_coeff: "i < length (coeffs f) \ coeffs f ! i = coeff f i" by (metis nth_default_coeffs_eq nth_default_def) lemma degree_prod_eq_sum_degree: fixes A :: "'a set" and f :: "'a \ 'b::field poly" assumes f0: "\i\A. f i \ 0" shows "degree (\i\A. (f i)) = (\i\A. degree (f i))" using f0 proof (induct A rule: infinite_finite_induct) case (insert x A) have "(\i\insert x A. degree (f i)) = degree (f x) + (\i\A. degree (f i))" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... = degree (f x) + degree (\i\A. (f i))" by (simp add: insert.hyps insert.prems) also have "... = degree (f x * (\i\A. (f i)))" proof (rule degree_mult_eq[symmetric]) show "f x \ 0" using insert.prems by auto show "prod f A \ 0" by (simp add: insert.hyps(1) insert.prems) qed also have "... = degree (\i\insert x A. (f i))" by (simp add: insert.hyps) finally show ?case .. qed auto definition monom_mult :: "nat \ 'a :: comm_semiring_1 poly \ 'a poly" where "monom_mult n f = monom 1 n * f" lemma monom_mult_unfold [code_unfold]: "monom 1 n * f = monom_mult n f" "f * monom 1 n = monom_mult n f" by (auto simp: monom_mult_def ac_simps) lemma monom_mult_code [code abstract]: "coeffs (monom_mult n f) = (let xs = coeffs f in if xs = [] then xs else replicate n 0 @ xs)" by (rule coeffs_eqI) (auto simp add: Let_def monom_mult_def coeff_monom_mult nth_default_append nth_default_coeffs_eq) lemma coeff_pcompose_monom: fixes f :: "'a :: comm_ring_1 poly" assumes n: "j < n" shows "coeff (f \\<^sub>p monom 1 n) (n * i + j) = (if j = 0 then coeff f i else 0)" proof (induct f arbitrary: i) case (pCons a f i) note d = pcompose_pCons coeff_add coeff_monom_mult coeff_pCons show ?case proof (cases i) case 0 show ?thesis unfolding d 0 using n by (cases j, auto) next case (Suc ii) have id: "n * Suc ii + j - n = n * ii + j" using n by (simp add: diff_mult_distrib2) have id1: "(n \ n * Suc ii + j) = True" by auto have id2: "(case n * Suc ii + j of 0 \ a | Suc x \ coeff 0 x) = 0" using n by (cases "n * Suc ii + j", auto) show ?thesis unfolding d Suc id id1 id2 pCons(2) if_True by auto qed qed auto lemma coeff_pcompose_x_pow_n: fixes f :: "'a :: comm_ring_1 poly" assumes n: "n \ 0" shows "coeff (f \\<^sub>p monom 1 n) (n * i) = coeff f i" using coeff_pcompose_monom[of 0 n f i] n by auto lemma dvd_dvd_smult: "a dvd b \ f dvd g \ smult a f dvd smult b g" unfolding dvd_def by (metis mult_smult_left mult_smult_right smult_smult) definition sdiv_poly :: "'a :: idom_divide poly \ 'a \ 'a poly" where "sdiv_poly p a = (map_poly (\ c. c div a) p)" lemma smult_map_poly: "smult a = map_poly ((*) a)" by (rule ext, rule poly_eqI, subst coeff_map_poly, auto) lemma smult_exact_sdiv_poly: assumes "\ c. c \ set (coeffs p) \ a dvd c" shows "smult a (sdiv_poly p a) = p" unfolding smult_map_poly sdiv_poly_def by (subst map_poly_map_poly,simp,rule map_poly_idI, insert assms, auto) lemma coeff_sdiv_poly: "coeff (sdiv_poly f a) n = coeff f n div a" unfolding sdiv_poly_def by (rule coeff_map_poly, auto) lemma poly_pinfty_ge: fixes p :: "real poly" assumes "lead_coeff p > 0" "degree p \ 0" shows "\n. \ x \ n. poly p x \ b" proof - let ?p = "p - [:b - lead_coeff p :]" have id: "lead_coeff ?p = lead_coeff p" using assms(2) by (cases p, auto) with assms(1) have "lead_coeff ?p > 0" by auto from poly_pinfty_gt_lc[OF this, unfolded id] obtain n where "\ x. x \ n \ 0 \ poly p x - b" by auto thus ?thesis by auto qed lemma pderiv_sum: "pderiv (sum f I) = sum (\ i. (pderiv (f i))) I" by (induct I rule: infinite_finite_induct, auto simp: pderiv_add) lemma smult_sum2: "smult m (\i \ S. f i) = (\i \ S. smult m (f i))" by (induct S rule: infinite_finite_induct, auto simp add: smult_add_right) lemma degree_mult_not_eq: "degree (f * g) \ degree f + degree g \ lead_coeff f * lead_coeff g = 0" by (rule ccontr, auto simp: coeff_mult_degree_sum degree_mult_le le_antisym le_degree) lemma irreducible\<^sub>d_multD: fixes a b :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes l: "irreducible\<^sub>d (a*b)" shows "degree a = 0 \ a \ 0 \ irreducible\<^sub>d b \ degree b = 0 \ b \ 0 \ irreducible\<^sub>d a" proof- from l have a0: "a \ 0" and b0: "b \ 0" by auto note [simp] = degree_mult_eq[OF this] from l have "degree a = 0 \ degree b = 0" apply (unfold irreducible\<^sub>d_def) by force then show ?thesis proof(elim disjE) assume a: "degree a = 0" with l a0 have "irreducible\<^sub>d b" by (simp add: irreducible\<^sub>d_def) (metis degree_mult_eq degree_mult_eq_0 mult.left_commute plus_nat.add_0) with a a0 show ?thesis by auto next assume b: "degree b = 0" with l b0 have "irreducible\<^sub>d a" unfolding irreducible\<^sub>d_def by (smt add_cancel_left_right degree_mult_eq degree_mult_eq_0 neq0_conv semiring_normalization_rules(16)) with b b0 show ?thesis by auto qed qed lemma irreducible_connect_field[simp]: fixes f :: "'a :: field poly" shows "irreducible\<^sub>d f = irreducible f" (is "?l = ?r") proof show "?r \ ?l" apply (intro irreducible\<^sub>dI, force simp:is_unit_iff_degree) by (auto dest!: irreducible_multD simp: poly_dvd_1) next assume l: ?l show ?r proof (rule irreducibleI) from l show "f \ 0" "\ is_unit f" by (auto simp: poly_dvd_1) fix a b assume "f = a * b" from l[unfolded this] show "a dvd 1 \ b dvd 1" by (auto dest!: irreducible\<^sub>d_multD simp:is_unit_iff_degree) qed qed lemma is_unit_field_poly[simp]: fixes p :: "'a::field poly" shows "is_unit p \ p \ 0 \ degree p = 0" by (cases "p=0", auto simp: is_unit_iff_degree) lemma irreducible_smult_field[simp]: fixes c :: "'a :: field" shows "irreducible (smult c p) \ c \ 0 \ irreducible p" (is "?L \ ?R") proof (intro iffI conjI irreducible\<^sub>d_smult_not_zero_divisor_left[of c p, simplified]) assume "irreducible (smult c p)" then show "c \ 0" by auto next assume ?R then have c0: "c \ 0" and irr: "irreducible p" by auto show ?L proof (fold irreducible_connect_field, intro irreducible\<^sub>dI, unfold degree_smult_eq if_not_P[OF c0]) show "degree p > 0" using irr by auto fix q r from c0 have "p = smult (1/c) (smult c p)" by simp also assume "smult c p = q * r" finally have [simp]: "p = smult (1/c) \". assume main: "degree q < degree p" "degree r < degree p" have "\irreducible\<^sub>d p" by (rule reducible\<^sub>dI, rule exI[of _ "smult (1/c) q"], rule exI[of _ r], insert irr c0 main, simp) with irr show False by auto qed qed auto lemma irreducible_monic_factor: fixes p :: "'a :: field poly" assumes "degree p > 0" shows "\ q r. irreducible q \ p = q * r \ monic q" proof - from irreducible\<^sub>d_factorization_exists[OF assms] obtain fs where "fs \ []" and "set fs \ Collect irreducible" and "p = prod_list fs" by auto then have q: "irreducible (hd fs)" and p: "p = hd fs * prod_list (tl fs)" by (atomize(full), cases fs, auto) define c where "c = coeff (hd fs) (degree (hd fs))" from q have c: "c \ 0" unfolding c_def irreducible\<^sub>d_def by auto show ?thesis by (rule exI[of _ "smult (1/c) (hd fs)"], rule exI[of _ "smult c (prod_list (tl fs))"], unfold p, insert q c, auto simp: c_def) qed lemma monic_irreducible_factorization: fixes p :: "'a :: field poly" shows "monic p \ \ as f. finite as \ p = prod (\ a. a ^ Suc (f a)) as \ as \ {q. irreducible q \ monic q}" proof (induct "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "degree p > 0") case False with less(2) have "p = 1" by (simp add: coeff_eq_0 poly_eq_iff) thus ?thesis by (intro exI[of _ "{}"], auto) next case True from irreducible\<^sub>d_factor[OF this] obtain q r where p: "p = q * r" and q: "irreducible q" and deg: "degree r < degree p" by auto hence q0: "q \ 0" by auto define c where "c = coeff q (degree q)" let ?q = "smult (1/c) q" let ?r = "smult c r" from q0 have c: "c \ 0" "1 / c \ 0" unfolding c_def by auto hence p: "p = ?q * ?r" unfolding p by auto have deg: "degree ?r < degree p" using c deg by auto let ?Q = "{q. irreducible q \ monic (q :: 'a poly)}" have mon: "monic ?q" unfolding c_def using q0 by auto from monic_factor[OF \monic p\[unfolded p] this] have "monic ?r" . from less(1)[OF deg this] obtain f as where as: "finite as" "?r = (\ a \as. a ^ Suc (f a))" "as \ ?Q" by blast from q c have irred: "irreducible ?q" by simp show ?thesis proof (cases "?q \ as") case False let ?as = "insert ?q as" let ?f = "\ a. if a = ?q then 0 else f a" have "p = ?q * (\ a \as. a ^ Suc (f a))" unfolding p as by simp also have "(\ a \as. a ^ Suc (f a)) = (\ a \as. a ^ Suc (?f a))" by (rule prod.cong, insert False, auto) also have "?q * \ = (\ a \ ?as. a ^ Suc (?f a))" by (subst prod.insert, insert as False, auto) finally have p: "p = (\ a \ ?as. a ^ Suc (?f a))" . from as(1) have fin: "finite ?as" by auto from as mon irred have Q: "?as \ ?Q" by auto from fin p Q show ?thesis by(intro exI[of _ ?as] exI[of _ ?f], auto) next case True let ?f = "\ a. if a = ?q then Suc (f a) else f a" have "p = ?q * (\ a \as. a ^ Suc (f a))" unfolding p as by simp also have "(\ a \as. a ^ Suc (f a)) = ?q ^ Suc (f ?q) * (\ a \(as - {?q}). a ^ Suc (f a))" by (subst prod.remove[OF _ True], insert as, auto) also have "(\ a \(as - {?q}). a ^ Suc (f a)) = (\ a \(as - {?q}). a ^ Suc (?f a))" by (rule prod.cong, auto) also have "?q * (?q ^ Suc (f ?q) * \ ) = ?q ^ Suc (?f ?q) * \" by (simp add: ac_simps) also have "\ = (\ a \ as. a ^ Suc (?f a))" by (subst prod.remove[OF _ True], insert as, auto) finally have "p = (\ a \ as. a ^ Suc (?f a))" . with as show ?thesis by (intro exI[of _ as] exI[of _ ?f], auto) qed qed qed lemma monic_irreducible_gcd: - "monic (f::'a::{field,euclidean_ring_gcd} poly) \ irreducible f \ gcd f u \ {1,f}" + "monic (f::'a::{field,euclidean_ring_gcd,semiring_gcd_mult_normalize, + normalization_euclidean_semiring_multiplicative} poly) \ + irreducible f \ gcd f u \ {1,f}" by (metis gcd_dvd1 irreducible_altdef insertCI is_unit_gcd_iff poly_dvd_antisym poly_gcd_monic) end diff --git a/thys/Polynomial_Interpolation/Missing_Unsorted.thy b/thys/Polynomial_Interpolation/Missing_Unsorted.thy --- a/thys/Polynomial_Interpolation/Missing_Unsorted.thy +++ b/thys/Polynomial_Interpolation/Missing_Unsorted.thy @@ -1,652 +1,652 @@ (* Author: RenĂ© Thiemann Akihisa Yamada Jose Divason License: BSD *) section \Missing Unsorted\ text \This theory contains several lemmas which might be of interest to the Isabelle distribution. For instance, we prove that $b^n \cdot n^k$ is bounded by a constant whenever $0 < b < 1$.\ theory Missing_Unsorted imports HOL.Complex "HOL-Computational_Algebra.Factorial_Ring" begin lemma bernoulli_inequality: assumes x: "-1 \ (x :: 'a :: linordered_field)" shows "1 + of_nat n * x \ (1 + x) ^ n" proof (induct n) case (Suc n) have "1 + of_nat (Suc n) * x = 1 + x + of_nat n * x" by (simp add: field_simps) also have "\ \ \ + of_nat n * x ^ 2" by simp also have "\ = (1 + of_nat n * x) * (1 + x)" by (simp add: field_simps power2_eq_square) also have "\ \ (1 + x) ^ n * (1 + x)" by (rule mult_right_mono[OF Suc], insert x, auto) also have "\ = (1 + x) ^ (Suc n)" by simp finally show ?case . qed simp context fixes b :: "'a :: archimedean_field" assumes b: "0 < b" "b < 1" begin private lemma pow_one: "b ^ x \ 1" using power_Suc_less_one[OF b, of "x - 1"] by (cases x, auto) private lemma pow_zero: "0 < b ^ x" using b(1) by simp lemma exp_tends_to_zero: assumes c: "c > 0" shows "\ x. b ^ x \ c" proof (rule ccontr) assume not: "\ ?thesis" define bb where "bb = inverse b" define cc where "cc = inverse c" from b have bb: "bb > 1" unfolding bb_def by (rule one_less_inverse) from c have cc: "cc > 0" unfolding cc_def by simp define bbb where "bbb = bb - 1" have id: "bb = 1 + bbb" and bbb: "bbb > 0" and bm1: "bbb \ -1" unfolding bbb_def using bb by auto have "\ n. cc / bbb < of_nat n" by (rule reals_Archimedean2) then obtain n where lt: "cc / bbb < of_nat n" by auto from not have "\ b ^ n \ c" by auto hence bnc: "b ^ n > c" by simp have "bb ^ n = inverse (b ^ n)" unfolding bb_def by (rule power_inverse) also have "\ < cc" unfolding cc_def by (rule less_imp_inverse_less[OF bnc c]) also have "\ < bbb * of_nat n" using lt bbb by (metis mult.commute pos_divide_less_eq) also have "\ \ bb ^ n" using bernoulli_inequality[OF bm1, folded id, of n] by (simp add: ac_simps) finally show False by simp qed lemma linear_exp_bound: "\ p. \ x. b ^ x * of_nat x \ p" proof - from b have "1 - b > 0" by simp from exp_tends_to_zero[OF this] obtain x0 where x0: "b ^ x0 \ 1 - b" .. { fix x assume "x \ x0" hence "\ y. x = x0 + y" by arith then obtain y where x: "x = x0 + y" by auto have "b ^ x = b ^ x0 * b ^ y" unfolding x by (simp add: power_add) also have "\ \ b ^ x0" using pow_one[of y] pow_zero[of x0] by auto also have "\ \ 1 - b" by (rule x0) finally have "b ^ x \ 1 - b" . } note x0 = this define bs where "bs = insert 1 { b ^ Suc x * of_nat (Suc x) | x . x \ x0}" have bs: "finite bs" unfolding bs_def by auto define p where "p = Max bs" have bs: "\ b. b \ bs \ b \ p" unfolding p_def using bs by simp hence p1: "p \ 1" unfolding bs_def by auto show ?thesis proof (rule exI[of _ p], intro allI) fix x show "b ^ x * of_nat x \ p" proof (induct x) case (Suc x) show ?case proof (cases "x \ x0") case True show ?thesis by (rule bs, unfold bs_def, insert True, auto) next case False let ?x = "of_nat x :: 'a" have "b ^ (Suc x) * of_nat (Suc x) = b * (b ^ x * ?x) + b ^ Suc x" by (simp add: field_simps) also have "\ \ b * p + b ^ Suc x" by (rule add_right_mono[OF mult_left_mono[OF Suc]], insert b, auto) also have "\ = p - ((1 - b) * p - b ^ (Suc x))" by (simp add: field_simps) also have "\ \ p - 0" proof - have "b ^ Suc x \ 1 - b" using x0[of "Suc x"] False by auto also have "\ \ (1 - b) * p" using b p1 by auto finally show ?thesis by (intro diff_left_mono, simp) qed finally show ?thesis by simp qed qed (insert p1, auto) qed qed lemma poly_exp_bound: "\ p. \ x. b ^ x * of_nat x ^ deg \ p" proof - show ?thesis proof (induct deg) case 0 show ?case by (rule exI[of _ 1], intro allI, insert pow_one, auto) next case (Suc deg) then obtain q where IH: "\ x. b ^ x * (of_nat x) ^ deg \ q" by auto define p where "p = max 0 q" from IH have IH: "\ x. b ^ x * (of_nat x) ^ deg \ p" unfolding p_def using le_max_iff_disj by blast have p: "p \ 0" unfolding p_def by simp show ?case proof (cases "deg = 0") case True thus ?thesis using linear_exp_bound by simp next case False note deg = this define p' where "p' = p*p * 2 ^ Suc deg * inverse b" let ?f = "\ x. b ^ x * (of_nat x) ^ Suc deg" define f where "f = ?f" { fix x let ?x = "of_nat x :: 'a" have "f (2 * x) \ (2 ^ Suc deg) * (p * p)" proof (cases "x = 0") case False hence x1: "?x \ 1" by (cases x, auto) from x1 have x: "?x ^ (deg - 1) \ 1" by simp from x1 have xx: "?x ^ Suc deg \ 1" by (rule one_le_power) define c where "c = b ^ x * b ^ x * (2 ^ Suc deg)" have c: "c > 0" unfolding c_def using b by auto have "f (2 * x) = ?f (2 * x)" unfolding f_def by simp also have "b ^ (2 * x) = (b ^ x) * (b ^ x)" by (simp add: power2_eq_square power_even_eq) also have "of_nat (2 * x) = 2 * ?x" by simp also have "(2 * ?x) ^ Suc deg = 2 ^ Suc deg * ?x ^ Suc deg" by simp finally have "f (2 * x) = c * ?x ^ Suc deg" unfolding c_def by (simp add: ac_simps) also have "\ \ c * ?x ^ Suc deg * ?x ^ (deg - 1)" proof - have "c * ?x ^ Suc deg > 0" using c xx by simp thus ?thesis unfolding mult_le_cancel_left1 using x by simp qed also have "\ = c * ?x ^ (Suc deg + (deg - 1))" by (simp add: power_add) also have "Suc deg + (deg - 1) = deg + deg" using deg by simp also have "?x ^ (deg + deg) = (?x ^ deg) * (?x ^ deg)" by (simp add: power_add) also have "c * \ = (2 ^ Suc deg) * ((b ^ x * ?x ^ deg) * (b ^ x * ?x ^ deg))" unfolding c_def by (simp add: ac_simps) also have "\ \ (2 ^ Suc deg) * (p * p)" by (rule mult_left_mono[OF mult_mono[OF IH IH p]], insert pow_zero[of x], auto) finally show "f (2 * x) \ (2 ^ Suc deg) * (p * p)" . qed (auto simp: f_def) hence "?f (2 * x) \ (2 ^ Suc deg) * (p * p)" unfolding f_def . } note even = this show ?thesis proof (rule exI[of _ p'], intro allI) fix y show "?f y \ p'" proof (cases "even y") case True define x where "x = y div 2" have "y = 2 * x" unfolding x_def using True by simp from even[of x, folded this] have "?f y \ 2 ^ Suc deg * (p * p)" . also have "\ \ \ * inverse b" unfolding mult_le_cancel_left1 using b p by (simp add: algebra_split_simps one_le_inverse) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . next case False define x where "x = y div 2" have "y = 2 * x + 1" unfolding x_def using False by simp hence "?f y = ?f (2 * x + 1)" by simp also have "\ \ b ^ (2 * x + 1) * of_nat (2 * x + 2) ^ Suc deg" by (rule mult_left_mono[OF power_mono], insert b, auto) also have "b ^ (2 * x + 1) = b ^ (2 * x + 2) * inverse b" using b by auto also have "b ^ (2 * x + 2) * inverse b * of_nat (2 * x + 2) ^ Suc deg = inverse b * ?f (2 * (x + 1))" by (simp add: ac_simps) also have "\ \ inverse b * ((2 ^ Suc deg) * (p * p))" by (rule mult_left_mono[OF even], insert b, auto) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . qed qed qed qed qed end lemma prod_list_replicate[simp]: "prod_list (replicate n a) = a ^ n" by (induct n, auto) lemma prod_list_power: fixes xs :: "'a :: comm_monoid_mult list" shows "prod_list xs ^ n = (\x\xs. x ^ n)" by (induct xs, auto simp: power_mult_distrib) lemma set_upt_Suc: "{0 ..< Suc i} = insert i {0 ..< i}" by (fact atLeast0_lessThan_Suc) lemma prod_pow[simp]: "(\i = 0..a\ * y dvd x \ a * y dvd x" for x y a :: int using abs_dvd_iff [of "a * y"] abs_dvd_iff [of "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_right_int [simp]: "gcd x (\a\ * y) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * y"] gcd_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma lcm_abs_mult_right_int [simp]: "lcm x (\a\ * y) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * y"] lcm_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_left_int [simp]: "gcd x (a * \y\) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * \y\"] gcd_abs2_int [of _ "a * y"] by (simp add: abs_mult) lemma lcm_abs_mult_left_int [simp]: "lcm x (a * \y\) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * \y\"] lcm_abs2_int [of _ "a * y"] by (simp add: abs_mult) abbreviation (input) list_gcd :: "'a :: semiring_gcd list \ 'a" where "list_gcd \ gcd_list" abbreviation (input) list_lcm :: "'a :: semiring_gcd list \ 'a" where "list_lcm \ lcm_list" lemma list_gcd_simps: "list_gcd [] = 0" "list_gcd (x # xs) = gcd x (list_gcd xs)" by simp_all lemma list_gcd: "x \ set xs \ list_gcd xs dvd x" by (fact Gcd_fin_dvd) lemma list_gcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd (list_gcd xs)" by (fact gcd_list_greatest) lemma list_gcd_mult_int [simp]: fixes xs :: "int list" shows "list_gcd (map (times a) xs) = \a\ * list_gcd xs" - by (simp add: Gcd_mult) + by (simp add: Gcd_mult abs_mult) lemma list_lcm_simps: "list_lcm [] = 1" "list_lcm (x # xs) = lcm x (list_lcm xs)" by simp_all lemma list_lcm: "x \ set xs \ x dvd list_lcm xs" by (fact dvd_Lcm_fin) lemma list_lcm_least: "(\ x. x \ set xs \ x dvd y) \ list_lcm xs dvd y" by (fact lcm_list_least) lemma lcm_mult_distrib_nat: "(k :: nat) * lcm m n = lcm (k * m) (k * n)" by (simp add: lcm_mult_left) lemma lcm_mult_distrib_int: "abs (k::int) * lcm m n = lcm (k * m) (k * n)" - by (simp add: lcm_mult_left) + by (simp add: lcm_mult_left abs_mult) lemma list_lcm_mult_int [simp]: fixes xs :: "int list" shows "list_lcm (map (times a) xs) = (if xs = [] then 1 else \a\ * list_lcm xs)" - by (simp add: Lcm_mult) + by (simp add: Lcm_mult abs_mult) lemma list_lcm_pos: "list_lcm xs \ (0 :: int)" "0 \ set xs \ list_lcm xs \ 0" "0 \ set xs \ list_lcm xs > 0" proof - have "0 \ \Lcm (set xs)\" by (simp only: abs_ge_zero) then have "0 \ Lcm (set xs)" by simp then show "list_lcm xs \ 0" by simp assume "0 \ set xs" then show "list_lcm xs \ 0" by (simp add: Lcm_0_iff) with \list_lcm xs \ 0\ show "list_lcm xs > 0" by (simp add: le_less) qed lemma quotient_of_nonzero: "snd (quotient_of r) > 0" "snd (quotient_of r) \ 0" using quotient_of_denom_pos' [of r] by simp_all lemma quotient_of_int_div: assumes q: "quotient_of (of_int x / of_int y) = (a, b)" and y: "y \ 0" shows "\ z. z \ 0 \ x = a * z \ y = b * z" proof - let ?r = "rat_of_int" define z where "z = gcd x y" define x' where "x' = x div z" define y' where "y' = y div z" have id: "x = z * x'" "y = z * y'" unfolding x'_def y'_def z_def by auto from y have y': "y' \ 0" unfolding id by auto have z: "z \ 0" unfolding z_def using y by auto have cop: "coprime x' y'" unfolding x'_def y'_def z_def using div_gcd_coprime y by blast have "?r x / ?r y = ?r x' / ?r y'" unfolding id using z y y' by (auto simp: field_simps) from assms[unfolded this] have quot: "quotient_of (?r x' / ?r y') = (a, b)" by auto from quotient_of_coprime[OF quot] have cop': "coprime a b" . hence cop: "coprime b a" by (simp add: ac_simps) from quotient_of_denom_pos[OF quot] have b: "b > 0" "b \ 0" by auto from quotient_of_div[OF quot] quotient_of_denom_pos[OF quot] y' have "?r x' * ?r b = ?r a * ?r y'" by (auto simp: field_simps) hence id': "x' * b = a * y'" unfolding of_int_mult[symmetric] by linarith from id'[symmetric] have "b dvd y' * a" unfolding mult.commute[of y'] by auto with cop y' have "b dvd y'" by (simp add: coprime_dvd_mult_left_iff) then obtain z' where ybz: "y' = b * z'" unfolding dvd_def by auto from id[unfolded y' this] have y: "y = b * (z * z')" by auto with \y \ 0\ have zz: "z * z' \ 0" by auto from quotient_of_div[OF q] \y \ 0\ \b \ 0\ have "?r x * ?r b = ?r y * ?r a" by (auto simp: field_simps) hence id': "x * b = y * a" unfolding of_int_mult[symmetric] by linarith from this[unfolded y] b have x: "x = a * (z * z')" by auto show ?thesis unfolding x y using zz by blast qed fun max_list_non_empty :: "('a :: linorder) list \ 'a" where "max_list_non_empty [x] = x" | "max_list_non_empty (x # xs) = max x (max_list_non_empty xs)" lemma max_list_non_empty: "x \ set xs \ x \ max_list_non_empty xs" proof (induct xs) case (Cons y ys) note oCons = this show ?case proof (cases ys) case (Cons z zs) hence id: "max_list_non_empty (y # ys) = max y (max_list_non_empty ys)" by simp from oCons show ?thesis unfolding id by (auto simp: max.coboundedI2) qed (insert oCons, auto) qed simp lemma cnj_reals[simp]: "(cnj c \ \) = (c \ \)" using Reals_cnj_iff by fastforce lemma sgn_real_mono: "x \ y \ sgn x \ sgn (y :: real)" unfolding sgn_real_def by (auto split: if_splits) lemma sgn_minus_rat: "sgn (- (x :: rat)) = - sgn x" by (fact Rings.sgn_minus) lemma real_of_rat_sgn: "sgn (of_rat x) = real_of_rat (sgn x)" unfolding sgn_real_def sgn_rat_def by auto lemma inverse_le_iff_sgn: assumes sgn: "sgn x = sgn y" shows "(inverse (x :: real) \ inverse y) = (y \ x)" proof (cases "x = 0") case True with sgn have "sgn y = 0" by simp hence "y = 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0"; auto) thus ?thesis using True by simp next case False note x = this show ?thesis proof (cases "x < 0") case True with x sgn have "sgn y = -1" by simp hence "y < 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le_neg[OF True \y < 0\]) next case False with x have x: "x > 0" by auto with sgn have "sgn y = 1" by auto hence "y > 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le[OF x \y > 0\]) qed qed lemma inverse_le_sgn: assumes sgn: "sgn x = sgn y" and xy: "x \ (y :: real)" shows "inverse y \ inverse x" using xy inverse_le_iff_sgn[OF sgn] by auto lemma set_list_update: "set (xs [i := k]) = (if i < length xs then insert k (set (take i xs) \ set (drop (Suc i) xs)) else set xs)" proof (induct xs arbitrary: i) case (Cons x xs i) thus ?case by (cases i, auto) qed simp lemma prod_list_dvd: assumes "(x :: 'a :: comm_monoid_mult) \ set xs" shows "x dvd prod_list xs" proof - from assms[unfolded in_set_conv_decomp] obtain ys zs where xs: "xs = ys @ x # zs" by auto show ?thesis unfolding xs dvd_def by (intro exI[of _ "prod_list (ys @ zs)"], simp add: ac_simps) qed lemma dvd_prod: fixes A::"'b set" assumes "\b\A. a dvd f b" "finite A" shows "a dvd prod f A" using assms(2,1) proof (induct A) case (insert x A) thus ?case using comm_monoid_mult_class.dvd_mult dvd_mult2 insert_iff prod.insert by auto qed auto context fixes xs :: "'a :: comm_monoid_mult list" begin lemma prod_list_filter: "prod_list (filter f xs) * prod_list (filter (\ x. \ f x) xs) = prod_list xs" by (induct xs, auto simp: ac_simps) lemma prod_list_partition: assumes "partition f xs = (ys, zs)" shows "prod_list xs = prod_list ys * prod_list zs" using assms by (subst prod_list_filter[symmetric, of f], auto simp: o_def) end lemma dvd_imp_mult_div_cancel_left[simp]: assumes "(a :: 'a :: semidom_divide) dvd b" shows "a * (b div a) = b" proof(cases "b = 0") case True then show ?thesis by auto next case False with dvdE[OF assms] obtain c where *: "b = a * c" by auto also with False have "a \ 0" by auto then have "a * c div a = c" by auto also note *[symmetric] finally show ?thesis. qed lemma (in semidom) prod_list_zero_iff[simp]: "prod_list xs = 0 \ 0 \ set xs" by (induction xs, auto) context comm_monoid_mult begin lemma unit_prod [intro]: shows "a dvd 1 \ b dvd 1 \ (a * b) dvd 1" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff[simp]: shows "(a * b) dvd 1 \ a dvd 1 \ b dvd 1" by (auto dest: dvd_mult_left dvd_mult_right) end context comm_semiring_1 begin lemma irreducibleE[elim]: assumes "irreducible p" and "p \ 0 \ \ p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ thesis" shows thesis using assms by (auto simp: irreducible_def) lemma not_irreducibleE: assumes "\ irreducible x" and "x = 0 \ thesis" and "x dvd 1 \ thesis" and "\a b. x = a * b \ \ a dvd 1 \ \ b dvd 1 \ thesis" shows thesis using assms unfolding irreducible_def by auto lemma prime_elem_dvd_prod_list: assumes p: "prime_elem p" and pA: "p dvd prod_list A" shows "\a \ set A. p dvd a" proof(insert pA, induct A) case Nil with p show ?case by (simp add: prime_elem_not_unit) next case (Cons a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma prime_elem_dvd_prod_mset: assumes p: "prime_elem p" and pA: "p dvd prod_mset A" shows "\a \# A. p dvd a" proof(insert pA, induct A) case empty with p show ?case by (simp add: prime_elem_not_unit) next case (add a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma mult_unit_dvd_iff[simp]: assumes "b dvd 1" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" using dvd_mult_left[of a b c] by simp next assume "a dvd c" with assms mult_dvd_mono show "a * b dvd c" by fastforce qed lemma mult_unit_dvd_iff'[simp]: "a dvd 1 \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ b dvd 1" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "b dvd 1 \ c dvd 1" . thus ?thesis by (auto simp: c) qed end context idom begin text \ Following lemmas are adapted and generalized so that they don't use "algebraic" classes. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (auto simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma irreducibleI': assumes "a \ 0" "\ a dvd 1" "\b. b dvd a \ a dvd b \ b dvd 1" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ b dvd 1" by (intro assms) simp_all thus "b dvd 1 \ c dvd 1" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis using dvd_times_left_cancel_iff by fastforce qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: shows "irreducible x \ x \ 0 \ \ x dvd 1 \ (\b. b dvd x \ x dvd b \ b dvd 1)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma dvd_mult_unit_iff: assumes b: "b dvd 1" shows "a dvd c * b \ a dvd c" proof- from b obtain b' where 1: "b * b' = 1" by (elim dvdE, auto) then have b0: "b \ 0" by auto from 1 have "a = (a * b') * b" by (simp add: ac_simps) also have "\ dvd c * b \ a * b' dvd c" using b0 by auto finally show ?thesis by (auto intro: dvd_mult_left) qed lemma dvd_mult_unit_iff': "b dvd 1 \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma irreducible_mult_unit_left: shows "a dvd 1 \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma irreducible_mult_unit_right: shows "a dvd 1 \ irreducible (p * a) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma unit_imp_dvd [dest]: "b dvd 1 \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_mult_left_cancel: "a dvd 1 \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "a dvd 1 \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) text \New parts from here\ lemma irreducible_multD: assumes l: "irreducible (a*b)" shows "a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" proof- from l have "a dvd 1 \ b dvd 1" using irreducibleD by auto then show ?thesis proof(elim disjE) assume a: "a dvd 1" with l have "irreducible b" unfolding irreducible_def by (metis is_unit_mult_iff mult.left_commute mult_not_zero) with a show ?thesis by auto next assume a: "b dvd 1" with l have "irreducible a" unfolding irreducible_def by (meson is_unit_mult_iff mult_not_zero semiring_normalization_rules(16)) with a show ?thesis by auto qed qed end lemma (in field) irreducible_field[simp]: "irreducible x \ False" by (auto simp: dvd_field_iff irreducible_def) lemma (in idom) irreducible_mult: shows "irreducible (a*b) \ a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" by (auto dest: irreducible_multD simp: irreducible_mult_unit_left irreducible_mult_unit_right) end diff --git a/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy b/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy --- a/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy +++ b/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy @@ -1,450 +1,450 @@ (* Author: RenĂ© Thiemann Akihisa Yamada License: BSD *) section \Connecting Polynomials with Homomorphism Locales\ theory Ring_Hom_Poly imports "HOL-Computational_Algebra.Euclidean_Algorithm" Ring_Hom Missing_Polynomial begin text\@{term poly} as a homomorphism. Note that types differ.\ interpretation poly_hom: comm_semiring_hom "\p. poly p a" by (unfold_locales, auto) interpretation poly_hom: comm_ring_hom "\p. poly p a".. interpretation poly_hom: idom_hom "\p. poly p a".. text \@{term "(\\<^sub>p)"} as a homomorphism.\ interpretation pcompose_hom: comm_semiring_hom "\q. q \\<^sub>p p" using pcompose_add pcompose_mult pcompose_1 by (unfold_locales, auto) interpretation pcompose_hom: comm_ring_hom "\q. q \\<^sub>p p" .. interpretation pcompose_hom: idom_hom "\q. q \\<^sub>p p" .. definition eval_poly :: "('a \ 'b :: comm_semiring_1) \ 'a :: zero poly \ 'b \ 'b" where [code del]: "eval_poly h p = poly (map_poly h p)" lemma eval_poly_code[code]: "eval_poly h p x = fold_coeffs (\ a b. h a + x * b) p 0" by (induct p, auto simp: eval_poly_def) lemma eval_poly_as_sum: fixes h :: "'a :: zero \ 'b :: comm_semiring_1" assumes "h 0 = 0" shows "eval_poly h p x = (\i\degree p. x^i * h (coeff p i))" unfolding eval_poly_def proof (induct p) case 0 show ?case using assms by simp next case (pCons a p) thus ?case proof (cases "p = 0") case True show ?thesis by (simp add: True map_poly_simps assms) next case False show ?thesis unfolding degree_pCons_eq[OF False] unfolding sum.atMost_Suc_shift unfolding map_poly_pCons[OF pCons(1)] by (simp add: pCons(2) sum_distrib_left mult.assoc) qed qed lemma coeff_const: "coeff [: a :] i = (if i = 0 then a else 0)" by (metis coeff_monom monom_0) lemma x_as_monom: "[:0,1:] = monom 1 1" by (simp add: monom_0 monom_Suc) lemma x_pow_n: "monom 1 1 ^ n = monom 1 n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma map_poly_eval_poly: assumes h0: "h 0 = 0" shows "map_poly h p = eval_poly (\ a. [: h a :]) p [:0,1:]" (is "?mp = ?ep") proof (rule poly_eqI) fix i :: nat have 2: "(\x\i. \xa\degree p. (if xa = x then 1 else 0) * coeff [:h (coeff p xa):] (i - x)) = h (coeff p i)" (is "sum ?f ?s = ?r") proof - have "sum ?f ?s = ?f i + sum ?f ({..i} - {i})" by (rule sum.remove[of _ i], auto) also have "sum ?f ({..i} - {i}) = 0" by (rule sum.neutral, intro ballI, rule sum.neutral, auto simp: coeff_const) also have "?f i = (\xa\degree p. (if xa = i then 1 else 0) * h (coeff p xa))" (is "_ = ?m") unfolding coeff_const by simp also have "\ = ?r" proof (cases "i \ degree p") case True show ?thesis by (subst sum.remove[of _ i], insert True, auto) next case False hence [simp]: "coeff p i = 0" using le_degree by blast show ?thesis by (subst sum.neutral, auto simp: h0) qed finally show ?thesis by simp qed have h'0: "[: h 0 :] = 0" using h0 by auto show "coeff ?mp i = coeff ?ep i" unfolding coeff_map_poly[of h, OF h0] unfolding eval_poly_as_sum[of "\a. [: h a :]", OF h'0] unfolding coeff_sum unfolding x_as_monom x_pow_n coeff_mult unfolding sum.swap[of _ _ "{..degree p}"] unfolding coeff_monom using 2 by auto qed lemma smult_as_map_poly: "smult a = map_poly ((*) a)" by (rule ext, rule poly_eqI, subst coeff_map_poly, auto) subsection \@{const map_poly} of Homomorphisms\ context zero_hom begin text \We will consider @{term hom} is always simpler than @{term "map_poly hom"}.\ lemma map_poly_hom_monom[simp]: "map_poly hom (monom a i) = monom (hom a) i" by(rule map_poly_monom, auto) lemma coeff_map_poly_hom[simp]: "coeff (map_poly hom p) i = hom (coeff p i)" by (rule coeff_map_poly, rule hom_zero) end locale map_poly_zero_hom = base: zero_hom begin sublocale zero_hom "map_poly hom" by (unfold_locales, auto) end text \@{const map_poly} preserves homomorphisms over addition.\ context comm_monoid_add_hom begin lemma map_poly_hom_add[hom_distribs]: "map_poly hom (p + q) = map_poly hom p + map_poly hom q" by (rule map_poly_add; simp add: hom_distribs) end locale map_poly_comm_monoid_add_hom = base: comm_monoid_add_hom begin sublocale comm_monoid_add_hom "map_poly hom" by (unfold_locales, auto simp:hom_distribs) end text \To preserve homomorphisms over multiplication, it demands commutative ring homomorphisms.\ context comm_semiring_hom begin lemma map_poly_pCons_hom[hom_distribs]: "map_poly hom (pCons a p) = pCons (hom a) (map_poly hom p)" unfolding map_poly_simps by auto lemma map_poly_hom_smult[hom_distribs]: "map_poly hom (smult c p) = smult (hom c) (map_poly hom p)" by (induct p, auto simp: hom_distribs) lemma poly_map_poly[simp]: "poly (map_poly hom p) (hom x) = hom (poly p x)" by (induct p; simp add: hom_distribs) end locale map_poly_comm_semiring_hom = base: comm_semiring_hom begin sublocale map_poly_comm_monoid_add_hom.. sublocale comm_semiring_hom "map_poly hom" proof show "map_poly hom 1 = 1" by simp fix p q show "map_poly hom (p * q) = map_poly hom p * map_poly hom q" by (induct p, auto simp: hom_distribs) qed end locale map_poly_comm_ring_hom = base: comm_ring_hom begin sublocale map_poly_comm_semiring_hom.. sublocale comm_ring_hom "map_poly hom".. end locale map_poly_idom_hom = base: idom_hom begin sublocale map_poly_comm_ring_hom.. sublocale idom_hom "map_poly hom".. end subsubsection \Injectivity\ locale map_poly_inj_zero_hom = base: inj_zero_hom begin sublocale inj_zero_hom "map_poly hom" proof (unfold_locales) fix p q :: "'a poly" assume "map_poly hom p = map_poly hom q" from cong[of "\p. coeff p _", OF refl this] show "p = q" by (auto intro: poly_eqI) qed simp end locale map_poly_inj_comm_monoid_add_hom = base: inj_comm_monoid_add_hom begin sublocale map_poly_comm_monoid_add_hom.. sublocale map_poly_inj_zero_hom.. sublocale inj_comm_monoid_add_hom "map_poly hom".. end locale map_poly_inj_comm_semiring_hom = base: inj_comm_semiring_hom begin sublocale map_poly_comm_semiring_hom.. sublocale map_poly_inj_zero_hom.. sublocale inj_comm_semiring_hom "map_poly hom".. end locale map_poly_inj_comm_ring_hom = base: inj_comm_ring_hom begin sublocale map_poly_inj_comm_semiring_hom.. sublocale inj_comm_ring_hom "map_poly hom".. end locale map_poly_inj_idom_hom = base: inj_idom_hom begin sublocale map_poly_inj_comm_ring_hom.. sublocale inj_idom_hom "map_poly hom".. end lemma degree_map_poly_le: "degree (map_poly f p) \ degree p" by(induct p;auto) lemma coeffs_map_poly: (* An exact variant *) assumes "f (lead_coeff p) = 0 \ p = 0" shows "coeffs (map_poly f p) = map f (coeffs p)" unfolding coeffs_map_poly using assms by (simp add:coeffs_def) lemma degree_map_poly: (* An exact variant *) assumes "f (lead_coeff p) = 0 \ p = 0" shows "degree (map_poly f p) = degree p" unfolding degree_eq_length_coeffs unfolding coeffs_map_poly[of f, OF assms] by simp context zero_hom_0 begin lemma degree_map_poly_hom[simp]: "degree (map_poly hom p) = degree p" by (rule degree_map_poly, auto) lemma coeffs_map_poly_hom[simp]: "coeffs (map_poly hom p) = map hom (coeffs p)" by (rule coeffs_map_poly, auto) lemma hom_lead_coeff[simp]: "lead_coeff (map_poly hom p) = hom (lead_coeff p)" by simp end context comm_semiring_hom begin interpretation map_poly_hom: map_poly_comm_semiring_hom.. lemma poly_map_poly_0[simp]: "poly (map_poly hom p) 0 = hom (poly p 0)" (is "?l = ?r") proof- have "?l = poly (map_poly hom p) (hom 0)" by auto then show ?thesis unfolding poly_map_poly. qed lemma poly_map_poly_1[simp]: "poly (map_poly hom p) 1 = hom (poly p 1)" (is "?l = ?r") proof- have "?l = poly (map_poly hom p) (hom 1)" by auto then show ?thesis unfolding poly_map_poly. qed lemma map_poly_hom_as_monom_sum: "(\j\degree p. monom (hom (coeff p j)) j) = map_poly hom p" proof - show ?thesis by (subst(6) poly_as_sum_of_monoms'[OF le_refl, symmetric], simp add: hom_distribs) qed lemma map_poly_pcompose[hom_distribs]: "map_poly hom (f \\<^sub>p g) = map_poly hom f \\<^sub>p map_poly hom g" by (induct f arbitrary: g; auto simp: hom_distribs) end context comm_semiring_hom begin lemma eval_poly_0[simp]: "eval_poly hom 0 x = 0" unfolding eval_poly_def by simp lemma eval_poly_monom: "eval_poly hom (monom a n) x = hom a * x ^ n" unfolding eval_poly_def unfolding map_poly_monom[of hom, OF hom_zero] using poly_monom. lemma poly_map_poly_eval_poly: "poly (map_poly hom p) = eval_poly hom p" unfolding eval_poly_def.. lemma map_poly_eval_poly: "map_poly hom p = eval_poly (\ a. [: hom a :]) p [:0,1:]" by (rule map_poly_eval_poly, simp) lemma degree_extension: assumes "degree p \ n" shows "(\i\degree p. x ^ i * hom (coeff p i)) = (\i\n. x ^ i * hom (coeff p i))" (is "?l = ?r") proof - let ?f = "\ i. x ^ i * hom (coeff p i)" define m where "m = n - degree p" have n: "n = degree p + m" unfolding m_def using assms by auto have "?r = (\ i \ degree p + m. ?f i)" unfolding n .. also have "\ = ?l + sum ?f {Suc (degree p) .. degree p + m}" by (subst sum.union_disjoint[symmetric], auto intro: sum.cong) also have "sum ?f {Suc (degree p) .. degree p + m} = 0" by (rule sum.neutral, auto simp: coeff_eq_0) finally show ?thesis by simp qed lemma eval_poly_add[simp]: "eval_poly hom (p + q) x = eval_poly hom p x + eval_poly hom q x" unfolding eval_poly_def hom_distribs.. lemma eval_poly_sum: "eval_poly hom (\k\A. p k) x = (\k\A. eval_poly hom (p k) x)" proof (induct A rule: infinite_finite_induct) case (insert a A) show ?case unfolding sum.insert[OF insert(1-2)] insert(3)[symmetric] by simp qed (auto simp: eval_poly_def) lemma eval_poly_poly: "eval_poly hom p (hom x) = hom (poly p x)" unfolding eval_poly_def by auto end context comm_ring_hom begin interpretation map_poly_hom: map_poly_comm_ring_hom.. lemma pseudo_divmod_main_hom: "pseudo_divmod_main (hom lc) (map_poly hom q) (map_poly hom r) (map_poly hom d) dr i = map_prod (map_poly hom) (map_poly hom) (pseudo_divmod_main lc q r d dr i)" proof- show ?thesis by (induct lc q r d dr i rule:pseudo_divmod_main.induct, auto simp: Let_def hom_distribs) qed end lemma(in inj_comm_ring_hom) pseudo_divmod_hom: "pseudo_divmod (map_poly hom p) (map_poly hom q) = map_prod (map_poly hom) (map_poly hom) (pseudo_divmod p q)" unfolding pseudo_divmod_def using pseudo_divmod_main_hom[of _ 0] by (cases "q = 0",auto) lemma(in inj_idom_hom) pseudo_mod_hom: "pseudo_mod (map_poly hom p) (map_poly hom q) = map_poly hom (pseudo_mod p q)" using pseudo_divmod_hom unfolding pseudo_mod_def by auto lemma(in idom_hom) map_poly_pderiv[hom_distribs]: "map_poly hom (pderiv p) = pderiv (map_poly hom p)" proof (induct p rule: pderiv.induct) case (1 a p) then show ?case unfolding pderiv.simps map_poly_pCons_hom by (cases "p = 0", auto simp: hom_distribs) qed context field_hom begin lemma map_poly_pdivmod[hom_distribs]: "map_prod (map_poly hom) (map_poly hom) (p div q, p mod q) = (map_poly hom p div map_poly hom q, map_poly hom p mod map_poly hom q)" (is "?l = ?r") proof - let ?mp = "map_poly hom" interpret map_poly_hom: map_poly_idom_hom.. obtain r s where dm: "(p div q, p mod q) = (r, s)" by force hence r: "r = p div q" and s: "s = p mod q" by simp_all from dm [folded pdivmod_pdivmodrel] have "eucl_rel_poly p q (r, s)" by auto from this[unfolded eucl_rel_poly_iff] have eq: "p = r * q + s" and cond: "(if q = 0 then r = 0 else s = 0 \ degree s < degree q)" by auto from arg_cong[OF eq, of ?mp, unfolded map_poly_add] have eq: "?mp p = ?mp q * ?mp r + ?mp s" by (auto simp: hom_distribs) from cond have cond: "(if ?mp q = 0 then ?mp r = 0 else ?mp s = 0 \ degree (?mp s) < degree (?mp q))" by simp from eq cond have "eucl_rel_poly (?mp p) (?mp q) (?mp r, ?mp s)" unfolding eucl_rel_poly_iff by auto from this[unfolded pdivmod_pdivmodrel] show ?thesis unfolding dm prod.simps by simp qed lemma map_poly_div[hom_distribs]: "map_poly hom (p div q) = map_poly hom p div map_poly hom q" using map_poly_pdivmod[of p q] by simp lemma map_poly_mod[hom_distribs]: "map_poly hom (p mod q) = map_poly hom p mod map_poly hom q" using map_poly_pdivmod[of p q] by simp end locale field_hom' = field_hom hom - for hom :: "'a :: {field,euclidean_ring_gcd} \ 'b :: {field,euclidean_ring_gcd}" + for hom :: "'a :: {field_gcd} \ 'b :: {field_gcd}" begin lemma map_poly_normalize[hom_distribs]: "map_poly hom (normalize p) = normalize (map_poly hom p)" by (simp add: normalize_poly_def hom_distribs) lemma map_poly_gcd[hom_distribs]: "map_poly hom (gcd p q) = gcd (map_poly hom p) (map_poly hom q)" by (induct p q rule: eucl_induct) (simp_all add: map_poly_normalize ac_simps hom_distribs) end definition div_poly :: "'a :: euclidean_semiring \ 'a poly \ 'a poly" where "div_poly a p = map_poly (\ c. c div a) p" lemma smult_div_poly: assumes "\ c. c \ set (coeffs p) \ a dvd c" shows "smult a (div_poly a p) = p" unfolding smult_as_map_poly div_poly_def by (subst map_poly_map_poly, force, subst map_poly_idI, insert assms, auto) lemma coeff_div_poly: "coeff (div_poly a f) n = coeff f n div a" unfolding div_poly_def by (rule coeff_map_poly, auto) locale map_poly_inj_idom_divide_hom = base: inj_idom_divide_hom begin sublocale map_poly_idom_hom .. sublocale map_poly_inj_zero_hom .. sublocale inj_idom_hom "map_poly hom" .. lemma divide_poly_main_hom: defines "hh \ map_poly hom" shows "hh (divide_poly_main lc f g h i j) = divide_poly_main (hom lc) (hh f) (hh g) (hh h) i j" unfolding hh_def proof (induct j arbitrary: lc f g h i) case (Suc j lc f g h i) let ?h = "map_poly hom" show ?case unfolding divide_poly_main.simps Let_def unfolding base.coeff_map_poly_hom base.hom_div[symmetric] base.hom_mult[symmetric] base.eq_iff if_distrib[of ?h] hom_zero by (rule if_cong[OF refl _ refl], subst Suc, simp add: hom_minus hom_add hom_mult) qed simp sublocale inj_idom_divide_hom "map_poly hom" proof fix f g :: "'a poly" let ?h = "map_poly hom" show "?h (f div g) = (?h f) div (?h g)" unfolding divide_poly_def if_distrib[of ?h] divide_poly_main_hom by simp qed lemma order_hom: "order (hom x) (map_poly hom f) = order x f" unfolding Polynomial.order_def unfolding hom_dvd_iff[symmetric] unfolding hom_power by (simp add: base.hom_uminus) end subsection \Example Interpretations\ abbreviation "of_int_poly \ map_poly of_int" interpretation of_int_poly_hom: map_poly_comm_semiring_hom of_int.. interpretation of_int_poly_hom: map_poly_comm_ring_hom of_int.. interpretation of_int_poly_hom: map_poly_idom_hom of_int.. interpretation of_int_poly_hom: map_poly_inj_comm_ring_hom "of_int :: int \ 'a :: {comm_ring_1,ring_char_0}" .. interpretation of_int_poly_hom: map_poly_inj_idom_hom "of_int :: int \ 'a :: {idom,ring_char_0}" .. text \The following operations are homomorphic w.r.t. only @{class monoid_add}.\ interpretation pCons_0_hom: injective "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: zero_hom_0 "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: inj_comm_monoid_add_hom "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: inj_ab_group_add_hom "pCons 0" by (unfold_locales, auto) interpretation monom_hom: injective "\x. monom x d" by (unfold_locales, auto) interpretation monom_hom: inj_monoid_add_hom "\x. monom x d" by (unfold_locales, auto simp: add_monom) interpretation monom_hom: inj_comm_monoid_add_hom "\x. monom x d".. end diff --git a/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy b/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy --- a/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy +++ b/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy @@ -1,477 +1,480 @@ (* File: Algebraic_Auxiliaries.thy Authors: Daniel StĂ¼we Miscellaneous facts about algebra and number theory *) section \Auxiliary Material\ theory Algebraic_Auxiliaries imports "HOL-Algebra.Algebra" "HOL-Computational_Algebra.Squarefree" "HOL-Number_Theory.Number_Theory" begin hide_const (open) Divisibility.prime lemma sum_of_bool_eq_card: assumes "finite S" shows "(\a \ S. of_bool (P a)) = real (card {a \ S . P a })" proof - have "(\a \ S. of_bool (P a) :: real) = (\a \ {x\S. P x}. 1)" using assms by (intro sum.mono_neutral_cong_right) auto thus ?thesis by simp qed lemma mod_natE: fixes a n b :: nat assumes "a mod n = b" shows "\ l. a = n * l + b" using assms mod_mult_div_eq[of a n] by (metis add.commute) lemma (in group) r_coset_is_image: "H #> a = (\ x. x \ a) ` H" unfolding r_coset_def image_def by blast lemma (in group) FactGroup_order: assumes "subgroup H G" "finite H" shows "order G = order (G Mod H) * card H" using lagrange assms unfolding FactGroup_def order_def by simp corollary (in group) FactGroup_order_div: assumes "subgroup H G" "finite H" shows "order (G Mod H) = order G div card H" using assms FactGroup_order subgroupE(2)[OF \subgroup H G\] by (auto simp: order_def) lemma group_hom_imp_group_hom_image: assumes "group_hom G G h" shows "group_hom G (G\carrier := h ` carrier G\) h" using group_hom.axioms[OF assms] group_hom.img_is_subgroup[OF assms] group.subgroup_imp_group by(auto intro!: group_hom.intro simp: group_hom_axioms_def hom_def) theorem homomorphism_thm: assumes "group_hom G G h" shows "G Mod kernel G (G\carrier := h ` carrier G\) h \ G \carrier := h ` carrier G\" by (intro group_hom.FactGroup_iso group_hom_imp_group_hom_image assms) simp lemma is_iso_imp_same_card: assumes "H \ G " shows "order H = order G" proof - from assms obtain h where "bij_betw h (carrier H) (carrier G)" unfolding is_iso_def iso_def by blast then show ?thesis unfolding order_def by (rule bij_betw_same_card) qed corollary homomorphism_thm_order: assumes "group_hom G G h" shows "order (G\carrier := h ` carrier G\) * card (kernel G (G\carrier := h ` carrier G\) h) = order G " proof - have "order (G\carrier := h ` carrier G\) = order (G Mod (kernel G (G\carrier := h ` carrier G\) h))" using is_iso_imp_same_card[OF homomorphism_thm] \group_hom G G h\ by fastforce moreover have "group G" using \group_hom G G h\ group_hom.axioms by blast ultimately show ?thesis using \group_hom G G h\ and group_hom_imp_group_hom_image[OF \group_hom G G h\] unfolding FactGroup_def by (simp add: group.lagrange group_hom.subgroup_kernel order_def) qed lemma (in group_hom) kernel_subset: "kernel G H h \ carrier G" using subgroup_kernel G.subgroupE(1) by blast lemma (in group) proper_subgroup_imp_bound_on_card: assumes "H \ carrier G" "subgroup H G" "finite (carrier G)" shows "card H \ order G div 2" proof - from \finite (carrier G)\ have "finite (rcosets H)" by (simp add: RCOSETS_def) note subgroup.subgroup_in_rcosets[OF \subgroup H G\ is_group] then obtain J where "J \ H" "J \ rcosets H" using rcosets_part_G[OF \subgroup H G\] and \H \ carrier G\ by (metis Sup_le_iff inf.absorb_iff2 inf.idem inf.strict_order_iff) then have "2 \ card (rcosets H)" using \H \ rcosets H\ card_mono[OF \finite (rcosets H)\, of "{H, J}"] by simp then show ?thesis using mult_le_mono[of 2 "card (rcosets H)" "card H" "card H"] unfolding lagrange[OF \subgroup H G\] by force qed lemma cong_exp_trans[trans]: "[a ^ b = c] (mod n) \ [a = d] (mod n) \ [d ^ b = c] (mod n)" "[c = a ^ b] (mod n) \ [a = d] (mod n) \ [c = d ^ b] (mod n)" using cong_pow cong_sym cong_trans by blast+ lemma cong_exp_mod[simp]: "[(a mod n) ^ b = c] (mod n) \ [a ^ b = c] (mod n)" "[c = (a mod n) ^ b] (mod n) \ [c = a ^ b] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_mult_mod[simp]: "[(a mod n) * b = c] (mod n) \ [a * b = c] (mod n)" "[a * (b mod n) = c] (mod n) \ [a * b = c] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_add_mod[simp]: "[(a mod n) + b = c] (mod n) \ [a + b = c] (mod n)" "[a + (b mod n) = c] (mod n) \ [a + b = c] (mod n)" "[\i\A. f i mod n = c] (mod n) \ [\i\A. f i = c] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_add_trans[trans]: "[a = b + x] (mod n) \ [x = y] (mod n) \ [a = b + y] (mod n)" "[a = x + b] (mod n) \ [x = y] (mod n) \ [a = y + b] (mod n)" "[b + x = a] (mod n) \ [x = y] (mod n) \ [b + y = a] (mod n)" "[x + b = a] (mod n) \ [x = y] (mod n) \ [y + b = a] (mod n)" unfolding cong_def using mod_simps(1, 2) by metis+ lemma cong_mult_trans[trans]: "[a = b * x] (mod n) \ [x = y] (mod n) \ [a = b * y] (mod n)" "[a = x * b] (mod n) \ [x = y] (mod n) \ [a = y * b] (mod n)" "[b * x = a] (mod n) \ [x = y] (mod n) \ [b * y = a] (mod n)" "[x * b = a] (mod n) \ [x = y] (mod n) \ [y * b = a] (mod n)" unfolding cong_def using mod_simps(4, 5) by metis+ lemma cong_diff_trans[trans]: "[a = b - x] (mod n) \ [x = y] (mod n) \ [a = b - y] (mod n)" "[a = x - b] (mod n) \ [x = y] (mod n) \ [a = y - b] (mod n)" "[b - x = a] (mod n) \ [x = y] (mod n) \ [b - y = a] (mod n)" "[x - b = a] (mod n) \ [x = y] (mod n) \ [y - b = a] (mod n)" for a :: "'a :: {unique_euclidean_semiring, euclidean_ring_cancel}" unfolding cong_def by (metis mod_diff_eq)+ lemma eq_imp_eq_mod_int: "a = b \ [a = b] (mod m)" for a b :: int by simp lemma eq_imp_eq_mod_nat: "a = b \ [a = b] (mod m)" for a b :: nat by simp lemma cong_pow_I: "a = b \ [x^a = x^b](mod n)" by simp lemma gre1I: "(n = 0 \ False) \ (1 :: nat) \ n" by presburger lemma gre1I_nat: "(n = 0 \ False) \ (Suc 0 :: nat) \ n" by presburger lemma totient_less_not_prime: assumes "\ prime n" "1 < n" shows "totient n < n - 1" using totient_imp_prime totient_less assms by (metis One_nat_def Suc_pred le_less_trans less_SucE zero_le_one) lemma power2_diff_nat: "x \ y \ (x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" for x y :: nat by (simp add: algebra_simps power2_eq_square mult_2_right) (meson Nat.diff_diff_right le_add2 le_trans mult_le_mono order_refl) lemma square_inequality: "1 < n \ (n + n) \ (n * n)" for n :: nat by (metis Suc_eq_plus1_left Suc_leI mult_le_mono1 semiring_normalization_rules(4)) lemma square_one_cong_one: assumes "[x = 1](mod n)" shows "[x^2 = 1](mod n)" using assms cong_pow by fastforce lemma cong_square_alt_int: "prime p \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = p - 1] (mod p)" for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}" using dvd_add_triv_right_iff[of p "a - (p - 1)"] by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest!: prime_dvd_multD) lemma cong_square_alt: "prime p \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = p - 1] (mod p)" for a p :: nat using cong_square_alt_int and cong_int_iff prime_nat_int_transfer by (metis (mono_tags) int_ops(2) int_ops(7) less_imp_le_nat of_nat_diff prime_gt_1_nat) lemma square_minus_one_cong_one: fixes n x :: nat assumes "1 < n" "[x = n - 1](mod n)" shows "[x^2 = 1](mod n)" proof - have "[x^2 = (n - 1) * (n - 1)] (mod n)" using cong_mult[OF assms(2) assms(2)] by (simp add: algebra_simps power2_eq_square) also have "[(n - 1) * (n - 1) = Suc (n * n) - (n + n)] (mod n)" using power2_diff_nat[of 1 n] \1 < n\ by (simp add: algebra_simps power2_eq_square) also have "[Suc (n * n) - (n + n) = Suc (n * n)] (mod n)" proof - have "n * n + 0 * n = n * n" by linarith moreover have "n * n - (n + n) + (n + n) = n * n" using square_inequality[OF \1 < n\] le_add_diff_inverse2 by blast moreover have "(Suc 0 + 1) * n = n + n" by simp ultimately show ?thesis using square_inequality[OF \1 < n\] by (metis (no_types) Suc_diff_le add_Suc cong_iff_lin_nat) qed also have "[Suc (n * n) = 1] (mod n)" using cong_to_1'_nat by auto finally show ?thesis . qed lemma odd_prime_gt_2_int: "2 < p" if "odd p" "prime p" for p :: int using prime_ge_2_int[OF \prime p\] \odd p\ by (cases "p = 2") auto lemma odd_prime_gt_2_nat: "2 < p" if "odd p" "prime p" for p :: nat using prime_ge_2_nat[OF \prime p\] \odd p\ by (cases "p = 2") auto lemma gt_one_imp_gt_one_power_if_coprime: "1 \ x \ 1 < n \ coprime x n \ 1 \ x ^ (n - 1) mod n" by (rule gre1I) (auto simp: coprime_commute dest: coprime_absorb_left) lemma residue_one_dvd: "a mod n = 1 \ n dvd a - 1" for a n :: nat by (fastforce intro!: cong_to_1_nat simp: cong_def) lemma coprimeI_power_mod: fixes x r n :: nat assumes "x ^ r mod n = 1" "r \ 0" "n \ 0" shows "coprime x n" proof - have "coprime (x ^ r mod n) n" using coprime_1_right \x ^ r mod n = 1\ by (simp add: coprime_commute) thus ?thesis using \r \ 0\ \n \ 0\ by simp qed (* MOVE - EXTRA *) lemma prime_dvd_choose: assumes "0 < k" "k < p" "prime p" shows "p dvd (p choose k)" proof - have "k \ p" using \k < p\ by auto have "p dvd fact p" using \prime p\ by (simp add: prime_dvd_fact_iff) moreover have "\ p dvd fact k * fact (p - k)" unfolding prime_dvd_mult_iff[OF \prime p\] prime_dvd_fact_iff[OF \prime p\] using assms by simp ultimately show ?thesis unfolding binomial_fact_lemma[OF \k \ p\, symmetric] using assms prime_dvd_multD by blast qed lemma cong_eq_0_I: "(\i\A. [f i mod n = 0] (mod n)) \ [\i\A. f i = 0] (mod n)" using cong_sum by fastforce lemma power_mult_cong: assumes "[x^n = a](mod m)" "[y^n = b](mod m)" shows "[(x*y)^n = a*b](mod m)" using assms cong_mult[of "x^n" a m "y^n" b] power_mult_distrib by metis lemma fixes n :: nat assumes "n > 1" shows odd_pow_cong: "odd m \ [(n - 1) ^ m = n - 1] (mod n)" and even_pow_cong: "even m \ [(n - 1) ^ m = 1] (mod n)" proof (induction m) case (Suc m) case 1 with Suc have IH: "[(n - 1) ^ m = 1] (mod n)" by auto show ?case using \1 < n\ cong_mult[OF cong_refl IH] by simp next case (Suc m) case 2 with Suc have IH: "[(n - 1) ^ m = n - 1] (mod n)" by auto show ?case using cong_mult[OF cong_refl IH, of "(n - 1)"] and square_minus_one_cong_one[OF \1 < n\, of "n - 1"] by (auto simp: power2_eq_square intro: cong_trans) qed simp_all lemma cong_mult_uneq': fixes a :: "'a::{unique_euclidean_ring, ring_gcd}" assumes "coprime d a" shows "[b \ c] (mod a) \ [d = e] (mod a) \ [b * d \ c * e] (mod a)" using cong_mult_rcancel[OF assms] using cong_trans[of "b*d" "c*e" a "c*d"] using cong_scalar_left cong_sym by blast lemma p_coprime_right_nat: "prime p \ coprime a p = (\ p dvd a)" for p a :: nat by (meson coprime_absorb_left coprime_commute not_prime_unit prime_imp_coprime_nat) lemma squarefree_mult_imp_coprime: assumes "squarefree (a * b :: 'a :: semiring_gcd)" shows "coprime a b" proof (rule coprimeI) fix l assume "l dvd a" "l dvd b" then obtain a' b' where "a = l * a'" "b = l * b'" by (auto elim!: dvdE) with assms have "squarefree (l\<^sup>2 * (a' * b'))" by (simp add: power2_eq_square mult_ac) thus "l dvd 1" by (rule squarefreeD) auto qed lemma prime_divisor_exists_strong: fixes m :: int assumes "m > 1" "\prime m" shows "\n k. m = n * k \ 1 < n \ n < m \ 1 < k \ k < m" proof - from assms obtain n k where nk: "n * k > 1" "n \ 0" "m = n * k" "n \ 1" "n \ 0" "k \ 1" using assms unfolding prime_int_iff dvd_def by auto from nk have "n > 1" by linarith from nk assms have "n * k > 0" by simp with \n \ 0\ have "k > 0" using zero_less_mult_pos by force with \k \ 1\ have "k > 1" by linarith from nk have "n > 1" by linarith from \k > 1\ nk have "n < m" "k < m" by simp_all with nk \k > 1\ \n > 1\ show ?thesis by blast qed lemma prime_divisor_exists_strong_nat: fixes m :: nat assumes "1 < m" "\prime m" shows "\p k. m = p * k \ 1 < p \ p < m \ 1 < k \ k < m \ prime p" proof - obtain p where p_def: "prime p" "p dvd m" "p \ m" "1 < p" using assms prime_prime_factor and prime_gt_1_nat by blast moreover define k where "k = m div p" with \p dvd m\ have "m = p * k" by simp moreover have "p < m" using \p \ m\ dvd_imp_le[OF \p dvd m\] and \m > 1\ by simp moreover have "1 < k" "k < m" using \1 < m\ \1 < p\ and \p \ m\ unfolding \m = p * k\ by (force intro: Suc_lessI Nat.gr0I)+ ultimately show ?thesis using \1 < m\ by blast qed (* TODO Remove *) lemma prime_factorization_eqI: assumes "\p. p \# P \ prime p" "prod_mset P = n" shows "prime_factorization n = P" using prime_factorization_prod_mset_primes[of P] assms by simp lemma prime_factorization_prime_elem: assumes "prime_elem p" shows "prime_factorization p = {#normalize p#}" proof - have "prime_factorization p = prime_factorization (normalize p)" by (metis normalize_idem prime_factorization_cong) also have "\ = {#normalize p#}" by (rule prime_factorization_prime) (use assms in auto) finally show ?thesis . qed lemma size_prime_factorization_eq_Suc_0_iff [simp]: - "size (prime_factorization n) = Suc 0 \ prime_elem n" + fixes n :: "'a :: factorial_semiring_multiplicative" + shows "size (prime_factorization n) = Suc 0 \ prime_elem n" proof - define u where "u = unit_factor n" assume size: "size (prime_factorization n) = Suc 0" hence [simp]: "n \ 0" by auto from size obtain p where *: "prime_factorization n = {#p#}" by (auto elim!: size_mset_SucE) hence p: "p \ prime_factors n" by auto - have "n = u * prod_mset (prime_factorization n)" - unfolding u_def by (rule prime_decomposition [symmetric]) - with * have "n = u * p" by simp - also from p have "prime_elem \" - by (subst prime_elem_mult_unit_left) - (auto simp: u_def prime_imp_prime_elem in_prime_factors_iff) - finally show "prime_elem n" by auto + have "prime_elem (normalize p)" + using p by (auto simp: in_prime_factors_iff) + also have "p = prod_mset (prime_factorization n)" + using * by simp + also have "normalize \ = normalize n" + by (rule prod_mset_prime_factorization_weak) auto + finally show "prime_elem n" by simp qed (auto simp: prime_factorization_prime_elem) (* END TODO *) (* TODO Move *) lemma squarefree_prime_elem [simp, intro]: fixes p :: "'a :: algebraic_semidom" assumes "prime_elem p" shows "squarefree p" proof (rule squarefreeI) fix x assume "x\<^sup>2 dvd p" show "is_unit x" proof (rule ccontr) assume "\is_unit x" hence "\is_unit (x\<^sup>2)" by (simp add: is_unit_power_iff) from assms and this and \x\<^sup>2 dvd p\ have "prime_elem (x\<^sup>2)" by (rule prime_elem_mono) thus False by (simp add: prime_elem_power_iff) qed qed lemma squarefree_prime [simp, intro]: "prime p \ squarefree p" by auto lemma not_squarefree_primepow: assumes "primepow n" shows "squarefree n \ prime n" using assms by (auto simp: primepow_def squarefree_power_iff prime_power_iff) lemma prime_factorization_normalize [simp]: "prime_factorization (normalize n) = prime_factorization n" by (rule prime_factorization_cong) auto -lemma one_prime_factor_iff_primepow: "card (prime_factors n) = Suc 0 \ primepow (normalize n)" +lemma one_prime_factor_iff_primepow: + fixes n :: "'a :: factorial_semiring_multiplicative" + shows "card (prime_factors n) = Suc 0 \ primepow (normalize n)" proof assume "primepow (normalize n)" then obtain p k where pk: "prime p" "normalize n = p ^ k" "k > 0" by (auto simp: primepow_def) hence "card (prime_factors (normalize n)) = Suc 0" by (subst pk) (simp add: prime_factors_power prime_factorization_prime) thus "card (prime_factors n) = Suc 0" by simp next assume *: "card (prime_factors n) = Suc 0" from * have "(\p\prime_factors n. p ^ multiplicity p n) = normalize n" by (intro prod_prime_factors) auto also from * have "card (prime_factors n) = 1" by simp then obtain p where p: "prime_factors n = {p}" by (elim card_1_singletonE) finally have "normalize n = p ^ multiplicity p n" by simp moreover from p have "prime p" "multiplicity p n > 0" by (auto simp: prime_factors_multiplicity) ultimately show "primepow (normalize n)" unfolding primepow_def by blast qed lemma squarefree_imp_prod_prime_factors_eq: + fixes x :: "'a :: factorial_semiring_multiplicative" assumes "squarefree x" shows "\(prime_factors x) = normalize x" proof - from assms have [simp]: "x \ 0" by auto have "(\p\prime_factors x. p ^ multiplicity p x) = normalize x" by (intro prod_prime_factors) auto also have "(\p\prime_factors x. p ^ multiplicity p x) = (\p\prime_factors x. p)" using assms by (intro prod.cong refl) (auto simp: squarefree_factorial_semiring') finally show ?thesis by simp qed (* END TODO *) end \ No newline at end of file diff --git a/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy b/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy --- a/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy +++ b/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy @@ -1,961 +1,961 @@ (* Author: Manuel Eberl *) theory Misc_Polynomial imports "HOL-Computational_Algebra.Polynomial" "HOL-Computational_Algebra.Polynomial_Factorial" begin subsection \Analysis\ lemma fun_eq_in_ivl: assumes "a \ b" "\x::real. a \ x \ x \ b \ eventually (\\. f \ = f x) (at x)" shows "f a = f b" proof (rule connected_local_const) show "connected {a..b}" "a \ {a..b}" "b \ {a..b}" using \a \ b\ by (auto intro: connected_Icc) show "\aa\{a..b}. eventually (\b. f aa = f b) (at aa within {a..b})" proof fix x assume "x \ {a..b}" with assms(2)[rule_format, of x] show "eventually (\b. f x = f b) (at x within {a..b})" by (auto simp: eventually_at_filter elim: eventually_mono) qed qed subsection \Polynomials\ subsubsection \General simplification lemmas\ lemma pderiv_div: assumes [simp]: "q dvd p" "q \ 0" shows "pderiv (p div q) = (q * pderiv p - p * pderiv q) div (q * q)" "q*q dvd (q * pderiv p - p * pderiv q)" proof- from assms obtain r where "p = q * r" unfolding dvd_def by blast hence "q * pderiv p - p * pderiv q = (q * q) * pderiv r" by (simp add: algebra_simps pderiv_mult) thus "q*q dvd (q * pderiv p - p * pderiv q)" by simp note 0 = pderiv_mult[of q "p div q"] have 1: "q * (p div q) = p" by (metis assms(1) assms(2) dvd_def nonzero_mult_div_cancel_left) have f1: "pderiv (p div q) * (q * q) div (q * q) = pderiv (p div q)" by simp have f2: "pderiv p = q * pderiv (p div q) + p div q * pderiv q" by (metis 0 1) have "p * pderiv q = pderiv q * (q * (p div q))" by (metis 1 mult.commute) then have "p * pderiv q = q * (p div q * pderiv q)" by fastforce then have "q * pderiv p - p * pderiv q = q * (q * pderiv (p div q))" using f2 by (metis add_diff_cancel_right' distrib_left) then show "pderiv (p div q) = (q * pderiv p - p * pderiv q) div (q * q)" using f1 by (metis mult.commute mult.left_commute) qed subsubsection \Divisibility of polynomials\ text \ Two polynomials that are coprime have no common roots. \ lemma coprime_imp_no_common_roots: "\ (poly p x = 0 \ poly q x = 0)" if "coprime p q" for x :: "'a :: field" proof clarify assume "poly p x = 0" "poly q x = 0" then have "[:-x, 1:] dvd p" "[:-x, 1:] dvd q" by (simp_all add: poly_eq_0_iff_dvd) with that have "is_unit [:-x, 1:]" by (rule coprime_common_divisor) then show False by (auto simp add: is_unit_pCons_iff) qed lemma poly_div: assumes "poly q x \ 0" and "(q::'a :: field poly) dvd p" shows "poly (p div q) x = poly p x / poly q x" proof- from assms have [simp]: "q \ 0" by force have "poly q x * poly (p div q) x = poly (q * (p div q)) x" by simp also have "q * (p div q) = p" using assms by (simp add: div_mult_swap) finally show "poly (p div q) x = poly p x / poly q x" using assms by (simp add: field_simps) qed (* TODO: make this less ugly *) lemma poly_div_gcd_squarefree_aux: - assumes "pderiv (p::('a::{field_char_0,euclidean_ring_gcd}) poly) \ 0" + assumes "pderiv (p::('a::{field_char_0,field_gcd}) poly) \ 0" defines "d \ gcd p (pderiv p)" shows "coprime (p div d) (pderiv (p div d))" and "\x. poly (p div d) x = 0 \ poly p x = 0" proof - obtain r s where "bezout_coefficients p (pderiv p) = (r, s)" by (auto simp add: prod_eq_iff) then have "r * p + s * pderiv p = gcd p (pderiv p)" by (rule bezout_coefficients) then have rs: "d = r * p + s * pderiv p" by (simp add: d_def) define t where "t = p div d" define p' where [simp]: "p' = pderiv p" define d' where [simp]: "d' = pderiv d" define u where "u = p' div d" have A: "p = t * d" and B: "p' = u * d" by (simp_all add: t_def u_def d_def algebra_simps) from poly_squarefree_decomp[OF assms(1) A B[unfolded p'_def] rs] show "\x. poly (p div d) x = 0 \ poly p x = 0" by (auto simp: t_def) from rs have C: "s*t*d' = d * (1 - r*t - s*pderiv t)" by (simp add: A B algebra_simps pderiv_mult) from assms have [simp]: "p \ 0" "d \ 0" "t \ 0" by (force, force, subst (asm) A, force) have "\x. \x dvd t; x dvd (pderiv t)\ \ x dvd 1" proof - fix x assume "x dvd t" "x dvd (pderiv t)" then obtain v w where vw: "t = x*v" "pderiv t = x*w" unfolding dvd_def by blast define x' v' where [simp]: "x' = pderiv x" and [simp]: "v' = pderiv v" from vw have "x*v' + v*x' = x*w" by (simp add: pderiv_mult) hence "v*x' = x*(w - v')" by (simp add: algebra_simps) hence "x dvd v*pderiv x" by simp then obtain y where y: "v*x' = x*y" unfolding dvd_def by force from \t \ 0\ and vw have "x \ 0" by simp have x_pow_n_dvd_d: "\n. x^n dvd d" proof- fix n show "x ^ n dvd d" proof (induction n, simp, rename_tac n, case_tac n) fix n assume "n = (0::nat)" from vw and C have "d = x*(d*r*v + d*s*w + s*v*d')" by (simp add: algebra_simps) with \n = 0\ show "x^Suc n dvd d" by (force intro: dvdI) next fix n n' assume IH: "x^n dvd d" and "n = Suc n'" hence [simp]: "Suc n' = n" "x * x^n' = x^n" by simp_all define c :: "'a poly" where "c = [:of_nat n:]" from pderiv_power_Suc[of x n'] have [simp]: "pderiv (x^n) = c*x^n' * x'" unfolding c_def by simp from IH obtain z where d: "d = x^n * z" unfolding dvd_def by blast define z' where [simp]: "z' = pderiv z" from d \d \ 0\ have "x^n \ 0" "z \ 0" by force+ from C d have "x^n*z = z*r*v*x^Suc n + z*s*c*x^n*(v*x') + s*v*z'*x^Suc n + s*z*(v*x')*x^n + s*z*v'*x^Suc n" by (simp add: algebra_simps vw pderiv_mult) also have "... = x^n*x * (z*r*v + z*s*c*y + s*v*z' + s*z*y + s*z*v')" by (simp only: y, simp add: algebra_simps) finally have "z = x*(z*r*v+z*s*c*y+s*v*z'+s*z*y+s*z*v')" using \x^n \ 0\ by force hence "x dvd z" by (metis dvd_triv_left) with d show "x^Suc n dvd d" by simp qed qed have "degree x = 0" proof (cases "degree x", simp) case (Suc n) hence "x \ 0" by auto with Suc have "degree (x ^ (Suc (degree d))) > degree d" by (subst degree_power_eq, simp_all) moreover from x_pow_n_dvd_d[of "Suc (degree d)"] and \d \ 0\ have "degree (x^Suc (degree d)) \ degree d" by (simp add: dvd_imp_degree_le) ultimately show ?thesis by simp qed then obtain c where [simp]: "x = [:c:]" by (cases x, simp split: if_split_asm) moreover from \x \ 0\ have "c \ 0" by simp ultimately show "x dvd 1" using dvdI[of 1 x "[:inverse c:]"] by simp qed then show "coprime t (pderiv t)" by (rule coprimeI) qed lemma normalize_field: "normalize (x :: 'a :: {field,normalization_semidom}) = (if x = 0 then 0 else 1)" by (auto simp: is_unit_normalize dvd_field_iff) lemma normalize_field_eq_1 [simp]: "x \ 0 \ normalize (x :: 'a :: {field,normalization_semidom}) = 1" by (simp add: normalize_field) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {field,normalization_semidom}) = x" by (cases "x = 0") (auto simp: is_unit_unit_factor dvd_field_iff) text \ Dividing a polynomial by its gcd with its derivative yields a squarefree polynomial with the same roots. \ lemma poly_div_gcd_squarefree: - assumes "(p :: ('a::{field_char_0,euclidean_ring_gcd}) poly) \ 0" + assumes "(p :: ('a::{field_char_0,field_gcd}) poly) \ 0" defines "d \ gcd p (pderiv p)" shows "coprime (p div d) (pderiv (p div d))" (is ?A) and "\x. poly (p div d) x = 0 \ poly p x = 0" (is "\x. ?B x") proof- have "?A \ (\x. ?B x)" proof (cases "pderiv p = 0") case False from poly_div_gcd_squarefree_aux[OF this] show ?thesis unfolding d_def by auto next case True then obtain c where [simp]: "p = [:c:]" using pderiv_iszero by blast from assms(1) have "c \ 0" by simp from True have "d = smult (inverse c) p" by (simp add: d_def normalize_poly_def map_poly_pCons field_simps) with \p \ 0\ \c \ 0\ have "p div d = [:c:]" by (simp add: pCons_one) with \c \ 0\ show ?thesis by (simp add: normalize_const_poly is_unit_triv) qed thus ?A and "\x. ?B x" by simp_all qed subsubsection \Sign changes of a polynomial\ text \ If a polynomial has different signs at two points, it has a root inbetween. \ lemma poly_different_sign_imp_root: assumes "a < b" and "sgn (poly p a) \ sgn (poly p (b::real))" shows "\x. a \ x \ x \ b \ poly p x = 0" proof (cases "poly p a = 0 \ poly p b = 0") case True thus ?thesis using assms(1) by (elim disjE, rule_tac exI[of _ a], simp, rule_tac exI[of _ b], simp) next case False hence [simp]: "poly p a \ 0" "poly p b \ 0" by simp_all show ?thesis proof (cases "poly p a < 0") case True hence "sgn (poly p a) = -1" by simp with assms True have "poly p b > 0" by (auto simp: sgn_real_def split: if_split_asm) from poly_IVT_pos[OF \a < b\ True this] guess x .. thus ?thesis by (intro exI[of _ x], simp) next case False hence "poly p a > 0" by (simp add: not_less less_eq_real_def) hence "sgn (poly p a) = 1" by simp with assms False have "poly p b < 0" by (auto simp: sgn_real_def not_less less_eq_real_def split: if_split_asm) from poly_IVT_neg[OF \a < b\ \poly p a > 0\ this] guess x .. thus ?thesis by (intro exI[of _ x], simp) qed qed lemma poly_different_sign_imp_root': assumes "sgn (poly p a) \ sgn (poly p (b::real))" shows "\x. poly p x = 0" using assms by (cases "a < b", auto dest!: poly_different_sign_imp_root simp: less_eq_real_def not_less) lemma no_roots_inbetween_imp_same_sign: assumes "a < b" "\x. a \ x \ x \ b \ poly p x \ (0::real)" shows "sgn (poly p a) = sgn (poly p b)" using poly_different_sign_imp_root assms by auto subsubsection \Limits of polynomials\ lemma poly_neighbourhood_without_roots: assumes "(p :: real poly) \ 0" shows "eventually (\x. poly p x \ 0) (at x\<^sub>0)" proof- { fix \ :: real assume "\ > 0" have fin: "finite {x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" using poly_roots_finite[OF assms] by simp with \\ > 0\have "\\>0. \\\ \ (\x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x \ 0)" proof (induction "card {x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" arbitrary: \ rule: less_induct) case (less \) let ?A = "{x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" show ?case proof (cases "card ?A") case 0 hence "?A = {}" using less by auto thus ?thesis using less(2) by (rule_tac exI[of _ \], auto) next case (Suc _) with less(3) have "{x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0} \ {}" by force then obtain x where x_props: "\x - x\<^sub>0\ < \" "x \ x\<^sub>0" "poly p x = 0" by blast define \' where "\' = \x - x\<^sub>0\ / 2" have "\' > 0" "\' < \" unfolding \'_def using x_props by simp_all from x_props(1,2) and \\ > 0\ have "x \ {x'. \x' - x\<^sub>0\ < \' \ x' \ x\<^sub>0 \ poly p x' = 0}" (is "_ \ ?B") by (auto simp: \'_def) moreover from x_props have "x \ {x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" by blast ultimately have "?B \ ?A" by auto hence "card ?B < card ?A" "finite ?B" by (rule psubset_card_mono[OF less(3)], blast intro: finite_subset[OF _ less(3)]) from less(1)[OF this(1) \\' > 0\ this(2)] show ?thesis using \\' < \\ by force qed qed } from this[of "1"] show ?thesis by (auto simp: eventually_at dist_real_def) qed lemma poly_neighbourhood_same_sign: assumes "poly p (x\<^sub>0 :: real) \ 0" shows "eventually (\x. sgn (poly p x) = sgn (poly p x\<^sub>0)) (at x\<^sub>0)" proof - have cont: "isCont (\x. sgn (poly p x)) x\<^sub>0" by (rule isCont_sgn, rule poly_isCont, rule assms) then have "eventually (\x. \sgn (poly p x) - sgn (poly p x\<^sub>0)\ < 1) (at x\<^sub>0)" by (auto simp: isCont_def tendsto_iff dist_real_def) then show ?thesis by (rule eventually_mono) (simp add: sgn_real_def split: if_split_asm) qed lemma poly_lhopital: assumes "poly p (x::real) = 0" "poly q x = 0" "q \ 0" assumes "(\x. poly (pderiv p) x / poly (pderiv q) x) \x\ y" shows "(\x. poly p x / poly q x) \x\ y" using assms proof (rule_tac lhopital) have "isCont (poly p) x" "isCont (poly q) x" by simp_all with assms(1,2) show "poly p \x\ 0" "poly q \x\ 0" by (simp_all add: isCont_def) from \q \ 0\ and \poly q x = 0\ have "pderiv q \ 0" by (auto dest: pderiv_iszero) from poly_neighbourhood_without_roots[OF this] show "eventually (\x. poly (pderiv q) x \ 0) (at x)" . qed (auto intro: poly_DERIV poly_neighbourhood_without_roots) lemma poly_roots_bounds: assumes "p \ 0" obtains l u where "l \ (u :: real)" and "poly p l \ 0" and "poly p u \ 0" and "{x. x > l \ x \ u \ poly p x = 0 } = {x. poly p x = 0}" and "\x. x \ l \ sgn (poly p x) = sgn (poly p l)" and "\x. x \ u \ sgn (poly p x) = sgn (poly p u)" proof from assms have "finite {x. poly p x = 0}" (is "finite ?roots") using poly_roots_finite by fast let ?roots' = "insert 0 ?roots" define l where "l = Min ?roots' - 1" define u where "u = Max ?roots' + 1" from \finite ?roots\ have A: "finite ?roots'" by auto from Min_le[OF this, of 0] and Max_ge[OF this, of 0] show "l \ u" by (simp add: l_def u_def) from Min_le[OF A] have l_props: "\x. x\l \ poly p x \ 0" by (fastforce simp: l_def) from Max_ge[OF A] have u_props: "\x. x\u \ poly p x \ 0" by (fastforce simp: u_def) from l_props u_props show [simp]: "poly p l \ 0" "poly p u \ 0" by auto from l_props have "\x. poly p x = 0 \ x > l" by (metis not_le) moreover from u_props have "\x. poly p x = 0 \ x \ u" by (metis linear) ultimately show "{x. x > l \ x \ u \ poly p x = 0} = ?roots" by auto { fix x assume A: "x < l" "sgn (poly p x) \ sgn (poly p l)" with poly_IVT_pos[OF A(1), of p] poly_IVT_neg[OF A(1), of p] A(2) have False by (auto split: if_split_asm simp: sgn_real_def l_props not_less less_eq_real_def) } thus "\x. x \ l \ sgn (poly p x) = sgn (poly p l)" by (case_tac "x = l", auto simp: less_eq_real_def) { fix x assume A: "x > u" "sgn (poly p x) \ sgn (poly p u)" with u_props poly_IVT_neg[OF A(1), of p] poly_IVT_pos[OF A(1), of p] A(2) have False by (auto split: if_split_asm simp: sgn_real_def l_props not_less less_eq_real_def) } thus "\x. x \ u \ sgn (poly p x) = sgn (poly p u)" by (case_tac "x = u", auto simp: less_eq_real_def) qed definition poly_inf :: "('a::real_normed_vector) poly \ 'a" where "poly_inf p \ sgn (coeff p (degree p))" definition poly_neg_inf :: "('a::real_normed_vector) poly \ 'a" where "poly_neg_inf p \ if even (degree p) then sgn (coeff p (degree p)) else -sgn (coeff p (degree p))" lemma poly_inf_0_iff[simp]: "poly_inf p = 0 \ p = 0" "poly_neg_inf p = 0 \ p = 0" by (auto simp: poly_inf_def poly_neg_inf_def sgn_zero_iff) lemma poly_inf_mult[simp]: fixes p :: "('a::real_normed_field) poly" shows "poly_inf (p*q) = poly_inf p * poly_inf q" "poly_neg_inf (p*q) = poly_neg_inf p * poly_neg_inf q" unfolding poly_inf_def poly_neg_inf_def by ((cases "p = 0 \ q = 0",auto simp: sgn_zero_iff degree_mult_eq[of p q] coeff_mult_degree_sum Real_Vector_Spaces.sgn_mult)[])+ lemma poly_neq_0_at_infinity: assumes "(p :: real poly) \ 0" shows "eventually (\x. poly p x \ 0) at_infinity" proof- from poly_roots_bounds[OF assms] guess l u . note lu_props = this define b where "b = max (-l) u" show ?thesis proof (subst eventually_at_infinity, rule exI[of _ b], clarsimp) fix x assume A: "\x\ \ b" and B: "poly p x = 0" show False proof (cases "x \ 0") case True with A have "x \ u" unfolding b_def by simp with lu_props(3, 6) show False by (metis sgn_zero_iff B) next case False with A have "x \ l" unfolding b_def by simp with lu_props(2, 5) show False by (metis sgn_zero_iff B) qed qed qed lemma poly_limit_aux: fixes p :: "real poly" defines "n \ degree p" shows "((\x. poly p x / x ^ n) \ coeff p n) at_infinity" proof (subst filterlim_cong, rule refl, rule refl) show "eventually (\x. poly p x / x^n = (\i\n. coeff p i / x^(n-i))) at_infinity" proof (rule eventually_mono) show "eventually (\x::real. x \ 0) at_infinity" by (simp add: eventually_at_infinity, rule exI[of _ 1], auto) fix x :: real assume [simp]: "x \ 0" show "poly p x / x ^ n = (\i\n. coeff p i / x ^ (n - i))" by (simp add: n_def sum_divide_distrib power_diff poly_altdef) qed let ?a = "\i. if i = n then coeff p n else 0" have "\i\{..n}. ((\x. coeff p i / x ^ (n - i)) \ ?a i) at_infinity" proof fix i assume "i \ {..n}" hence "i \ n" by simp show "((\x. coeff p i / x ^ (n - i)) \ ?a i) at_infinity" proof (cases "i = n") case True thus ?thesis by (intro tendstoI, subst eventually_at_infinity, intro exI[of _ 1], simp add: dist_real_def) next case False hence "n - i > 0" using \i \ n\ by simp from tendsto_inverse_0 and divide_real_def[of 1] have "((\x. 1 / x :: real) \ 0) at_infinity" by simp from tendsto_power[OF this, of "n - i"] have "((\x::real. 1 / x ^ (n - i)) \ 0) at_infinity" using \n - i > 0\ by (simp add: power_0_left power_one_over) from tendsto_mult_right_zero[OF this, of "coeff p i"] have "((\x. coeff p i / x ^ (n - i)) \ 0) at_infinity" by (simp add: field_simps) thus ?thesis using False by simp qed qed hence "((\x. \i\n. coeff p i / x^(n-i)) \ (\i\n. ?a i)) at_infinity" by (force intro!: tendsto_sum) also have "(\i\n. ?a i) = coeff p n" by (subst sum.delta, simp_all) finally show "((\x. \i\n. coeff p i / x^(n-i)) \ coeff p n) at_infinity" . qed lemma poly_at_top_at_top: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) > 0" shows "LIM x at_top. poly p x :> at_top" proof- let ?n = "degree p" define f g where "f x = poly p x / x^?n" and "g x = x ^ ?n" for x :: real from poly_limit_aux have "(f \ coeff p (degree p)) at_top" using tendsto_mono at_top_le_at_infinity unfolding f_def by blast moreover from assms have "LIM x at_top. g x :> at_top" by (auto simp add: g_def intro!: filterlim_pow_at_top filterlim_ident) ultimately have "LIM x at_top. f x * g x :> at_top" using filterlim_tendsto_pos_mult_at_top assms by simp also have "eventually (\x. f x * g x = poly p x) at_top" unfolding f_def g_def by (subst eventually_at_top_linorder, rule exI[of _ 1], simp add: poly_altdef field_simps sum_distrib_left power_diff) note filterlim_cong[OF refl refl this] finally show ?thesis . qed lemma poly_at_bot_at_top: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) < 0" shows "LIM x at_top. poly p x :> at_bot" proof- from poly_at_top_at_top[of "-p"] and assms have "LIM x at_top. -poly p x :> at_top" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma poly_lim_inf: "eventually (\x::real. sgn (poly p x) = poly_inf p) at_top" proof (cases "degree p \ 1") case False hence "degree p = 0" by simp then obtain c where "p = [:c:]" by (cases p, auto split: if_split_asm) thus ?thesis by (simp add: eventually_at_top_linorder poly_inf_def) next case True note deg = this let ?lc = "coeff p (degree p)" from True have "?lc \ 0" by force show ?thesis proof (cases "?lc > 0") case True from poly_at_top_at_top[OF deg this] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp: filterlim_at_top eventually_at_top_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp only: eventually_at_top_linorder poly_inf_def, intro exI[of _ x\<^sub>0], simp add: True) next case False hence "?lc < 0" using \?lc \ 0\ by linarith from poly_at_bot_at_top[OF deg this] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp: filterlim_at_bot eventually_at_top_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp only: eventually_at_top_linorder poly_inf_def, intro exI[of _ x\<^sub>0], simp add: \?lc < 0\) qed qed lemma poly_at_top_or_bot_at_bot: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) > 0" shows "LIM x at_bot. poly p x :> (if even (degree p) then at_top else at_bot)" proof- let ?n = "degree p" define f g where "f x = poly p x / x ^ ?n" and "g x = x ^ ?n" for x :: real from poly_limit_aux have "(f \ coeff p (degree p)) at_bot" using tendsto_mono at_bot_le_at_infinity by (force simp: f_def [abs_def]) moreover from assms have "LIM x at_bot. g x :> (if even (degree p) then at_top else at_bot)" by (auto simp add: g_def split: if_split_asm intro: filterlim_pow_at_bot_even filterlim_pow_at_bot_odd filterlim_ident) ultimately have "LIM x at_bot. f x * g x :> (if even ?n then at_top else at_bot)" by (auto simp: assms intro: filterlim_tendsto_pos_mult_at_top filterlim_tendsto_pos_mult_at_bot) also have "eventually (\x. f x * g x = poly p x) at_bot" unfolding f_def g_def by (subst eventually_at_bot_linorder, rule exI[of _ "-1"], simp add: poly_altdef field_simps sum_distrib_left power_diff) note filterlim_cong[OF refl refl this] finally show ?thesis . qed lemma poly_at_bot_or_top_at_bot: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) < 0" shows "LIM x at_bot. poly p x :> (if even (degree p) then at_bot else at_top)" proof- from poly_at_top_or_bot_at_bot[of "-p"] and assms have "LIM x at_bot. -poly p x :> (if even (degree p) then at_top else at_bot)" by simp thus ?thesis by (auto simp: filterlim_uminus_at_bot) qed lemma poly_lim_neg_inf: "eventually (\x::real. sgn (poly p x) = poly_neg_inf p) at_bot" proof (cases "degree p \ 1") case False hence "degree p = 0" by simp then obtain c where "p = [:c:]" by (cases p, auto split: if_split_asm) thus ?thesis by (simp add: eventually_at_bot_linorder poly_neg_inf_def) next case True note deg = this let ?lc = "coeff p (degree p)" from True have "?lc \ 0" by force show ?thesis proof (cases "?lc > 0") case True note lc_pos = this show ?thesis proof (cases "even (degree p)") case True from poly_at_top_or_bot_at_bot[OF deg lc_pos] and True obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp add: filterlim_at_top filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp add: True eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_pos) next case False from poly_at_top_or_bot_at_bot[OF deg lc_pos] and False obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp add: filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp add: False eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_pos) qed next case False hence lc_neg: "?lc < 0" using \?lc \ 0\ by linarith show ?thesis proof (cases "even (degree p)") case True with poly_at_bot_or_top_at_bot[OF deg lc_neg] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp: filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp only: True eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_neg) next case False with poly_at_bot_or_top_at_bot[OF deg lc_neg] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp: filterlim_at_top eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp only: False eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_neg) qed qed qed subsubsection \Signs of polynomials for sufficiently large values\ lemma polys_inf_sign_thresholds: assumes "finite (ps :: real poly set)" obtains l u where "l \ u" and "\p. \p \ ps; p \ 0\ \ {x. l < x \ x \ u \ poly p x = 0} = {x. poly p x = 0}" and "\p x. \p \ ps; x \ u\ \ sgn (poly p x) = poly_inf p" and "\p x. \p \ ps; x \ l\ \ sgn (poly p x) = poly_neg_inf p" proof goal_cases case prems: 1 have "\l u. l \ u \ (\p x. p \ ps \ x \ u \ sgn (poly p x) = poly_inf p) \ (\p x. p \ ps \ x \ l \ sgn (poly p x) = poly_neg_inf p)" (is "\l u. ?P ps l u") proof (induction rule: finite_subset_induct[OF assms(1), where A = UNIV]) case 1 show ?case by simp next case 2 show ?case by (intro exI[of _ 42], simp) next case prems: (3 p ps) from prems(4) obtain l u where lu_props: "?P ps l u" by blast from poly_lim_inf obtain u' where u'_props: "\x\u'. sgn (poly p x) = poly_inf p" by (force simp add: eventually_at_top_linorder) from poly_lim_neg_inf obtain l' where l'_props: "\x\l'. sgn (poly p x) = poly_neg_inf p" by (force simp add: eventually_at_bot_linorder) show ?case by (rule exI[of _ "min l l'"], rule exI[of _ "max u u'"], insert lu_props l'_props u'_props, auto) qed then obtain l u where lu_props: "l \ u" "\p x. p \ ps \ u \ x \ sgn (poly p x) = poly_inf p" "\p x. p \ ps \ x \ l \ sgn (poly p x) = poly_neg_inf p" by blast moreover { fix p x assume A: "p \ ps" "p \ 0" "poly p x = 0" from A have "l < x" "x < u" by (auto simp: not_le[symmetric] dest: lu_props(2,3)) } note A = this have "\p. p \ ps \ p \ 0 \ {x. l < x \ x \ u \ poly p x = 0} = {x. poly p x = 0}" by (auto dest: A) from prems[OF lu_props(1) this lu_props(2,3)] show thesis . qed subsubsection \Positivity of polynomials\ lemma poly_pos: "(\x::real. poly p x > 0) \ poly_inf p = 1 \ (\x. poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. poly p x \ 0" by simp from poly_lim_inf obtain x where "sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) with A show "poly_inf p = 1" by (simp add: sgn_real_def split: if_split_asm) next assume "poly_inf p = 1 \ (\x. poly p x \ 0)" hence A: "poly_inf p = 1" and B: "(\x. poly p x \ 0)" by simp_all from poly_lim_inf obtain x where C: "sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) show "\x. poly p x > 0" proof (rule ccontr) assume "\(\x. poly p x > 0)" then obtain x' where "poly p x' \ 0" by (auto simp: not_less) with A and C have "sgn (poly p x') \ sgn (poly p x)" by (auto simp: sgn_real_def split: if_split_asm) from poly_different_sign_imp_root'[OF this] and B show False by blast qed qed lemma poly_pos_greater: "(\x::real. x > a \ poly p x > 0) \ poly_inf p = 1 \ (\x. x > a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x > a \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. x > a \ poly p x \ 0" by auto from poly_lim_inf obtain x\<^sub>0 where "\x\x\<^sub>0. sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) hence "poly_inf p = sgn (poly p (max x\<^sub>0 (a + 1)))" by simp also from A have "... = 1" by force finally show "poly_inf p = 1" . next assume "poly_inf p = 1 \ (\x. x > a \ poly p x \ 0)" hence A: "poly_inf p = 1" and B: "(\x. x > a \ poly p x \ 0)" by simp_all from poly_lim_inf obtain x\<^sub>0 where C: "\x\x\<^sub>0. sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) hence "sgn (poly p (max x\<^sub>0 (a+1))) = poly_inf p" by simp with A have D: "sgn (poly p (max x\<^sub>0 (a+1))) = 1" by simp show "\x. x > a \ poly p x > 0" proof (rule ccontr) assume "\(\x. x > a \ poly p x > 0)" then obtain x' where "x' > a" "poly p x' \ 0" by (auto simp: not_less) with A and D have E: "sgn (poly p x') \ sgn (poly p (max x\<^sub>0(a+1)))" by (auto simp: sgn_real_def split: if_split_asm) show False apply (cases x' "max x\<^sub>0 (a+1)" rule: linorder_cases) using B E \x' > a\ apply (force dest!: poly_different_sign_imp_root[of _ _ p])+ done qed qed lemma poly_pos_geq: "(\x::real. x \ a \ poly p x > 0) \ poly_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x \ a \ poly p x > 0" hence "\x::real. x > a \ poly p x > 0" by simp also note poly_pos_greater finally have "poly_inf p = 1" "(\x>a. poly p x \ 0)" by simp_all moreover from A have "poly p a > 0" by simp ultimately show "poly_inf p = 1" "\x\a. poly p x \ 0" by (auto simp: less_eq_real_def) next assume "poly_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" hence A: "poly_inf p = 1" and B: "poly p a \ 0" and C: "\x>a. poly p x \ 0" by simp_all from A and C and poly_pos_greater have "\x>a. poly p x > 0" by simp moreover with B C poly_IVT_pos[of a "a+1" p] have "poly p a > 0" by force ultimately show "\x\a. poly p x > 0" by (auto simp: less_eq_real_def) qed lemma poly_pos_less: "(\x::real. x < a \ poly p x > 0) \ poly_neg_inf p = 1 \ (\x. x < a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x < a \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. x < a \ poly p x \ 0" by auto from poly_lim_neg_inf obtain x\<^sub>0 where "\x\x\<^sub>0. sgn (poly p x) = poly_neg_inf p" by (auto simp: eventually_at_bot_linorder) hence "poly_neg_inf p = sgn (poly p (min x\<^sub>0 (a - 1)))" by simp also from A have "... = 1" by force finally show "poly_neg_inf p = 1" . next assume "poly_neg_inf p = 1 \ (\x. x < a \ poly p x \ 0)" hence A: "poly_neg_inf p = 1" and B: "(\x. x < a \ poly p x \ 0)" by simp_all from poly_lim_neg_inf obtain x\<^sub>0 where C: "\x\x\<^sub>0. sgn (poly p x) = poly_neg_inf p" by (auto simp: eventually_at_bot_linorder) hence "sgn (poly p (min x\<^sub>0 (a - 1))) = poly_neg_inf p" by simp with A have D: "sgn (poly p (min x\<^sub>0 (a - 1))) = 1" by simp show "\x. x < a \ poly p x > 0" proof (rule ccontr) assume "\(\x. x < a \ poly p x > 0)" then obtain x' where "x' < a" "poly p x' \ 0" by (auto simp: not_less) with A and D have E: "sgn (poly p x') \ sgn (poly p (min x\<^sub>0 (a - 1)))" by (auto simp: sgn_real_def split: if_split_asm) show False apply (cases x' "min x\<^sub>0 (a - 1)" rule: linorder_cases) using B E \x' < a\ apply (auto dest!: poly_different_sign_imp_root[of _ _ p])+ done qed qed lemma poly_pos_leq: "(\x::real. x \ a \ poly p x > 0) \ poly_neg_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x \ a \ poly p x > 0" hence "\x::real. x < a \ poly p x > 0" by simp also note poly_pos_less finally have "poly_neg_inf p = 1" "(\x 0)" by simp_all moreover from A have "poly p a > 0" by simp ultimately show "poly_neg_inf p = 1" "\x\a. poly p x \ 0" by (auto simp: less_eq_real_def) next assume "poly_neg_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" hence A: "poly_neg_inf p = 1" and B: "poly p a \ 0" and C: "\x 0" by simp_all from A and C and poly_pos_less have "\x 0" by simp moreover with B C poly_IVT_neg[of "a - 1" a p] have "poly p a > 0" by force ultimately show "\x\a. poly p x > 0" by (auto simp: less_eq_real_def) qed lemma poly_pos_between_less_less: "(\x::real. a < x \ x < b \ poly p x > 0) \ (a \ b \ poly p ((a+b)/2) > 0) \ (\x. a < x \ x < b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a < x \ x < b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a < x \ x < b \ poly p x \ 0" by auto from A show "a \ b \ poly p ((a+b)/2) > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p ((a+b)/2)) \ (\x. a x poly p x \ 0)" hence A: "b \ a \ 0 < poly p ((a+b)/2)" and B: "\x. a x poly p x \ 0" by simp_all show "\x. a < x \ x < b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a < x" "x < b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p ((a+b)/2) > 0" by simp ultimately have "sgn (poly p x) \ sgn (poly p ((a+b)/2))" by simp thus False using B apply (cases x "(a+b)/2" rule: linorder_cases) apply (drule poly_different_sign_imp_root[of _ _ p], assumption, insert \a < b\ \a < x\ \x < b\ , force) [] apply simp apply (drule poly_different_sign_imp_root[of _ _ p], simp, insert \a < b\ \a < x\ \x < b\ , force) done qed qed lemma poly_pos_between_less_leq: "(\x::real. a < x \ x \ b \ poly p x > 0) \ (a \ b \ poly p b > 0) \ (\x. a < x \ x \ b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a < x \ x \ b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a < x \ x \ b \ poly p x \ 0" by auto from A show "a \ b \ poly p b > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p b) \ (\x. a x\b \ poly p x \ 0)" hence A: "b \ a \ 0 < poly p b" and B: "\x. a x\b \ poly p x \ 0" by simp_all show "\x. a < x \ x \ b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a < x" "x \ b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p b > 0" by simp ultimately have "x < b" using \x \ b\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p b > 0\ have "sgn (poly p x) \ sgn (poly p b)" by simp from poly_different_sign_imp_root[OF \x < b\ this] and B and \x > a\ show False by auto qed qed lemma poly_pos_between_leq_less: "(\x::real. a \ x \ x < b \ poly p x > 0) \ (a \ b \ poly p a > 0) \ (\x. a \ x \ x < b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a \ x \ x < b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a \ x \ x < b \ poly p x \ 0" by auto from A show "a \ b \ poly p a > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p a) \ (\x. a\x \ x poly p x \ 0)" hence A: "b \ a \ 0 < poly p a" and B: "\x. a\x \ x poly p x \ 0" by simp_all show "\x. a \ x \ x < b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a \ x" "x < b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p a > 0" by simp ultimately have "x > a" using \x \ a\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p a > 0\ have "sgn (poly p a) \ sgn (poly p x)" by simp from poly_different_sign_imp_root[OF \x > a\ this] and B and \x < b\ show False by auto qed qed lemma poly_pos_between_leq_leq: "(\x::real. a \ x \ x \ b \ poly p x > 0) \ (a > b \ poly p a > 0) \ (\x. a \ x \ x \ b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a \ x \ x \ b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a \ x \ x \ b \ poly p x \ 0" by auto from A show "a > b \ poly p a > 0" by (cases "a \ b", auto) next assume "(b < a \ 0 < poly p a) \ (\x. a\x \ x\b \ poly p x \ 0)" hence A: "b < a \ 0 < poly p a" and B: "\x. a\x \ x\b \ poly p x \ 0" by simp_all show "\x. a \ x \ x \ b \ poly p x > 0" proof (cases "a > b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a \ b" "a \ x" "x \ b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a \ b\ have "poly p a > 0" by simp ultimately have "x > a" using \x \ a\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p a > 0\ have "sgn (poly p a) \ sgn (poly p x)" by simp from poly_different_sign_imp_root[OF \x > a\ this] and B and \x \ b\ show False by auto qed qed end diff --git a/thys/Sturm_Tarski/PolyMisc.thy b/thys/Sturm_Tarski/PolyMisc.thy --- a/thys/Sturm_Tarski/PolyMisc.thy +++ b/thys/Sturm_Tarski/PolyMisc.thy @@ -1,165 +1,165 @@ (* Author: Wenda Li *) theory PolyMisc imports "HOL-Computational_Algebra.Polynomial_Factorial" begin lemma coprime_poly_0: "poly p x \ 0 \ poly q x \ 0" if "coprime p q" for x :: "'a :: field" proof (rule ccontr) assume " \ (poly p x \ 0 \ poly q x \ 0)" then have "[:-x, 1:] dvd p" "[:-x, 1:] dvd q" by (simp_all add: poly_eq_0_iff_dvd) with that have "is_unit [:-x, 1:]" by (rule coprime_common_divisor) then show False by (auto simp add: is_unit_pCons_iff) qed lemma smult_cancel: fixes p::"'a::idom poly" assumes "c\0" and smult: "smult c p = smult c q" shows "p=q" proof - have "smult c (p-q)=0" using smult by (metis diff_self smult_diff_right) thus ?thesis using \c\0\ by auto qed lemma dvd_monic: fixes p q:: "'a :: idom poly" assumes monic:"lead_coeff p=1" and "p dvd (smult c q)" and "c\0" shows "p dvd q" using assms proof (cases "q=0 \ degree p=0") case True thus ?thesis using assms by (auto elim!: degree_eq_zeroE simp add: const_poly_dvd_iff) next case False hence "q\0" and "degree p\0" by auto obtain k where k:"smult c q = p*k" using assms dvd_def by metis hence "k\0" by (metis False assms(3) mult_zero_right smult_eq_0_iff) hence deg_eq:"degree q=degree p + degree k" by (metis False assms(3) degree_0 degree_mult_eq degree_smult_eq k) have c_dvd:"\n\degree k. c dvd coeff k (degree k - n)" proof (rule,rule) fix n assume "n \ degree k " thus "c dvd coeff k (degree k - n)" proof (induct n rule:nat_less_induct) case (1 n) define T where "T\(\i. coeff p i * coeff k (degree p+degree k - n - i))" have "c * coeff q (degree q - n) = (\i\degree q - n. coeff p i * coeff k (degree q - n - i))" using coeff_mult[of p k "degree q - n"] k coeff_smult[of c q "degree q -n"] by auto also have "...=(\i\degree p+degree k - n. T i)" using deg_eq unfolding T_def by auto also have "...=(\i\{0..{{0..A\C. finite A" unfolding C_def by auto moreover have "\A\C. \B\C. A \ B \ A \ B = {}" unfolding C_def by auto ultimately have "sum T (\C) = sum (sum T) C" using sum.Union_disjoint by auto moreover have "\C={..degree p + degree k - n}" using \n \ degree k\ unfolding C_def by auto moreover have "sum (sum T) C= sum T {0..{degree p}" by (metis atLeast0LessThan insertI1 lessThan_iff less_imp_not_eq) moreover have "{degree p}\{degree p + 1..degree p + degree k - n}" by (metis add.commute add_diff_cancel_right' atLeastAtMost_singleton_iff diff_self_eq_0 eq_imp_le not_one_le_zero) moreover have "{0..{degree p + 1..degree p + degree k - n}" using \degree k\n\ \degree p\0\ by fastforce ultimately show ?thesis unfolding C_def by auto qed ultimately show ?thesis by auto qed also have "...=(\i\{0..x\{degree p + 1..degree p + degree k - n}. T x=0" using coeff_eq_0[of p] unfolding T_def by simp hence "sum T {degree p + 1..degree p + degree k - n}=0" by auto moreover have "T (degree p)=coeff k (degree k - n)" using monic by (simp add: T_def) ultimately show ?thesis by auto qed finally have c_coeff: "c * coeff q (degree q - n) = sum T {0..0\c dvd sum T {0.. {0..0" hence "(n+i-degree p)\degree k" using \n \ degree k\ by auto moreover have "n + i - degree p n\0\ by auto ultimately have "c dvd coeff k (degree k - (n+i-degree p))" using 1(1) by auto hence "c dvd coeff k (degree p + degree k - n - i)" by (metis add_diff_cancel_left' deg_eq diff_diff_left dvd_0_right le_degree le_diff_conv add.commute ordered_cancel_comm_monoid_diff_class.diff_diff_right) thus "c dvd T i" unfolding T_def by auto qed moreover have "n=0 \?case" proof - assume "n=0" hence "\i\{0..n. c dvd coeff k n" by (metis diff_diff_cancel dvd_0_right le_add2 le_add_diff_inverse le_degree) then obtain f where f:"\n. c * f n=coeff k n" unfolding dvd_def by metis have " \\<^sub>\ n. f n = 0 " by (metis (mono_tags, lifting) MOST_coeff_eq_0 MOST_mono assms(3) f mult_eq_0_iff) hence "smult c (Abs_poly f)=k" using f smult.abs_eq[of c "Abs_poly f"] Abs_poly_inverse[of f] coeff_inverse[of k] by simp hence "q=p* Abs_poly f" using k \c\0\ smult_cancel by auto thus ?thesis unfolding dvd_def by auto qed lemma poly_power_n_eq: fixes x::"'a :: idom" assumes "n\0" shows "poly ([:-a,1:]^n) x=0 \ (x=a)" using assms by (induct n,auto) lemma poly_power_n_odd: fixes x a:: real assumes "odd n" shows "poly ([:-a,1:]^n) x>0 \ (x>a)" using assms proof - have "poly ([:-a,1:]^n) x\0 = (poly [:- a, 1:] x \0)" unfolding poly_power using zero_le_odd_power[OF \odd n\] by blast also have "(poly [:- a, 1:] x \0) = (x\a)" by fastforce finally have "poly ([:-a,1:]^n) x\0 = (x\a)" . moreover have "poly ([:-a,1:]^n) x=0 = (x=a)" by(rule poly_power_n_eq, metis assms even_zero) ultimately show ?thesis by linarith qed lemma gcd_coprime_poly: - fixes p q::"'a::factorial_ring_gcd poly" + fixes p q::"'a::{factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes nz: "p \ 0 \ q \ 0" and p': "p = p' * gcd p q" and q': "q = q' * gcd p q" shows "coprime p' q'" using gcd_coprime nz p' q' by auto lemma poly_mod: "poly (p mod q) x = poly p x" if "poly q x = 0" proof - from that have "poly (p mod q) x = poly (p div q * q) x + poly (p mod q) x" by simp also have "\ = poly p x" by (simp only: poly_add [symmetric]) simp finally show ?thesis . qed end diff --git a/thys/Subresultants/Subresultant_Gcd.thy b/thys/Subresultants/Subresultant_Gcd.thy --- a/thys/Subresultants/Subresultant_Gcd.thy +++ b/thys/Subresultants/Subresultant_Gcd.thy @@ -1,327 +1,327 @@ section \Computing the Gcd via the subresultant PRS\ text \This theory now formalizes how the subresultant PRS can be used to calculate the gcd of two polynomials. Moreover, it proves the connection between resultants and gcd, namely that the resultant is 0 iff the degree of the gcd is non-zero.\ theory Subresultant_Gcd imports Subresultant Polynomial_Factorization.Missing_Polynomial_Factorial begin subsection \Algorithm\ definition gcd_impl_primitive where [code del]: "gcd_impl_primitive G1 G2 = normalize (primitive_part (fst (subresultant_prs dichotomous_Lazard G1 G2)))" definition gcd_impl_main where [code del]: "gcd_impl_main G1 G2 = (if G1 = 0 then 0 else if G2 = 0 then normalize G1 else smult (gcd (content G1) (content G2)) (gcd_impl_primitive (primitive_part G1) (primitive_part G2)))" definition gcd_impl where "gcd_impl f g = (if length (coeffs f) \ length (coeffs g) then gcd_impl_main f g else gcd_impl_main g f)" subsection \Soundness Proof for @{term "gcd_impl = gcd"}\ locale subresultant_prs_gcd = subresultant_prs_locale2 F n \ f k \ G1 G2 for - F :: "nat \ 'a :: factorial_ring_gcd fract poly" + F :: "nat \ 'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} fract poly" and n :: "nat \ nat" and \ :: "nat \ nat" and f :: "nat \ 'a fract" and k :: nat and \ :: "nat \ 'a fract" and G1 G2 :: "'a poly" begin text \The subresultant PRS computes the gcd up to a scalar multiple.\ lemma subresultant_prs_gcd: assumes "subresultant_prs dichotomous_Lazard G1 G2 = (Gk, hk)" shows "\ a b. a \ 0 \ b \ 0 \ smult a (gcd G1 G2) = smult b (normalize Gk)" proof - from subresultant_prs[OF dichotomous_Lazard assms] have Fk: "F k = ffp Gk" and "\ i. \ H. i \ 0 \ F i = ffp H" and "\ i. \ b. 3 \ i \ i \ Suc k \ \ i = ff b" by auto from choice[OF this(2)] choice[OF this(3)] obtain H beta where FH: "\ i. i \ 0 \ F i = ffp (H i)" and beta: "\ i. 3 \ i \ i \ Suc k \ \ i = ff (beta i)" by auto from Fk FH[OF k0] FH[of 1] FH[of 2] FH[of "Suc k"] F0[of "Suc k"] F1 F2 have border: "H k = Gk" "H 1 = G1" "H 2 = G2" "H (Suc k) = 0" by auto have "i \ 0 \ i \ k \ \ a b. a \ 0 \ b \ 0 \ smult a (gcd G1 G2) = smult b (gcd (H i) (H (Suc i)))" for i proof (induct i rule: less_induct) case (less i) from less(3) have ik: "i \ k" . from less(2) have "i = 1 \ i \ 2" by auto thus ?case proof assume "i = 1" thus ?thesis unfolding border[symmetric] by (intro exI[of _ 1], auto simp: numeral_2_eq_2) next assume i2: "i \ 2" with ik have "i - 1 < i" "i - 1 \ 0" and imk: "i - 1 \ k" by auto from less(1)[OF this] i2 obtain a b where a: "a \ 0" and b: "b \ 0" and IH: "smult a (gcd G1 G2) = smult b (gcd (H (i - 1)) (H i))" by auto define M where "M = pseudo_mod (H (i - 1)) (H i)" define c where "c = \ (Suc i)" have M: "pseudo_mod (F (i - 1)) (F i) = ffp M" unfolding to_fract_hom.pseudo_mod_hom[symmetric] M_def using i2 FH by auto have c: "c \ 0" using \0 unfolding c_def . from i2 ik have 3: "Suc i \ 3" "Suc i \ Suc k" by auto from pmod[OF 3] have pm: "smult c (F (Suc i)) = pseudo_mod (F (i - 1)) (F i)" unfolding c_def by simp from beta[OF 3, folded c_def] obtain d where cd: "c = ff d" by auto with c have d: "d \ 0" by auto from pm[unfolded cd M] FH[of "Suc i"] have "ffp (smult d (H (Suc i))) = ffp M" by auto hence pm: "smult d (H (Suc i)) = M" by (rule map_poly_hom.injectivity) from ik F0[of i] i2 FH[of i] have Hi0: "H i \ 0" by auto from pseudo_mod[OF this, of "H (i - 1)", folded M_def] obtain c Q where c: "c \ 0" and "smult c (H (i - 1)) = H i * Q + M" by auto from this[folded pm] have "smult c (H (i - 1)) = Q * H i + smult d (H (Suc i))" by simp from gcd_add_mult[of "H i" Q "smult d (H (Suc i))", folded this] have "gcd (H i) (smult c (H (i - 1))) = gcd (H i) (smult d (H (Suc i)))" . with gcd_smult_ex[OF c, of "H (i - 1)" "H i"] obtain e where e: "e \ 0" and "gcd (H i) (smult d (H (Suc i))) = smult e (gcd (H i) (H (i - 1)))" unfolding gcd.commute[of "H i"] by auto with gcd_smult_ex[OF d, of "H (Suc i)" "H i"] obtain c where c: "c \ 0" and "smult c (gcd (H i) (H (Suc i))) = smult e (gcd (H (i - 1)) (H i))" unfolding gcd.commute[of "H i"] by auto from arg_cong[OF this(2), of "smult b"] arg_cong[OF IH, of "smult e"] have "smult (e * a) (gcd G1 G2) = smult (b * c) (gcd (H i) (H (Suc i)))" unfolding smult_smult by (simp add: ac_simps) moreover have "e * a \ 0" "b * c \ 0" using a b c e by auto ultimately show ?thesis by blast qed qed from this[OF k0 le_refl, unfolded border] obtain a b where "a \ 0" "b \ 0" and "smult a (gcd G1 G2) = smult b (normalize Gk)" by auto thus ?thesis by auto qed lemma gcd_impl_primitive: assumes "primitive_part G1 = G1" and "primitive_part G2 = G2" shows "gcd_impl_primitive G1 G2 = gcd G1 G2" proof - let ?pp = primitive_part let ?c = "content" let ?n = normalize from F2 F0[of 2] k2 have G2: "G2 \ 0" by auto obtain Gk hk where sub: "subresultant_prs dichotomous_Lazard G1 G2 = (Gk, hk)" by force have impl: "gcd_impl_primitive G1 G2 = ?n (?pp Gk)" unfolding gcd_impl_primitive_def sub by auto from subresultant_prs_gcd[OF sub] obtain a b where a: "a \ 0" and b: "b \ 0" and id: "smult a (gcd G1 G2) = smult b (?n Gk)" by auto define c where "c = unit_factor (gcd G1 G2)" define d where "d = smult (unit_factor a) c" from G2 have c: "is_unit c" unfolding c_def by auto from arg_cong[OF id, of ?pp, unfolded primitive_part_smult primitive_part_gcd assms primitive_part_normalize c_def[symmetric]] have id: "d * gcd G1 G2 = smult (unit_factor b) (?n (?pp Gk))" unfolding d_def by simp have d: "is_unit d" unfolding d_def using c a by (simp add: is_unit_smult_iff) from is_unitE[OF d] obtain e where e: "is_unit e" and de: "d * e = 1" by metis define a where "a = smult (unit_factor b) e" from arg_cong[OF id, of "\ x. e * x"] have "(d * e) * gcd G1 G2 = a * (?n (?pp Gk))" by (simp add: ac_simps a_def) hence id: "gcd G1 G2 = a * (?n (?pp Gk))" using de by simp have a: "is_unit a" unfolding a_def using b e by (simp add: is_unit_smult_iff) define b where "b = unit_factor (?pp Gk)" have "Gk \ 0" using subresultant_prs[OF dichotomous_Lazard sub] F0[OF k0] by auto hence b: "is_unit b" unfolding b_def by auto from is_unitE[OF b] obtain c where c: "is_unit c" and bc: "b * c = 1" by metis obtain d where d: "is_unit d" and dac: "d = a * c" using c a by auto have "gcd G1 G2 = d * (b * ?n (?pp Gk))" unfolding id dac using bc by (simp add: ac_simps) also have "b * ?n (?pp Gk) = ?pp Gk" unfolding b_def by simp finally have "gcd G1 G2 = d * ?pp Gk" by simp from arg_cong[OF this, of ?n] have "gcd G1 G2 = ?n (d * ?pp Gk)" by simp also have "\ = ?n (?pp Gk)" using d unfolding normalize_mult by (simp add: is_unit_normalize) finally show ?thesis unfolding impl .. qed end lemma gcd_impl_main: assumes len: "length (coeffs G1) \ length (coeffs G2)" shows "gcd_impl_main G1 G2 = gcd G1 G2" proof (cases "G1 = 0") case G1: False show ?thesis proof (cases "G2 = 0") case G2: False let ?pp = "primitive_part" from G2 have G2: "?pp G2 \ 0" and id: "(G2 = 0) = False" by auto from len have len: "length (coeffs (?pp G1)) \ length (coeffs (?pp G2))" by simp from enter_subresultant_prs[OF len G2] obtain F n d f k b where "subresultant_prs_locale2 F n d f k b (?pp G1) (?pp G2)" by auto interpret subresultant_prs_locale2 F n d f k b "?pp G1" "?pp G2" by fact interpret subresultant_prs_gcd F n d f k b "?pp G1" "?pp G2" .. show ?thesis unfolding gcd_impl_main_def gcd_poly_decompose[of G1] id if_False using G1 by (subst gcd_impl_primitive, auto) next case True thus ?thesis unfolding gcd_impl_main_def by simp qed next case True with len have "G2 = 0" by auto thus ?thesis using True unfolding gcd_impl_main_def by simp qed theorem gcd_impl[simp]: "gcd_impl = gcd" proof (intro ext) fix f g :: "'a poly" show "gcd_impl f g = gcd f g" proof (cases "length (coeffs f) \ length (coeffs g)") case True thus ?thesis unfolding gcd_impl_def gcd_impl_main[OF True] by auto next case False hence "length (coeffs g) \ length (coeffs f)" by auto from gcd_impl_main[OF this] show ?thesis unfolding gcd_impl_def gcd.commute[of f g] using False by auto qed qed text \The implementation also reveals an important connection between resultant and gcd.\ lemma resultant_0_gcd: "resultant f g = 0 \ degree (gcd f g) \ 0" proof - { fix f g :: "'a poly" assume len: "length (coeffs f) \ length (coeffs g)" { assume g: "g \ 0" with len have f: "f \ 0" by auto let ?f = "primitive_part f" let ?g = "primitive_part g" let ?c = "content" from len have len: "length (coeffs ?f) \ length (coeffs ?g)" by simp obtain Gk hk where sub: "subresultant_prs dichotomous_Lazard ?f ?g = (Gk,hk)" by force have cf: "?c f \ 0" and cg: "?c g \ 0" using f g by auto { from g have "?g \ 0" by auto from enter_subresultant_prs[OF len this] obtain F n d f k b where "subresultant_prs_locale2 F n d f k b ?f ?g" by auto interpret subresultant_prs_locale2 F n d f k b ?f ?g by fact from subresultant_prs[OF dichotomous_Lazard sub] have "h k = ff hk" by auto with h0[OF le_refl] have "hk \ 0" by auto } note hk0 = this have "resultant f g = 0 \ resultant (smult (?c f) ?f) (smult (?c g) ?g) = 0" by simp also have "\ \ resultant ?f ?g = 0" unfolding resultant_smult_left[OF cf] resultant_smult_right[OF cg] using cf cg by auto also have "\ \ resultant_impl_main dichotomous_Lazard ?f ?g = 0" unfolding resultant_impl[symmetric] resultant_impl_def resultant_impl_main_def resultant_impl_generic_def using len by auto also have "\ \ (degree Gk \ 0)" unfolding resultant_impl_main_def sub split using g hk0 by auto also have "degree Gk = degree (gcd_impl_primitive ?f ?g)" unfolding gcd_impl_primitive_def sub by simp also have "\ = degree (gcd_impl_main f g)" unfolding gcd_impl_main_def using f g by auto also have "\ = degree (gcd f g)" unfolding gcd_impl[symmetric] gcd_impl_def using len by auto finally have "(resultant f g = 0) = (degree (gcd f g) \ 0)" . } moreover { assume g: "g = 0" and f: "degree f \ 0" have "(resultant f g = 0) = (degree (gcd f g) \ 0)" unfolding g using f by auto } moreover { assume g: "g = 0" and f: "degree f = 0" have "(resultant f g = 0) = (degree (gcd f g) \ 0)" unfolding g using f by (auto simp: resultant_def sylvester_mat_def sylvester_mat_sub_def) } ultimately have "(resultant f g = 0) = (degree (gcd f g) \ 0)" by blast } note main = this show ?thesis proof (cases "length (coeffs f) \ length (coeffs g)") case True from main[OF True] show ?thesis . next case False hence "length (coeffs g) \ length (coeffs f)" by auto from main[OF this] show ?thesis unfolding gcd.commute[of g f] resultant_swap[of g f] by (simp split: if_splits) qed qed subsection \Code Equations\ definition [code del]: "gcd_impl_rec = subresultant_prs_main_impl fst" definition [code del]: "gcd_impl_start = subresultant_prs_impl fst" lemma gcd_impl_rec_code[code]: "gcd_impl_rec Gi_1 Gi ni_1 d1_1 hi_2 = ( let pmod = pseudo_mod Gi_1 Gi in if pmod = 0 then Gi else let ni = degree Gi; d1 = ni_1 - ni; gi_1 = lead_coeff Gi_1; hi_1 = (if d1_1 = 1 then gi_1 else dichotomous_Lazard gi_1 hi_2 d1_1); divisor = if d1 = 1 then gi_1 * hi_1 else if even d1 then - gi_1 * hi_1 ^ d1 else gi_1 * hi_1 ^ d1; Gi_p1 = sdiv_poly pmod divisor in gcd_impl_rec Gi Gi_p1 ni d1 hi_1)" unfolding gcd_impl_rec_def subresultant_prs_main_impl.simps[of _ Gi_1] split Let_def unfolding gcd_impl_rec_def[symmetric] by (rule if_cong, auto) lemma gcd_impl_start_code[code]: "gcd_impl_start G1 G2 = (let pmod = pseudo_mod G1 G2 in if pmod = 0 then G2 else let n2 = degree G2; n1 = degree G1; d1 = n1 - n2; G3 = if even d1 then - pmod else pmod; pmod = pseudo_mod G2 G3 in if pmod = 0 then G3 else let g2 = lead_coeff G2; n3 = degree G3; h2 = (if d1 = 1 then g2 else g2 ^ d1); d2 = n2 - n3; divisor = (if d2 = 1 then g2 * h2 else if even d2 then - g2 * h2 ^ d2 else g2 * h2 ^ d2); G4 = sdiv_poly pmod divisor in gcd_impl_rec G3 G4 n3 d2 h2)" proof - obtain d1 where d1: "degree G1 - degree G2 = d1" by auto have id1: "(if even d1 then - pmod else pmod) = (-1)^ (d1 + 1) * (pmod :: 'a poly)" for pmod by simp show ?thesis unfolding gcd_impl_start_def subresultant_prs_impl_def gcd_impl_rec_def[symmetric] Let_def split unfolding d1 unfolding id1 by (rule if_cong, auto) qed lemma gcd_impl_main_code[code]: "gcd_impl_main G1 G2 = (if G1 = 0 then 0 else if G2 = 0 then normalize G1 else let c1 = content G1; c2 = content G2; p1 = map_poly (\ x. x div c1) G1; p2 = map_poly (\ x. x div c2) G2 in smult (gcd c1 c2) (normalize (primitive_part (gcd_impl_start p1 p2))))" unfolding gcd_impl_main_def Let_def primitive_part_def gcd_impl_start_def gcd_impl_primitive_def subresultant_prs_impl by simp corollary gcd_via_subresultant: "gcd f g = gcd_impl f g" by simp text \Note that we did not activate @{thm gcd_via_subresultant} as code-equation, since according to our experiments, the subresultant-gcd algorithm is not always more efficient than the currently active equation. In particular, on @{typ "int poly"} @{const gcd_impl} performs worse, but on multi-variate polynomials, e.g., @{typ "int poly poly poly"}, @{const gcd_impl} is preferable.\ end