diff --git a/thys/Algebraic_Numbers/Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Algebraic_Numbers.thy @@ -1,746 +1,742 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Algebraic Numbers: Addition and Multiplication\ text \This theory contains the remaining field operations for algebraic numbers, namely addition and multiplication.\ theory Algebraic_Numbers -imports + imports Algebraic_Numbers_Prelim Resultant Polynomial_Factorization.Polynomial_Divisibility begin interpretation coeff_hom: monoid_add_hom "\p. coeff p i" by (unfold_locales, auto) interpretation coeff_hom: comm_monoid_add_hom "\p. coeff p i".. interpretation coeff_hom: group_add_hom "\p. coeff p i".. interpretation coeff_hom: ab_group_add_hom "\p. coeff p i".. interpretation coeff_0_hom: monoid_mult_hom "\p. coeff p 0" by (unfold_locales, auto simp: coeff_mult) interpretation coeff_0_hom: semiring_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_monoid_mult_hom "\p. coeff p 0".. interpretation coeff_0_hom: comm_semiring_hom "\p. coeff p 0".. subsection \Addition of Algebraic Numbers\ definition "x_y \ [: [: 0, 1 :], -1 :]" definition "poly_x_minus_y p = poly_lift p \\<^sub>p x_y" text \The following polynomial represents the sum of two algebraic numbers.\ definition poly_add :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_add p q = resultant (poly_x_minus_y p) (poly_lift q)" subsubsection \@{term poly_add} has desired root\ interpretation poly_x_minus_y_hom: comm_ring_hom poly_x_minus_y by (unfold_locales; simp add: poly_x_minus_y_def hom_distribs) lemma poly2_x_y[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 x_y x y = x - y" unfolding poly2_def by (simp add: x_y_def) lemma degree_poly_x_minus_y[simp]: fixes p :: "'a::idom poly" shows "degree (poly_x_minus_y p) = degree p" unfolding poly_x_minus_y_def x_y_def by auto lemma poly_x_minus_y_pCons[simp]: "poly_x_minus_y (pCons a p) = [:[: a :]:] + poly_x_minus_y p * x_y" unfolding poly_x_minus_y_def x_y_def by simp lemma poly_poly_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly (poly (poly_x_minus_y p) q) x = poly p (x - poly q x)" by (induct p; simp add: ring_distribs x_y_def) lemma poly2_poly_x_minus_y[simp]: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_minus_y p) x y = poly p (x-y)" unfolding poly2_def by simp interpretation x_y_mult_hom: zero_hom_0 "\p :: 'a :: comm_ring_1 poly poly. x_y * p" proof (unfold_locales) fix p :: "'a poly poly" assume "x_y * p = 0" then show "p = 0" apply (simp add: x_y_def) by (metis eq_neg_iff_add_eq_0 minus_equation_iff minus_pCons synthetic_div_unique_lemma) qed lemma x_y_nonzero[simp]: "x_y \ 0" by (simp add: x_y_def) lemma degree_x_y[simp]: "degree x_y = 1" by (simp add: x_y_def) interpretation x_y_mult_hom: inj_comm_monoid_add_hom "\p :: 'a :: idom poly poly. x_y * p" proof (unfold_locales) show "x_y * p = x_y * q \ p = q" for p q :: "'a poly poly" proof (induct p arbitrary:q) case 0 then show ?case by simp next case p: (pCons a p) from p(3)[unfolded mult_pCons_right] have "x_y * (monom a 0 + pCons 0 1 * p) = x_y * q" apply (subst(asm) pCons_0_as_mult) apply (subst(asm) smult_prod) by (simp only: field_simps distrib_left) then have "monom a 0 + pCons 0 1 * p = q" by simp then show "pCons a p = q" using pCons_as_add by (simp add: monom_0 monom_Suc) qed qed interpretation poly_x_minus_y_hom: inj_idom_hom poly_x_minus_y proof fix p :: "'a poly" assume 0: "poly_x_minus_y p = 0" then have "poly_lift p \\<^sub>p x_y = 0" by (simp add: poly_x_minus_y_def) then show "p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) note p = this[unfolded poly_lift_pCons pcompose_pCons] show ?case proof (cases "a=0") case a0: True with p have "x_y * poly_lift p \\<^sub>p x_y = 0" by simp then have "poly_lift p \\<^sub>p x_y = 0" by simp then show ?thesis using p by simp next case a0: False with p have p0: "p \ 0" by auto from p have "[:[:a:]:] = - x_y * poly_lift p \\<^sub>p x_y" by (simp add: eq_neg_iff_add_eq_0) then have "degree [:[:a:]:] = degree (x_y * poly_lift p \\<^sub>p x_y)" by simp also have "... = degree (x_y::'a poly poly) + degree (poly_lift p \\<^sub>p x_y)" apply (subst degree_mult_eq) apply simp apply (subst pcompose_eq_0) apply (simp add: x_y_def) apply (simp add: p0) apply simp done finally have False by simp then show ?thesis.. qed qed qed lemma poly_add: fixes p q :: "'a ::comm_ring_1 poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" shows "poly (poly_add p q) (x+y) = 0" proof (unfold poly_add_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y, simp_all) subsubsection \@{const poly_add} is nonzero\ text \ We first prove that @{const poly_lift} preserves factorization. The result will be essential also in the next section for division of algebraic numbers. \ -interpretation coeff_lift_hom: - factor_preserving_hom "coeff_lift :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} \ _" - by (unfold_locales, auto) - interpretation poly_lift_hom: unit_preserving_hom "poly_lift :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly \ _" proof fix x :: "'a poly" assume "poly_lift x dvd 1" then have "poly_y_x (poly_lift x) dvd poly_y_x 1" by simp then show "x dvd 1" by (auto simp add: poly_y_x_poly_lift) qed interpretation poly_lift_hom: factor_preserving_hom "poly_lift::'a::idom poly \ 'a poly poly" proof unfold_locales fix p :: "'a poly" assume p: "irreducible p" show "irreducible (poly_lift p)" proof(rule ccontr) from p have p0: "p \ 0" and "\ p dvd 1" by (auto dest: irreducible_not_unit) with poly_lift_hom.hom_dvd[of p 1] have p1: "\ poly_lift p dvd 1" by auto assume "\ irreducible (poly_lift p)" from this[unfolded irreducible_altdef,simplified] p0 p1 obtain q where "q dvd poly_lift p" and pq: "\ poly_lift p dvd q" and q: "\ q dvd 1" by auto then obtain r where "q * r = poly_lift p" by (elim dvdE, auto) then have "poly_y_x (q * r) = poly_y_x (poly_lift p)" by auto also have "... = [:p:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "poly_y_x (q * r) = poly_y_x q * poly_y_x r" by (auto simp: hom_distribs) finally have "... = [:p:]" by auto then have qp: "poly_y_x q dvd [:p:]" by (metis dvdI) from dvd_const[OF this] p0 have "degree (poly_y_x q) = 0" by auto from degree_0_id[OF this,symmetric] obtain s where qs: "poly_y_x q = [:s:]" by auto have "poly_lift s = poly_y_x (poly_y_x (poly_lift s))" by auto also have "... = poly_y_x [:s:]" by (auto simp: poly_y_x_poly_lift monom_0) also have "... = q" by (auto simp: qs[symmetric]) finally have sq: "poly_lift s = q" by auto from qp[unfolded qs] have sp: "s dvd p" by (auto simp: const_poly_dvd) from irreducibleD'[OF p this] sq q pq show False by auto qed qed text \ We now show that @{const poly_x_minus_y} is a factor-preserving homomorphism. This is essential for this section. This is easy since @{const poly_x_minus_y} can be represented as the composition of two factor-preserving homomorphisms. \ lemma poly_x_minus_y_as_comp: "poly_x_minus_y = (\p. p \\<^sub>p x_y) \ poly_lift" by (intro ext, unfold poly_x_minus_y_def, auto) context idom_isom begin sublocale comm_semiring_isom.. end interpretation poly_x_minus_y_hom: factor_preserving_hom "poly_x_minus_y :: 'a :: idom poly \ 'a poly poly" proof- interpret x_y_hom: bijective "\p :: 'a poly poly. p \\<^sub>p x_y" proof (unfold bijective_eq_bij, rule id_imp_bij) fix p :: "'a poly poly" show "p \\<^sub>p x_y \\<^sub>p x_y = p" apply (induct p,simp) apply(unfold x_y_def hom_distribs pcompose_pCons) by (simp) qed interpret x_y_hom: idom_isom "\p :: 'a poly poly. p \\<^sub>p x_y" by (unfold_locales, auto) show "factor_preserving_hom (poly_x_minus_y :: 'a poly \ _)" by (unfold poly_x_minus_y_as_comp, rule factor_preserving_hom_comp, unfold_locales) qed text \ Now we show that results of @{const poly_x_minus_y} and @{const poly_lift} are coprime. \ lemma poly_y_x_const[simp]: "poly_y_x [:[:a:]:] = [:[:a:]:]" by (simp add: poly_y_x_def monom_0) context begin private abbreviation "y_x == [: [: 0, -1 :], 1 :]" lemma poly_y_x_x_y[simp]: "poly_y_x x_y = y_x" by (simp add: x_y_def poly_y_x_def monom_Suc monom_0) private lemma y_x[simp]: fixes x :: "'a :: comm_ring_1" shows "poly2 y_x x y = y - x" unfolding poly2_def by simp private definition "poly_y_minus_x p \ poly_lift p \\<^sub>p y_x" private lemma poly_y_minus_x_0[simp]: "poly_y_minus_x 0 = 0" by (simp add: poly_y_minus_x_def) private lemma poly_y_minus_x_pCons[simp]: "poly_y_minus_x (pCons a p) = [:[: a :]:] + poly_y_minus_x p * y_x" by (simp add: poly_y_minus_x_def) private lemma poly_y_x_poly_x_minus_y: fixes p :: "'a :: idom poly" shows "poly_y_x (poly_x_minus_y p) = poly_y_minus_x p" apply (induct p, simp) apply (unfold poly_x_minus_y_pCons hom_distribs) by simp lemma degree_poly_y_minus_x[simp]: fixes p :: "'a :: idom poly" shows "degree (poly_y_x (poly_x_minus_y p)) = degree p" by (simp add: poly_y_minus_x_def poly_y_x_poly_x_minus_y) end lemma dvd_all_coeffs_iff: fixes x :: "'a :: comm_semiring_1" (* No addition needed! *) shows "(\pi \ set (coeffs p). x dvd pi) \ (\i. x dvd coeff p i)" (is "?l = ?r") proof- have "?r = (\i\{..degree p} \ {Suc (degree p)..}. x dvd coeff p i)" by auto also have "... = (\i\degree p. x dvd coeff p i)" by (auto simp add: ball_Un coeff_eq_0) also have "... = ?l" by (auto simp: coeffs_def) finally show ?thesis.. qed lemma primitive_imp_no_constant_factor: fixes p :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} poly" assumes pr: "primitive p" and F: "mset_factors F p" and fF: "f \# F" shows "degree f \ 0" proof from F fF have irr: "irreducible f" and fp: "f dvd p" by (auto dest: mset_factors_imp_dvd) assume deg: "degree f = 0" then obtain f0 where f0: "f = [:f0:]" by (auto dest: degree0_coeffs) with fp have "[:f0:] dvd p" by simp then have "f0 dvd coeff p i" for i by (simp add: const_poly_dvd_iff) with primitiveD[OF pr] dvd_all_coeffs_iff have "f0 dvd 1" by (auto simp: coeffs_def) with f0 irr show False by auto qed lemma coprime_poly_x_minus_y_poly_lift: fixes p q :: "'a :: ufd poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and pr: "primitive p" shows "coprime (poly_x_minus_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto with poly_x_minus_y_hom.hom_mset_factors have pF: "mset_factors (image_mset poly_x_minus_y F) (poly_x_minus_y p)" by auto from degq have q: "\ q dvd 1" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto assume "\ coprime (poly_x_minus_y p) (poly_lift q)" from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd (poly_x_minus_y p)" and rq: "r dvd (poly_lift q)" and rU: "\ r dvd 1" by auto note poly_lift_hom.hom_dvd from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd (poly_x_minus_y p)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pF] p0 obtain f where f: "f \# F" and fh: "poly_x_minus_y f ddvd h" by auto from hr rq have "h dvd (poly_lift q)" by (rule dvd_trans) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "poly_lift g ddvd h" by auto from fh gh have "poly_x_minus_y f ddvd poly_lift g" using ddvd_trans by auto then have "poly_y_x (poly_x_minus_y f) ddvd poly_y_x (poly_lift g)" by simp also have "poly_y_x (poly_lift g) = [:g:]" unfolding poly_y_x_poly_lift monom_0 by auto finally have ddvd: "poly_y_x (poly_x_minus_y f) ddvd [:g:]" by auto then have "degree (poly_y_x (poly_x_minus_y f)) = 0" by (metis degree_pCons_0 dvd_0_left_iff dvd_const) then have "degree f = 0" by simp with primitive_imp_no_constant_factor[OF pr F f] show False by auto qed lemma poly_add_nonzero: fixes p q :: "'a :: ufd poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and pr: "primitive p" shows "poly_add p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_add p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_add_def]] degp and coprime_poly_x_minus_y_poly_lift[OF degp degq pr] show False by auto qed subsubsection \Summary for addition\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in comm_ring_hom) map_poly_x_minus_y: "map_poly (map_poly hom) (poly_x_minus_y p) = poly_x_minus_y (map_poly hom p)" proof- interpret mp: map_poly_comm_ring_hom hom.. interpret mmp: map_poly_comm_ring_hom "map_poly hom".. show ?thesis apply (induct p, simp) apply(unfold x_y_def hom_distribs poly_x_minus_y_pCons, simp) done qed lemma (in comm_ring_hom) hom_poly_lift[simp]: "map_poly (map_poly hom) (poly_lift q) = poly_lift (map_poly hom q)" proof - show ?thesis unfolding poly_lift_def unfolding map_poly_map_poly[of coeff_lift,OF coeff_lift_hom.hom_zero] unfolding map_poly_coeff_lift_hom by simp qed lemma lead_coeff_poly_x_minus_y: fixes p :: "'a::idom poly" shows "lead_coeff (poly_x_minus_y p) = [:lead_coeff p * ((- 1) ^ degree p):]" (is "?l = ?r") proof- have "?l = Polynomial.smult (lead_coeff p) ((- 1) ^ degree p)" by (unfold poly_x_minus_y_def, subst lead_coeff_comp; simp add: x_y_def) also have "... = ?r" by (unfold hom_distribs, simp add: smult_as_map_poly[symmetric]) finally show ?thesis. qed lemma (in idom_hom) poly_add_hom: assumes p0: "hom (lead_coeff p) \ 0" and q0: "hom (lead_coeff q) \ 0" shows "map_poly hom (poly_add p q) = poly_add (map_poly hom p) (map_poly hom q)" proof - interpret mh: map_poly_idom_hom.. show ?thesis unfolding poly_add_def apply (subst mh.resultant_map_poly(1)[symmetric]) apply (subst degree_map_poly_2) apply (unfold lead_coeff_poly_x_minus_y, unfold hom_distribs, simp add: p0) apply simp apply (subst degree_map_poly_2) apply (simp_all add: q0 map_poly_x_minus_y) done qed lemma(in zero_hom) hom_lead_coeff_nonzero_imp_map_poly_hom: assumes "hom (lead_coeff p) \ 0" shows "map_poly hom p \ 0" proof assume "map_poly hom p = 0" then have "coeff (map_poly hom p) (degree p) = 0" by simp with assms show False by simp qed lemma ipoly_poly_add: fixes x y :: "'a :: idom" assumes p0: "(of_int (lead_coeff p) :: 'a) \ 0" and q0: "(of_int (lead_coeff q) :: 'a) \ 0" and x: "ipoly p x = 0" and y: "ipoly q y = 0" shows "ipoly (poly_add p q) (x+y) = 0" using assms of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom[OF q0] by (auto intro: poly_add simp: of_int_hom.poly_add_hom[OF p0 q0]) lemma (in comm_monoid_gcd) gcd_list_eq_0_iff[simp]: "listgcd xs = 0 \ (\x \ set xs. x = 0)" by (induct xs, auto) lemma primitive_field_poly[simp]: "primitive (p :: 'a :: field poly) \ p \ 0" by (unfold primitive_iff_some_content_dvd_1,auto simp: dvd_field_iff coeffs_def) lemma ipoly_poly_add_nonzero: fixes x y :: "'a :: field" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "(of_int (lead_coeff p) :: 'a) \ 0" and "(of_int (lead_coeff q) :: 'a) \ 0" shows "poly_add p q \ 0" proof- from assms have "(of_int_poly (poly_add p q) :: 'a poly) \ 0" apply (subst of_int_hom.poly_add_hom,simp,simp) by (rule poly_add_nonzero, auto dest:of_int_hom.hom_lead_coeff_nonzero_imp_map_poly_hom) then show ?thesis by auto qed lemma represents_add: assumes x: "p represents x" and y: "q represents y" shows "(poly_add p q) represents (x + y)" using assms by (intro representsI ipoly_poly_add ipoly_poly_add_nonzero, auto) subsection \Division of Algebraic Numbers\ definition poly_x_mult_y where [code del]: "poly_x_mult_y p \ (\ i \ degree p. monom (monom (coeff p i) i) i)" lemma coeff_poly_x_mult_y: shows "coeff (poly_x_mult_y p) i = monom (coeff p i) i" (is "?l = ?r") proof(cases "degree p < i") case i: False have "?l = sum (\j. if j = i then (monom (coeff p j) j) else 0) {..degree p}" (is "_ = sum ?f ?A") by (simp add: poly_x_mult_y_def coeff_sum) also have "... = sum ?f {i}" using i by (intro sum.mono_neutral_right, auto) also have "... = ?f i" by simp also have "... = ?r" by auto finally show ?thesis. next case True then show ?thesis by (auto simp: poly_x_mult_y_def coeff_eq_0 coeff_sum) qed lemma poly_x_mult_y_code[code]: "poly_x_mult_y p = (let cs = coeffs p in poly_of_list (map (\ (i, ai). monom ai i) (zip [0 ..< length cs] cs)))" unfolding Let_def poly_of_list_def proof (rule poly_eqI, unfold coeff_poly_x_mult_y) fix n let ?xs = "zip [0.. degree p \ p = 0" unfolding degree_eq_length_coeffs by (cases n, auto) hence "monom (coeff p n) n = 0" using coeff_eq_0[of p n] by auto thus ?thesis unfolding id by simp qed qed definition poly_div :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_div p q = resultant (poly_x_mult_y p) (poly_lift q)" text \@{const poly_div} has desired roots.\ lemma poly2_poly_x_mult_y: fixes p :: "'a :: comm_ring_1 poly" shows "poly2 (poly_x_mult_y p) x y = poly p (x * y)" apply (subst(3) poly_as_sum_of_monoms[symmetric]) apply (unfold poly_x_mult_y_def hom_distribs) by (auto simp: poly2_monom poly_monom power_mult_distrib ac_simps) lemma poly_div: fixes p q :: "'a ::field poly" assumes q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and y0: "y \ 0" shows "poly (poly_div p q) (x/y) = 0" proof (unfold poly_div_def, rule poly_resultant_zero[OF disjI2]) have "degree q > 0" using poly_zero q0 y by auto thus degq: "degree (poly_lift q) > 0" by auto qed (insert x y y0, simp_all add: poly2_poly_x_mult_y) text \@{const poly_div} is nonzero.\ interpretation poly_x_mult_y_hom: ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" by (unfold_locales, auto intro: poly2_ext simp: poly2_poly_x_mult_y hom_distribs) interpretation poly_x_mult_y_hom: inj_ring_hom "poly_x_mult_y :: 'a :: {idom,ring_char_0} poly \ _" proof let ?h = poly_x_mult_y fix f :: "'a poly" assume "?h f = 0" then have "poly2 (?h f) x 1 = 0" for x by simp from this[unfolded poly2_poly_x_mult_y] show "f = 0" by auto qed lemma degree_poly_x_mult_y[simp]: fixes p :: "'a :: {idom, ring_char_0} poly" shows "degree (poly_x_mult_y p) = degree p" (is "?l = ?r") proof(rule antisym) show "?r \ ?l" by (cases "p=0", auto intro: le_degree simp: coeff_poly_x_mult_y) show "?l \ ?r" unfolding poly_x_mult_y_def by (auto intro: degree_sum_le le_trans[OF degree_monom_le]) qed interpretation poly_x_mult_y_hom: unit_preserving_hom "poly_x_mult_y :: 'a :: field_char_0 poly \ _" proof(unfold_locales) let ?h = "poly_x_mult_y :: 'a poly \ _" fix f :: "'a poly" assume unit: "?h f dvd 1" then have "degree (?h f) = 0" and "coeff (?h f) 0 dvd 1" unfolding poly_dvd_1 by auto then have deg: "degree f = 0" by (auto simp add: degree_monom_eq) with unit show "f dvd 1" by(cases "f = 0", auto) qed lemmas poly_y_x_o_poly_lift = o_def[of poly_y_x poly_lift, unfolded poly_y_x_poly_lift] lemma irreducible_dvd_degree: assumes "(f::'a::field poly) dvd g" "irreducible g" "degree f > 0" shows "degree f = degree g" using assms by (metis irreducible_altdef degree_0 dvd_refl is_unit_field_poly linorder_neqE_nat poly_divides_conv0) lemma coprime_poly_x_mult_y_poly_lift: fixes p q :: "'a :: field_char_0 poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and nz: "poly p 0 \ 0 \ poly q 0 \ 0" shows "coprime (poly_x_mult_y p) (poly_lift q)" proof(rule ccontr) from degp have p: "\ p dvd 1" by (auto simp: dvd_const) from degp have p0: "p \ 0" by auto from mset_factors_exist[of p, OF p0 p] obtain F where F: "mset_factors F p" by auto then have pF: "prod_mset (image_mset poly_x_mult_y F) = poly_x_mult_y p" by (auto simp: hom_distribs) from degq have q: "\ is_unit q" by (auto simp: dvd_const) from degq have q0: "q \ 0" by auto from mset_factors_exist[OF q0 q] obtain G where G: "mset_factors G q" by auto with poly_lift_hom.hom_mset_factors have pG: "mset_factors (image_mset poly_lift G) (poly_lift q)" by auto from poly_y_x_hom.hom_mset_factors[OF this] have pG: "mset_factors (image_mset coeff_lift G) [:q:]" by (auto simp: poly_y_x_poly_lift monom_0 image_mset.compositionality poly_y_x_o_poly_lift) assume "\ coprime (poly_x_mult_y p) (poly_lift q)" then have "\ coprime (poly_y_x (poly_x_mult_y p)) (poly_y_x (poly_lift q))" by (simp del: coprime_iff_coprime) from this[unfolded not_coprime_iff_common_factor] obtain r where rp: "r dvd poly_y_x (poly_x_mult_y p)" and rq: "r dvd poly_y_x (poly_lift q)" and rU: "\ r dvd 1" by auto from rp p0 have r0: "r \ 0" by auto from mset_factors_exist[OF r0 rU] obtain H where H: "mset_factors H r" by auto then have "H \ {#}" by auto then obtain h where hH: "h \# H" by fastforce with H mset_factors_imp_dvd have hr: "h dvd r" and h: "irreducible h" by auto from irreducible_not_unit[OF h] have hU: "\ h dvd 1" by auto from hr rp have "h dvd poly_y_x (poly_x_mult_y p)" by (rule dvd_trans) note this[folded pF,unfolded poly_y_x_hom.hom_prod_mset image_mset.compositionality] from prime_elem_dvd_prod_mset[OF h[folded prime_elem_iff_irreducible] this] obtain f where f: "f \# F" and hf: "h dvd poly_y_x (poly_x_mult_y f)" by auto have irrF: "irreducible f" using f F by blast from dvd_trans[OF hr rq] have "h dvd [:q:]" by (simp add: poly_y_x_poly_lift monom_0) from irreducible_dvd_imp_factor[OF this h pG] q0 obtain g where g: "g \# G" and gh: "[:g:] dvd h" by auto from dvd_trans[OF gh hf] have *: "[:g:] dvd poly_y_x (poly_x_mult_y f)" using dvd_trans by auto show False proof (cases "poly f 0 = 0") case f_0: False from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) [:0:]" by simp also have "... = [:poly f 0:]" by (intro poly_ext, fold poly2_def, simp add: poly2_poly_x_mult_y) also have "... dvd 1" using f_0 by auto finally have "g dvd 1". with g G show False by (auto elim!: mset_factorsE dest!: irreducible_not_unit) next case True hence "[:0,1:] dvd f" by (unfold dvd_iff_poly_eq_0, simp) from irreducible_dvd_degree[OF this irrF] have "degree f = 1" by auto from degree1_coeffs[OF this] True obtain c where c: "c \ 0" and f: "f = [:0,c:]" by auto from g G have irrG: "irreducible g" by auto from poly_hom.hom_dvd[OF *] have "g dvd poly (poly_y_x (poly_x_mult_y f)) 1" by simp also have "\ = f" by (auto simp: f poly_x_mult_y_code Let_def c poly_y_x_pCons map_poly_monom poly_monom poly_lift_def) also have "\ dvd [:0,1:]" unfolding f dvd_def using c by (intro exI[of _ "[: inverse c :]"], auto) finally have g01: "g dvd [:0,1:]" . from divides_degree[OF this] irrG have "degree g = 1" by auto from degree1_coeffs[OF this] obtain a b where g: "g = [:b,a:]" and a: "a \ 0" by auto from g01[unfolded dvd_def] g obtain k where id: "[:0,1:] = g * k" by auto from id have 0: "g \ 0" "k \ 0" by auto from arg_cong[OF id, of degree] have "degree k = 0" unfolding degree_mult_eq[OF 0] unfolding g using a by auto from degree0_coeffs[OF this] obtain kk where k: "k = [:kk:]" by auto from id[unfolded g k] a have "b = 0" by auto hence "poly g 0 = 0" by (auto simp: g) from True this nz \f \# F\ \g \# G\ F G show False by (auto dest!:mset_factors_imp_dvd elim:dvdE) qed qed lemma poly_div_nonzero: fixes p q :: "'a :: field_char_0 poly" assumes p0: "p \ 0" and q0: "q \ 0" and x: "poly p x = 0" and y: "poly q y = 0" and p_0: "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof have degp: "degree p > 0" using le_0_eq order_degree order_root p0 x by (metis gr0I) have degq: "degree q > 0" using le_0_eq order_degree order_root q0 y by (metis gr0I) assume 0: "poly_div p q = 0" from resultant_zero_imp_common_factor[OF _ this[unfolded poly_div_def]] degp and coprime_poly_x_mult_y_poly_lift[OF degp degq] p_0 show False by auto qed subsubsection \Summary for division\ text \Now we lift the results to one that uses @{const ipoly}, by showing some homomorphism lemmas.\ lemma (in inj_comm_ring_hom) poly_x_mult_y_hom: "poly_x_mult_y (map_poly hom p) = map_poly (map_poly hom) (poly_x_mult_y p)" proof - interpret mh: map_poly_inj_comm_ring_hom.. interpret mmh: map_poly_inj_comm_ring_hom "map_poly hom".. show ?thesis unfolding poly_x_mult_y_def by (simp add: hom_distribs) qed lemma (in inj_comm_ring_hom) poly_div_hom: "map_poly hom (poly_div p q) = poly_div (map_poly hom p) (map_poly hom q)" proof - have zero: "\x. hom x = 0 \ x = 0" by simp interpret mh: map_poly_inj_comm_ring_hom.. show ?thesis unfolding poly_div_def mh.resultant_hom[symmetric] by (simp add: poly_x_mult_y_hom) qed lemma ipoly_poly_div: fixes x y :: "'a :: field_char_0" assumes "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "y \ 0" shows "ipoly (poly_div p q) (x/y) = 0" by (unfold of_int_hom.poly_div_hom, rule poly_div, insert assms, auto) lemma ipoly_poly_div_nonzero: fixes x y :: "'a :: field_char_0" assumes "p \ 0" and "q \ 0" and "ipoly p x = 0" and "ipoly q y = 0" and "poly p 0 \ 0 \ poly q 0 \ 0" shows "poly_div p q \ 0" proof- from assms have "(of_int_poly (poly_div p q) :: 'a poly) \ 0" using of_int_hom.poly_map_poly[of p] by (subst of_int_hom.poly_div_hom, subst poly_div_nonzero, auto) then show ?thesis by auto qed lemma represents_div: fixes x y :: "'a :: field_char_0" assumes "p represents x" and "q represents y" and "poly q 0 \ 0" shows "(poly_div p q) represents (x / y)" using assms by (intro representsI ipoly_poly_div ipoly_poly_div_nonzero, auto) subsection \Multiplication of Algebraic Numbers\ definition poly_mult where "poly_mult p q \ poly_div p (reflect_poly q)" lemma represents_mult: assumes px: "p represents x" and qy: "q represents y" and q_0: "poly q 0 \ 0" shows "(poly_mult p q) represents (x * y)" proof- from q_0 qy have y0: "y \ 0" by auto from represents_inverse[OF y0 qy] y0 px q_0 have "poly_mult p q represents x / (inverse y)" unfolding poly_mult_def by (intro represents_div, auto) with y0 show ?thesis by (simp add: field_simps) qed subsection \Summary: Closure Properties of Algebraic Numbers\ lemma algebraic_representsI: "p represents x \ algebraic x" unfolding represents_def algebraic_altdef_ipoly by auto lemma algebraic_of_rat: "algebraic (of_rat x)" by (rule algebraic_representsI[OF poly_rat_represents_of_rat]) lemma algebraic_uminus: "algebraic x \ algebraic (-x)" by (auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_uminus) lemma algebraic_inverse: "algebraic x \ algebraic (inverse x)" using algebraic_of_rat[of 0] by (cases "x = 0", auto dest: algebraic_imp_represents_irreducible intro: algebraic_representsI represents_inverse) lemma algebraic_plus: "algebraic x \ algebraic y \ algebraic (x + y)" by (auto dest!: algebraic_imp_represents_irreducible_cf_pos intro!: algebraic_representsI[OF represents_add]) lemma algebraic_div: assumes x: "algebraic x" and y: "algebraic y" shows "algebraic (x/y)" proof(cases "y = 0 \ x = 0") case True then show ?thesis using algebraic_of_rat[of 0] by auto next case False then have x0: "x \ 0" and y0: "y \ 0" by auto from x y obtain p q where px: "p represents x" and irr: "irreducible q" and qy: "q represents y" by (auto dest!: algebraic_imp_represents_irreducible) show ?thesis using False px represents_irr_non_0[OF irr qy] by (auto intro!: algebraic_representsI[OF represents_div] qy) qed lemma algebraic_times: "algebraic x \ algebraic y \ algebraic (x * y)" using algebraic_div[OF _ algebraic_inverse, of x y] by (simp add: field_simps) lemma algebraic_root: "algebraic x \ algebraic (root n x)" proof - assume "algebraic x" then obtain p where p: "p represents x" by (auto dest: algebraic_imp_represents_irreducible_cf_pos) from algebraic_representsI[OF represents_nth_root_neg_real[OF _ this, of n]] algebraic_representsI[OF represents_nth_root_pos_real[OF _ this, of n]] algebraic_of_rat[of 0] show ?thesis by (cases "n = 0", force, cases "n > 0", force, cases "n < 0", auto) qed lemma algebraic_nth_root: "n \ 0 \ algebraic x \ y^n = x \ algebraic y" by (auto dest: algebraic_imp_represents_irreducible_cf_pos intro: algebraic_representsI represents_nth_root) hide_const x_y end 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,1082 @@ (* 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 + Polynomial_Interpolation.Newton_Interpolation (* for lemma smult_1 *) + Polynomial_Factorization.Gauss_Lemma Berlekamp_Zassenhaus.Unique_Factorization_Poly + Polynomial_Factorization.Square_Free_Factorization 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)" +lemma irr_cf_poly_rat[simp]: "irreducible (poly_rat x)" + "lead_coeff (poly_rat x) > 0" "primitive (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" + from irr_cf_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, 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 :: {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,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,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,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,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 + then show "reflect_poly p \ 0" by (metis reflect_poly_0 reflect_poly_at_0_eq_0_iff) 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,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,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/Complex_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy @@ -1,1130 +1,1135 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Complex Algebraic Numbers\ text \Since currently there is no immediate analog of Sturm's theorem for the complex numbers, we implement complex algebraic numbers via their real and imaginary part. The major algorithm in this theory is a factorization algorithm which factors a rational polynomial over the complex numbers. This algorithm is then combined with explicit root algorithms to try to factor arbitrary complex polymials.\ theory Complex_Algebraic_Numbers imports Real_Roots Complex_Roots_Real_Poly Compare_Complex Jordan_Normal_Form.Char_Poly Berlekamp_Zassenhaus.Code_Abort_Gcd Interval_Arithmetic begin subsection \Complex Roots\ + +hide_const (open) UnivPoly.coeff +hide_const (open) Module.smult +hide_const (open) Coset.order + abbreviation complex_of_int_poly :: "int poly \ complex poly" where "complex_of_int_poly \ map_poly of_int" abbreviation complex_of_rat_poly :: "rat poly \ complex poly" where "complex_of_rat_poly \ map_poly of_rat" lemma poly_complex_to_real: "(poly (complex_of_int_poly p) (complex_of_real x) = 0) = (poly (real_of_int_poly p) x = 0)" proof - have id: "of_int = complex_of_real o real_of_int" by auto interpret cr: semiring_hom complex_of_real by (unfold_locales, auto) show ?thesis unfolding id by (subst map_poly_map_poly[symmetric], force+) qed lemma represents_cnj: assumes "p represents x" shows "p represents (cnj x)" proof - from assms have p: "p \ 0" and "ipoly p x = 0" by auto hence rt: "poly (complex_of_int_poly p) x = 0" by auto have "poly (complex_of_int_poly p) (cnj x) = 0" by (rule complex_conjugate_root[OF _ rt], subst coeffs_map_poly, auto) with p show ?thesis by auto qed definition poly_2i :: "int poly" where "poly_2i \ [: 4, 0, 1:]" lemma represents_2i: "poly_2i represents (2 * \)" unfolding represents_def poly_2i_def by simp definition root_poly_Re :: "int poly \ int poly" where "root_poly_Re p = cf_pos_poly (poly_mult_rat (inverse 2) (poly_add p p))" lemma root_poly_Re_code[code]: "root_poly_Re p = (let fs = coeffs (poly_add p p); k = length fs in cf_pos_poly (poly_of_list (map (\(fi, i). fi * 2 ^ i) (zip fs [0.. int poly list" where "root_poly_Im p = (let fs = factors_of_int_poly (poly_add p (poly_uminus p)) in remdups ((if (\ f \ set fs. coeff f 0 = 0) then [[:0,1:]] else [])) @ [ cf_pos_poly (poly_div f poly_2i) . f \ fs, coeff f 0 \ 0])" lemma represents_root_poly: assumes "ipoly p x = 0" and p: "p \ 0" shows "(root_poly_Re p) represents (Re x)" and "\ q \ set (root_poly_Im p). q represents (Im x)" proof - let ?Rep = "root_poly_Re p" let ?Imp = "root_poly_Im p" from assms have ap: "p represents x" by auto from represents_cnj[OF this] have apc: "p represents (cnj x)" . from represents_mult_rat[OF _ represents_add[OF ap apc], of "inverse 2"] have "?Rep represents (1 / 2 * (x + cnj x))" unfolding root_poly_Re_def Let_def by (auto simp: hom_distribs) also have "1 / 2 * (x + cnj x) = of_real (Re x)" by (simp add: complex_add_cnj) finally have Rep: "?Rep \ 0" and rt: "ipoly ?Rep (complex_of_real (Re x)) = 0" unfolding represents_def by auto from rt[unfolded poly_complex_to_real] have "ipoly ?Rep (Re x) = 0" . with Rep show "?Rep represents (Re x)" by auto let ?q = "poly_add p (poly_uminus p)" from represents_add[OF ap, of "poly_uminus p" "- cnj x"] represents_uminus[OF apc] have apq: "?q represents (x - cnj x)" by auto from factors_int_poly_represents[OF this] obtain pi where pi: "pi \ set (factors_of_int_poly ?q)" and appi: "pi represents (x - cnj x)" and irr_pi: "irreducible pi" by auto have id: "inverse (2 * \) * (x - cnj x) = of_real (Im x)" apply (cases x) by (simp add: complex_split imaginary_unit.ctr legacy_Complex_simps) from represents_2i have 12: "poly_2i represents (2 * \)" by simp have "\ qi \ set ?Imp. qi represents (inverse (2 * \) * (x - cnj x))" proof (cases "x - cnj x = 0") case False have "poly poly_2i 0 \ 0" unfolding poly_2i_def by auto from represents_div[OF appi 12 this] represents_irr_non_0[OF irr_pi appi False, unfolded poly_0_coeff_0] pi show ?thesis unfolding root_poly_Im_def Let_def by (auto intro: bexI[of _ "cf_pos_poly (poly_div pi poly_2i)"]) next case True hence id2: "Im x = 0" by (simp add: complex_eq_iff) from appi[unfolded True represents_def] have "coeff pi 0 = 0" by (cases pi, auto) with pi have mem: "[:0,1:] \ set ?Imp" unfolding root_poly_Im_def Let_def by auto have "[:0,1:] represents (complex_of_real (Im x))" unfolding id2 represents_def by simp with mem show ?thesis unfolding id by auto qed then obtain qi where qi: "qi \ set ?Imp" "qi \ 0" and rt: "ipoly qi (complex_of_real (Im x)) = 0" unfolding id represents_def by auto from qi rt[unfolded poly_complex_to_real] show "\ qi \ set ?Imp. qi represents (Im x)" by auto qed text \Determine complex roots of a polynomial, intended for polynomials of degree 3 or higher, for lower degree polynomials use @{const roots1} or @{const croots2}\ hide_const (open) eq primrec remdups_gen :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "remdups_gen eq [] = []" | "remdups_gen eq (x # xs) = (if (\ y \ set xs. eq x y) then remdups_gen eq xs else x # remdups_gen eq xs)" lemma real_of_3_remdups_equal_3[simp]: "real_of_3 ` set (remdups_gen equal_3 xs) = real_of_3 ` set xs" by (induct xs, auto simp: equal_3) lemma distinct_remdups_equal_3: "distinct (map real_of_3 (remdups_gen equal_3 xs))" by (induct xs, auto, auto simp: equal_3) lemma real_of_3_code [code]: "real_of_3 x = real_of (Real_Alg_Quotient x)" by (transfer, auto) definition "real_parts_3 p = roots_of_3 (root_poly_Re p)" definition "pos_imaginary_parts_3 p = remdups_gen equal_3 (filter (\ x. sgn_3 x = 1) (concat (map roots_of_3 (root_poly_Im p))))" lemma real_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" shows "Re x \ real_of_3 ` set (real_parts_3 p)" unfolding real_parts_3_def using represents_root_poly(1)[OF assms(2,1)] roots_of_3(1) unfolding represents_def by auto lemma distinct_real_parts_3: "distinct (map real_of_3 (real_parts_3 p))" unfolding real_parts_3_def using roots_of_3(2) . lemma pos_imaginary_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" and "Im x > 0" shows "Im x \ real_of_3 ` set (pos_imaginary_parts_3 p)" proof - from represents_root_poly(2)[OF assms(2,1)] obtain q where q: "q \ set (root_poly_Im p)" "q represents Im x" by auto from roots_of_3(1)[of q] have "Im x \ real_of_3 ` set (roots_of_3 q)" using q unfolding represents_def by auto then obtain i3 where i3: "i3 \ set (roots_of_3 q)" and id: "Im x = real_of_3 i3" by auto from \Im x > 0\ have "sgn (Im x) = 1" by simp hence sgn: "sgn_3 i3 = 1" unfolding id by (metis of_rat_eq_1_iff sgn_3) show ?thesis unfolding pos_imaginary_parts_3_def real_of_3_remdups_equal_3 id using sgn i3 q(1) by auto qed lemma distinct_pos_imaginary_parts_3: "distinct (map real_of_3 (pos_imaginary_parts_3 p))" unfolding pos_imaginary_parts_3_def by (rule distinct_remdups_equal_3) lemma remdups_gen_subset: "set (remdups_gen eq xs) \ set xs" by (induct xs, auto) lemma positive_pos_imaginary_parts_3: assumes "x \ set (pos_imaginary_parts_3 p)" shows "0 < real_of_3 x" proof - from subsetD[OF remdups_gen_subset assms[unfolded pos_imaginary_parts_3_def]] have "sgn_3 x = 1" by auto thus ?thesis using sgn_3[of x] by (simp add: sgn_1_pos) qed definition "pair_to_complex ri \ case ri of (r,i) \ Complex (real_of_3 r) (real_of_3 i)" fun get_itvl_2 :: "real_alg_2 \ real interval" where "get_itvl_2 (Irrational n (p,l,r)) = Interval (of_rat l) (of_rat r)" | "get_itvl_2 (Rational r) = (let rr = of_rat r in Interval rr rr)" lemma get_bounds_2: assumes "invariant_2 x" shows "real_of_2 x \\<^sub>i get_itvl_2 x" proof (cases x) case (Irrational n plr) with assms obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) from assms Irrational plr have inv1: "invariant_1 (p,l,r)" and id: "real_of_2 x = real_of_1 (p,l,r)" by auto show ?thesis unfolding id using invariant_1D(1)[OF inv1] by (auto simp: plr Irrational) qed (insert assms, auto simp: Let_def) lift_definition get_itvl_3 :: "real_alg_3 \ real interval" is get_itvl_2 . lemma get_itvl_3: "real_of_3 x \\<^sub>i get_itvl_3 x" by (transfer, insert get_bounds_2, auto) fun tighten_bounds_2 :: "real_alg_2 \ real_alg_2" where "tighten_bounds_2 (Irrational n (p,l,r)) = (case tighten_poly_bounds p l r (sgn (ipoly p r)) of (l',r',_) \ Irrational n (p,l',r'))" | "tighten_bounds_2 (Rational r) = Rational r" lemma tighten_bounds_2: assumes inv: "invariant_2 x" shows "real_of_2 (tighten_bounds_2 x) = real_of_2 x" "invariant_2 (tighten_bounds_2 x)" "get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r-l) / 2" proof (atomize(full), cases x) case (Irrational n plr) show "real_of_2 (tighten_bounds_2 x) = real_of_2 x \ invariant_2 (tighten_bounds_2 x) \ (get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r - l) / 2)" proof - obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) let ?tb = "tighten_poly_bounds p l r (sgn (ipoly p r))" obtain l' r' sr' where tb: "?tb = (l',r',sr')" by (cases ?tb, auto) have id: "tighten_bounds_2 x = Irrational n (p,l',r')" unfolding Irrational plr using tb by auto from inv[unfolded Irrational plr] have inv: "invariant_1_2 (p, l, r)" "n = card {y. y \ real_of_1 (p, l, r) \ ipoly p y = 0}" by auto have rof: "real_of_2 x = real_of_1 (p, l, r)" "real_of_2 (tighten_bounds_2 x) = real_of_1 (p, l', r')" using Irrational plr id by auto from inv have inv1: "invariant_1 (p, l, r)" and "poly_cond2 p" by auto hence rc: "\!x. root_cond (p, l, r) x" "poly_cond2 p" by auto note tb' = tighten_poly_bounds[OF tb rc refl] have eq: "real_of_1 (p, l, r) = real_of_1 (p, l', r')" using tb' inv1 using invariant_1_sub_interval(2) by presburger from inv1 tb' have "invariant_1 (p, l', r')" by (metis invariant_1_sub_interval(1)) hence inv2: "invariant_2 (tighten_bounds_2 x)" unfolding id using inv eq by auto thus ?thesis unfolding rof eq unfolding id unfolding Irrational plr using tb'(1-4) arg_cong[OF tb'(5), of real_of_rat] by (auto simp: hom_distribs) qed qed (auto simp: Let_def) lift_definition tighten_bounds_3 :: "real_alg_3 \ real_alg_3" is tighten_bounds_2 using tighten_bounds_2 by auto lemma tighten_bounds_3: "real_of_3 (tighten_bounds_3 x) = real_of_3 x" "get_itvl_3 x = Interval l r \ get_itvl_3 (tighten_bounds_3 x) = Interval l' r' \ r' - l' = (r-l) / 2" by (transfer, insert tighten_bounds_2, auto)+ partial_function (tailrec) filter_list_length :: "('a \ 'a) \ ('a \ bool) \ nat \ 'a list \ 'a list" where [code]: "filter_list_length f p n xs = (let ys = filter p xs in if length ys = n then ys else filter_list_length f p n (map f ys))" lemma filter_list_length: assumes "length (filter P xs) = n" and "\ i x. x \ set xs \ P x \ p ((f ^^ i) x)" and "\ x. x \ set xs \ \ P x \ \ i. \ p ((f ^^ i) x)" and g: "\ x. g (f x) = g x" and P: "\ x. P (f x) = P x" shows "map g (filter_list_length f p n xs) = map g (filter P xs)" proof - from assms(3) have "\ x. \ i. x \ set xs \ \ P x \ \ p ((f ^^ i) x)" by auto from choice[OF this] obtain i where i: "\ x. x \ set xs \ \ P x \ \ p ((f ^^ (i x)) x)" by auto define m where "m = max_list (map i xs)" have m: "\ x. x \ set xs \ \ P x \ \ i \ m. \ p ((f ^^ i) x)" using max_list[of _ "map i xs", folded m_def] i by auto show ?thesis using assms(1-2) m proof (induct m arbitrary: xs rule: less_induct) case (less m xs) define ys where "ys = filter p xs" have xs_ys: "filter P xs = filter P ys" unfolding ys_def filter_filter by (rule filter_cong[OF refl], insert less(3)[of _ 0], auto) have "filter (P \ f) ys = filter P ys" using P unfolding o_def by auto hence id3: "filter P (map f ys) = map f (filter P ys)" unfolding filter_map by simp hence id2: "map g (filter P (map f ys)) = map g (filter P ys)" by (simp add: g) show ?case proof (cases "length ys = n") case True hence id: "filter_list_length f p n xs = ys" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis using True unfolding id xs_ys using less(2) by (metis filter_id_conv length_filter_less less_le xs_ys) next case False { assume "m = 0" from less(4)[unfolded this] have Pp: "x \ set xs \ \ P x \ \ p x" for x by auto with xs_ys False[folded less(2)] have False by (metis (mono_tags, lifting) filter_True mem_Collect_eq set_filter ys_def) } note m0 = this then obtain M where mM: "m = Suc M" by (cases m, auto) hence m: "M < m" by simp from False have id: "filter_list_length f p n xs = filter_list_length f p n (map f ys)" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis unfolding id xs_ys id2[symmetric] proof (rule less(1)[OF m]) fix y assume "y \ set (map f ys)" then obtain x where x: "x \ set xs" "p x" and y: "y = f x" unfolding ys_def by auto { assume "\ P y" hence "\ P x" unfolding y P . from less(4)[OF x(1) this] obtain i where i: "i \ m" and p: "\ p ((f ^^ i) x)" by auto with x obtain j where ij: "i = Suc j" by (cases i, auto) with i have j: "j \ M" unfolding mM by auto have "\ p ((f ^^ j) y)" using p unfolding ij y funpow_Suc_right by simp thus "\i\ M. \ p ((f ^^ i) y)" using j by auto } { fix i assume "P y" hence "P x" unfolding y P . from less(3)[OF x(1) this, of "Suc i"] show "p ((f ^^ i) y)" unfolding y funpow_Suc_right by simp } next show "length (filter P (map f ys)) = n" unfolding id3 length_map using xs_ys less(2) by auto qed qed qed qed definition complex_roots_of_int_poly3 :: "int poly \ complex list" where "complex_roots_of_int_poly3 p \ let n = degree p; rrts = real_roots_of_int_poly p; nr = length rrts; crts = map (\ r. Complex r 0) rrts in if n = nr then crts else let nr_crts = n - nr in if nr_crts = 2 then let pp = real_of_int_poly p div (prod_list (map (\ x. [:-x,1:]) rrts)); cpp = map_poly (\ r. Complex r 0) pp in crts @ croots2 cpp else let nr_pos_crts = nr_crts div 2; rxs = real_parts_3 p; ixs = pos_imaginary_parts_3 p; rts = [(rx, ix). rx <- rxs, ix <- ixs]; crts' = map pair_to_complex (filter_list_length (map_prod tighten_bounds_3 tighten_bounds_3) (\ (r, i). 0 \\<^sub>c ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))) nr_pos_crts rts) in crts @ (concat (map (\ x. [x, cnj x]) crts'))" definition complex_roots_of_int_poly_all :: "int poly \ complex list" where "complex_roots_of_int_poly_all p = (let n = degree p in if n \ 3 then complex_roots_of_int_poly3 p else if n = 1 then [roots1 (map_poly of_int p)] else if n = 2 then croots2 (map_poly of_int p) else [])" lemma in_real_itvl_get_bounds_tighten: "real_of_3 x \\<^sub>i get_itvl_3 ((tighten_bounds_3 ^^ n) x)" proof (induct n arbitrary: x) case 0 thus ?case using get_itvl_3[of x] by simp next case (Suc n x) have id: "(tighten_bounds_3 ^^ (Suc n)) x = (tighten_bounds_3 ^^ n) (tighten_bounds_3 x)" by (metis comp_apply funpow_Suc_right) show ?case unfolding id tighten_bounds_3(1)[of x, symmetric] by (rule Suc) qed lemma sandwitch_real: fixes l r :: "nat \ real" assumes la: "l \ a" and ra: "r \ a" and lm: "\i. l i \ m i" and mr: "\i. m i \ r i" shows "m \ a" proof (rule LIMSEQ_I) fix e :: real assume "0 < e" hence e: "0 < e / 2" by simp from LIMSEQ_D[OF la e] obtain n1 where n1: "\ n. n \ n1 \ norm (l n - a) < e/2" by auto from LIMSEQ_D[OF ra e] obtain n2 where n2: "\ n. n \ n2 \ norm (r n - a) < e/2" by auto show "\no. \n\no. norm (m n - a) < e" proof (rule exI[of _ "max n1 n2"], intro allI impI) fix n assume "max n1 n2 \ n" with n1 n2 have *: "norm (l n - a) < e/2" "norm (r n - a) < e/2" by auto from lm[of n] mr[of n] have "norm (m n - a) \ norm (l n - a) + norm (r n - a)" by simp with * show "norm (m n - a) < e" by auto qed qed lemma real_of_tighten_bounds_many[simp]: "real_of_3 ((tighten_bounds_3 ^^ i) x) = real_of_3 x" apply (induct i) using tighten_bounds_3 by auto definition lower_3 where "lower_3 x i \ interval.lower (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" definition upper_3 where "upper_3 x i \ interval.upper (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" lemma interval_size_3: "upper_3 x i - lower_3 x i = (upper_3 x 0 - lower_3 x 0)/2^i" proof (induct i) case (Suc i) have "upper_3 x (Suc i) - lower_3 x (Suc i) = (upper_3 x i - lower_3 x i) / 2" unfolding upper_3_def lower_3_def using tighten_bounds_3 get_itvl_3 by auto with Suc show ?case by auto qed auto lemma interval_size_3_tendsto_0: "(\i. (upper_3 x i - lower_3 x i)) \ 0" by (subst interval_size_3, auto intro: LIMSEQ_divide_realpow_zero) lemma dist_tendsto_0_imp_tendsto: "(\i. \f i - a\ :: real) \ 0 \ f \ a" using LIM_zero_cancel tendsto_rabs_zero_iff by blast lemma upper_3_tendsto: "upper_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\(upper_3 x) i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma lower_3_tendsto: "lower_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\lower_3 x i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma tends_to_tight_bounds_3: "(\x. get_itvl_3 ((tighten_bounds_3 ^^ x) y)) \\<^sub>i real_of_3 y" using lower_3_tendsto[of y] upper_3_tendsto[of y] unfolding lower_3_def upper_3_def interval_tendsto_def o_def by auto lemma complex_roots_of_int_poly3: assumes p: "p \ 0" and sf: "square_free p" shows "set (complex_roots_of_int_poly3 p) = {x. ipoly p x = 0}" (is "?l = ?r") "distinct (complex_roots_of_int_poly3 p)" proof - interpret map_poly_inj_idom_hom of_real.. define q where "q = real_of_int_poly p" let ?q = "map_poly complex_of_real q" from p have q0: "q \ 0" unfolding q_def by auto hence q: "?q \ 0" by auto define rr where "rr = real_roots_of_int_poly p" define rrts where "rrts = map (\r. Complex r 0) rr" note d = complex_roots_of_int_poly3_def[of p, unfolded Let_def, folded rr_def, folded rrts_def] have rr: "set rr = {x. ipoly p x = 0}" unfolding rr_def using real_roots_of_int_poly(1)[OF p] . have rrts: "set rrts = {x. poly ?q x = 0 \ x \ \}" unfolding rrts_def set_map rr q_def complex_of_real_def[symmetric] by (auto elim: Reals_cases) have dist: "distinct rr" unfolding rr_def using real_roots_of_int_poly(2) . from dist have dist1: "distinct rrts" unfolding rrts_def distinct_map inj_on_def by auto have lrr: "length rr = card {x. poly (real_of_int_poly p) x = 0}" unfolding rr_def using real_roots_of_int_poly[of p] p distinct_card by fastforce have cr: "length rr = card {x. poly ?q x = 0 \ x \ \}" unfolding lrr q_def[symmetric] proof - have "card {x. poly q x = 0} \ card {x. poly (map_poly complex_of_real q) x = 0 \ x \ \}" (is "?l \ ?r") by (rule card_inj_on_le[of of_real], insert poly_roots_finite[OF q], auto simp: inj_on_def) moreover have "?l \ ?r" by (rule card_inj_on_le[of Re, OF _ _ poly_roots_finite[OF q0]], auto simp: inj_on_def elim!: Reals_cases) ultimately show "?l = ?r" by simp qed have conv: "\ x. ipoly p x = 0 \ poly ?q x = 0" unfolding q_def by (subst map_poly_map_poly, auto simp: o_def) have r: "?r = {x. poly ?q x = 0}" unfolding conv .. have "?l = {x. ipoly p x = 0} \ distinct (complex_roots_of_int_poly3 p)" proof (cases "degree p = length rr") case False note oFalse = this show ?thesis proof (cases "degree p - length rr = 2") case False let ?nr = "(degree p - length rr) div 2" define cpxI where "cpxI = pos_imaginary_parts_3 p" define cpxR where "cpxR = real_parts_3 p" let ?rts = "[(rx,ix). rx <- cpxR, ix <- cpxI]" define cpx where "cpx = map pair_to_complex (filter (\ c. ipoly p (pair_to_complex c) = 0) ?rts)" let ?LL = "cpx @ map cnj cpx" let ?LL' = "concat (map (\ x. [x,cnj x]) cpx)" let ?ll = "rrts @ ?LL" let ?ll' = "rrts @ ?LL'" have cpx: "set cpx \ ?r" unfolding cpx_def by auto have ccpx: "cnj ` set cpx \ ?r" using cpx unfolding r by (auto intro!: complex_conjugate_root[of ?q] simp: Reals_def) have "set ?ll \ ?r" using rrts cpx ccpx unfolding r by auto moreover { fix x :: complex assume rt: "ipoly p x = 0" { fix x assume rt: "ipoly p x = 0" and gt: "Im x > 0" define rx where "rx = Re x" let ?x = "Complex rx (Im x)" have x: "x = ?x" by (cases x, auto simp: rx_def) from rt x have rt': "ipoly p ?x = 0" by auto from real_parts_3[OF p rt, folded rx_def] pos_imaginary_parts_3[OF p rt gt] rt' have "?x \ set cpx" unfolding cpx_def cpxI_def cpxR_def by (force simp: pair_to_complex_def[abs_def]) hence "x \ set cpx" using x by simp } note gt = this have cases: "Im x = 0 \ Im x > 0 \ Im x < 0" by auto from rt have rt': "ipoly p (cnj x) = 0" unfolding conv by (intro complex_conjugate_root[of ?q x], auto simp: Reals_def) { assume "Im x > 0" from gt[OF rt this] have "x \ set ?ll" by auto } moreover { assume "Im x < 0" hence "Im (cnj x) > 0" by simp from gt[OF rt' this] have "cnj (cnj x) \ set ?ll" unfolding set_append set_map by blast hence "x \ set ?ll" by simp } moreover { assume "Im x = 0" hence "x \ \" using complex_is_Real_iff by blast with rt rrts have "x \ set ?ll" unfolding conv by auto } ultimately have "x \ set ?ll" using cases by blast } ultimately have lr: "set ?ll = {x. ipoly p x = 0}" by blast let ?rr = "map real_of_3 cpxR" let ?pi = "map real_of_3 cpxI" have dist2: "distinct ?rr" unfolding cpxR_def by (rule distinct_real_parts_3) have dist3: "distinct ?pi" unfolding cpxI_def by (rule distinct_pos_imaginary_parts_3) have idd: "concat (map (map pair_to_complex) (map (\rx. map (Pair rx) cpxI) cpxR)) = concat (map (\r. map (\ i. Complex (real_of_3 r) (real_of_3 i)) cpxI) cpxR)" unfolding pair_to_complex_def by (auto simp: o_def) have dist4: "distinct cpx" unfolding cpx_def proof (rule distinct_map_filter, unfold map_concat idd, unfold distinct_conv_nth, intro allI impI, goal_cases) case (1 i j) from nth_concat_diff[OF 1, unfolded length_map] dist2[unfolded distinct_conv_nth] dist3[unfolded distinct_conv_nth] show ?case by auto qed have dist5: "distinct (map cnj cpx)" using dist4 unfolding distinct_map by (auto simp: inj_on_def) { fix x :: complex have rrts: "x \ set rrts \ Im x = 0" unfolding rrts_def by auto have cpx: "\ x. x \ set cpx \ Im x > 0" unfolding cpx_def cpxI_def by (auto simp: pair_to_complex_def[abs_def] positive_pos_imaginary_parts_3) have cpx': "x \ cnj ` set cpx \ sgn (Im x) = -1" using cpx by auto have "x \ set rrts \ set cpx \ set rrts \ cnj ` set cpx \ set cpx \ cnj ` set cpx" using rrts cpx[of x] cpx' by auto } note dist6 = this have dist: "distinct ?ll" unfolding distinct_append using dist6 by (auto simp: dist1 dist4 dist5) let ?p = "complex_of_int_poly p" have pp: "?p \ 0" using p by auto from p square_free_of_int_poly[OF sf] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto from dist lr have "length ?ll = card {x. poly ?p x = 0}" by (metis distinct_card) also have "\ = degree p" using rsf unfolding rsquarefree_card_degree[OF pp] by simp finally have deg_len: "degree p = length ?ll" by simp let ?P = "\ c. ipoly p (pair_to_complex c) = 0" let ?itvl = "\ r i. ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))" let ?itv = "\ (r,i). ?itvl r i" let ?p = "(\ (r,i). 0 \\<^sub>c (?itvl r i))" let ?tb = tighten_bounds_3 let ?f = "map_prod ?tb ?tb" have filter: "map pair_to_complex (filter_list_length ?f ?p ?nr ?rts) = map pair_to_complex (filter ?P ?rts)" proof (rule filter_list_length) have "length (filter ?P ?rts) = length cpx" unfolding cpx_def by simp also have "\ = ?nr" unfolding deg_len by (simp add: rrts_def) finally show "length (filter ?P ?rts) = ?nr" by auto next fix n x assume x: "?P x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" unfolding xri by (induct n, auto) have px: "pair_to_complex x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto show "?p ((?f ^^ n) x)" unfolding id split by (rule ipoly_complex_interval[of "pair_to_complex x" _ p, unfolded x], unfold px, auto simp: in_complex_interval_def in_real_itvl_get_bounds_tighten) next fix x assume x: "x \ set ?rts" "\ ?P x" let ?x = "pair_to_complex x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" for n unfolding xri by (induct n, auto) have px: "?x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto have cvg: "(\ n. ?itv ((?f ^^ n) x)) \\<^sub>c ipoly p ?x" unfolding id split px proof (rule ipoly_complex_interval_tendsto) show "(\ia. Complex_Interval (get_itvl_3 ((?tb ^^ ia) r)) (get_itvl_3 ((?tb ^^ ia) i))) \\<^sub>c Complex (real_of_3 r) (real_of_3 i)" unfolding complex_interval_tendsto_def by (simp add: tends_to_tight_bounds_3 o_def) qed from complex_interval_tendsto_neq[OF this x(2)] show "\ i. \ ?p ((?f ^^ i) x)" unfolding id by auto next show "pair_to_complex (?f x) = pair_to_complex x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) next show "?P (?f x) = ?P x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) qed have l: "complex_roots_of_int_poly3 p = ?ll'" unfolding d filter cpx_def[symmetric] cpxI_def[symmetric] cpxR_def[symmetric] using False oFalse by auto have "distinct ?ll' = (distinct rrts \ distinct ?LL' \ set rrts \ set ?LL' = {})" unfolding distinct_append .. also have "set ?LL' = set ?LL" by auto also have "distinct ?LL' = distinct ?LL" by (induct cpx, auto) finally have "distinct ?ll' = distinct ?ll" unfolding distinct_append by auto with dist have "distinct ?ll'" by auto with lr l show ?thesis by auto next case True let ?cr = "map_poly of_real :: real poly \ complex poly" define pp where "pp = complex_of_int_poly p" have id: "pp = map_poly of_real q" unfolding q_def pp_def by (subst map_poly_map_poly, auto simp: o_def) let ?rts = "map (\ x. [:-x,1:]) rr" define rts where "rts = prod_list ?rts" let ?c2 = "?cr (q div rts)" have pq: "\ x. ipoly p x = 0 \ poly q x = 0" unfolding q_def by simp from True have 2: "degree q - card {x. poly q x = 0} = 2" unfolding pq[symmetric] lrr unfolding q_def by simp from True have id: "degree p = length rr \ False" "degree p - length rr = 2 \ True" by auto have l: "?l = of_real ` {x. poly q x = 0} \ set (croots2 ?c2)" unfolding d rts_def id if_False if_True set_append rrts Reals_def by (fold complex_of_real_def q_def, auto) from dist have len_rr: "length rr = card {x. poly q x = 0}" unfolding rr[unfolded pq, symmetric] by (simp add: distinct_card) have rr': "\ r. r \ set rr \ poly q r = 0" using rr unfolding q_def by simp with dist have "q = q div prod_list ?rts * prod_list ?rts" proof (induct rr arbitrary: q) case (Cons r rr q) note dist = Cons(2) let ?p = "q div [:-r,1:]" from Cons.prems(2) have "poly q r = 0" by simp hence "[:-r,1:] dvd q" using poly_eq_0_iff_dvd by blast from dvd_mult_div_cancel[OF this] have "q = ?p * [:-r,1:]" by simp moreover have "?p = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:])" proof (rule Cons.hyps) show "distinct rr" using dist by auto fix s assume "s \ set rr" with dist Cons(3) have "s \ r" "poly q s = 0" by auto hence "poly (?p * [:- 1 * r, 1:]) s = 0" using calculation by force thus "poly ?p s = 0" by (simp add: \s \ r\) qed ultimately have q: "q = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:]) * [:-r,1:]" by auto also have "\ = (?p div (\x\rr. [:- x, 1:])) * (\x\r # rr. [:- x, 1:])" unfolding mult.assoc by simp also have "?p div (\x\rr. [:- x, 1:]) = q div (\x\r # rr. [:- x, 1:])" unfolding poly_div_mult_right[symmetric] by simp finally show ?case . qed simp hence q_div: "q = q div rts * rts" unfolding rts_def . from q_div q0 have "q div rts \ 0" "rts \ 0" by auto from degree_mult_eq[OF this] have "degree q = degree (q div rts) + degree rts" using q_div by simp also have "degree rts = length rr" unfolding rts_def by (rule degree_linear_factors) also have "\ = card {x. poly q x = 0}" unfolding len_rr by simp finally have deg2: "degree ?c2 = 2" using 2 by simp note croots2 = croots2[OF deg2, symmetric] have "?q = ?cr (q div rts * rts)" using q_div by simp also have "\ = ?cr rts * ?c2" unfolding hom_distribs by simp finally have q_prod: "?q = ?cr rts * ?c2" . from croots2 l have l: "?l = of_real ` {x. poly q x = 0} \ {x. poly ?c2 x = 0}" by simp from r[unfolded q_prod] have r: "?r = {x. poly (?cr rts) x = 0} \ {x. poly ?c2 x = 0}" by auto also have "?cr rts = (\x\rr. ?cr [:- x, 1:])" by (simp add: rts_def o_def of_real_poly_hom.hom_prod_list) also have "{x. poly \ x = 0} = of_real ` set rr" unfolding poly_prod_list_zero_iff by auto also have "set rr = {x. poly q x = 0}" unfolding rr q_def by simp finally have lr: "?l = ?r" unfolding l by simp show ?thesis proof (intro conjI[OF lr]) from sf have sf: "square_free q" unfolding q_def by (rule square_free_of_int_poly) { interpret field_hom_0' complex_of_real .. from sf have "square_free ?q" unfolding square_free_map_poly . } note sf = this have l: "complex_roots_of_int_poly3 p = rrts @ croots2 ?c2" unfolding d rts_def id if_False if_True set_append rrts q_def complex_of_real_def by auto have dist2: "distinct (croots2 ?c2)" unfolding croots2_def Let_def by auto { fix x assume x: "x \ set (croots2 ?c2)" "x \ set rrts" from x(1)[unfolded croots2] have x1: "poly ?c2 x = 0" by auto from x(2) have x2: "poly (?cr rts) x = 0" unfolding rrts_def rts_def complex_of_real_def[symmetric] by (auto simp: poly_prod_list_zero_iff o_def) from square_free_multD(1)[OF sf[unfolded q_prod], of "[: -x, 1:]"] x1 x2 have False unfolding poly_eq_0_iff_dvd by auto } note dist3 = this show "distinct (complex_roots_of_int_poly3 p)" unfolding l distinct_append by (intro conjI dist1 dist2, insert dist3, auto) qed qed next case True have "card {x. poly ?q x = 0} \ degree ?q" by (rule poly_roots_degree[OF q]) also have "\ = degree p" unfolding q_def by simp also have "\ = card {x. poly ?q x = 0 \ x \ \}" using True cr by simp finally have le: "card {x. poly ?q x = 0} \ card {x. poly ?q x = 0 \ x \ \}" by auto have "{x. poly ?q x = 0 \ x \ \} = {x. poly ?q x = 0}" by (rule card_seteq[OF _ _ le], insert poly_roots_finite[OF q], auto) with True rrts dist1 show ?thesis unfolding r d by auto qed thus "distinct (complex_roots_of_int_poly3 p)" "?l = ?r" by auto qed lemma complex_roots_of_int_poly_all: assumes sf: "degree p \ 3 \ square_free p" shows "p \ 0 \ set (complex_roots_of_int_poly_all p) = {x. ipoly p x = 0}" (is "_ \ set ?l = ?r") and "distinct (complex_roots_of_int_poly_all p)" proof - note d = complex_roots_of_int_poly_all_def Let_def have "(p \ 0 \ set ?l = ?r) \ (distinct (complex_roots_of_int_poly_all p))" proof (cases "degree p \ 3") case True hence p: "p \ 0" by auto from True complex_roots_of_int_poly3[OF p] sf show ?thesis unfolding d by auto next case False let ?p = "map_poly (of_int :: int \ complex) p" have deg: "degree ?p = degree p" by (simp add: degree_map_poly) show ?thesis proof (cases "degree p = 1") case True hence l: "?l = [roots1 ?p]" unfolding d by auto from True have "degree ?p = 1" unfolding deg by auto from roots1[OF this] show ?thesis unfolding l roots1_def by auto next case False show ?thesis proof (cases "degree p = 2") case True hence l: "?l = croots2 ?p" unfolding d by auto from True have "degree ?p = 2" unfolding deg by auto from croots2[OF this] show ?thesis unfolding l by (simp add: croots2_def Let_def) next case False with \degree p \ 1\ \degree p \ 2\ \\ (degree p \ 3)\ have True: "degree p = 0" by auto hence l: "?l = []" unfolding d by auto from True have "degree ?p = 0" unfolding deg by auto from roots0[OF _ this] show ?thesis unfolding l by simp qed qed qed thus "p \ 0 \ set ?l = ?r" "distinct (complex_roots_of_int_poly_all p)" by auto qed text \It now comes the preferred function to compute complex roots of a integer polynomial.\ definition complex_roots_of_int_poly :: "int poly \ complex list" where "complex_roots_of_int_poly p = ( let ps = (if degree p \ 3 then factors_of_int_poly p else [p]) in concat (map complex_roots_of_int_poly_all ps))" definition complex_roots_of_rat_poly :: "rat poly \ complex list" where "complex_roots_of_rat_poly p = complex_roots_of_int_poly (snd (rat_to_int_poly p))" lemma complex_roots_of_int_poly: shows "p \ 0 \ set (complex_roots_of_int_poly p) = {x. ipoly p x = 0}" (is "_ \ ?l = ?r") and "distinct (complex_roots_of_int_poly p)" proof - have "(p \ 0 \ ?l = ?r) \ (distinct (complex_roots_of_int_poly p))" proof (cases "degree p \ 3") case False hence "complex_roots_of_int_poly p = complex_roots_of_int_poly_all p" unfolding complex_roots_of_int_poly_def Let_def by auto with complex_roots_of_int_poly_all[of p] False show ?thesis by auto next case True { fix q assume "q \ set (factors_of_int_poly p)" from factors_of_int_poly(1)[OF refl this] irreducible_imp_square_free[of q] have 0: "q \ 0" and sf: "square_free q" by auto from complex_roots_of_int_poly_all(1)[OF sf 0] complex_roots_of_int_poly_all(2)[OF sf] have "set (complex_roots_of_int_poly_all q) = {x. ipoly q x = 0}" "distinct (complex_roots_of_int_poly_all q)" by auto } note all = this from True have "?l = (\ ((\ p. set (complex_roots_of_int_poly_all p)) ` set (factors_of_int_poly p)))" unfolding complex_roots_of_int_poly_def Let_def by auto also have "\ = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" using all by blast finally have l: "?l = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" . have lr: "p \ 0 \ ?l = ?r" using l factors_of_int_poly(2)[OF refl, of p] by auto show ?thesis proof (rule conjI[OF lr]) from True have id: "complex_roots_of_int_poly p = concat (map complex_roots_of_int_poly_all (factors_of_int_poly p))" unfolding complex_roots_of_int_poly_def Let_def by auto show "distinct (complex_roots_of_int_poly p)" unfolding id distinct_conv_nth proof (intro allI impI, goal_cases) case (1 i j) let ?fp = "factors_of_int_poly p" let ?rr = "complex_roots_of_int_poly_all" let ?cc = "concat (map ?rr (factors_of_int_poly p))" from nth_concat_diff[OF 1, unfolded length_map] obtain j1 k1 j2 k2 where *: "(j1,k1) \ (j2,k2)" "j1 < length ?fp" "j2 < length ?fp" and "k1 < length (map ?rr ?fp ! j1)" "k2 < length (map ?rr ?fp ! j2)" "?cc ! i = map ?rr ?fp ! j1 ! k1" "?cc ! j = map ?rr ?fp ! j2 ! k2" by blast hence **: "k1 < length (?rr (?fp ! j1))" "k2 < length (?rr (?fp ! j2))" "?cc ! i = ?rr (?fp ! j1) ! k1" "?cc ! j = ?rr (?fp ! j2) ! k2" by auto from * have mem: "?fp ! j1 \ set ?fp" "?fp ! j2 \ set ?fp" by auto show "?cc ! i \ ?cc ! j" proof (cases "j1 = j2") case True with * have "k1 \ k2" by auto with all(2)[OF mem(2)] **(1-2) show ?thesis unfolding **(3-4) unfolding True distinct_conv_nth by auto next case False from \degree p \ 3\ have p: "p \ 0" by auto note fip = factors_of_int_poly(2-3)[OF refl this] show ?thesis unfolding **(3-4) proof define x where "x = ?rr (?fp ! j2) ! k2" assume id: "?rr (?fp ! j1) ! k1 = ?rr (?fp ! j2) ! k2" from ** have x1: "x \ set (?rr (?fp ! j1))" unfolding x_def id[symmetric] by auto from ** have x2: "x \ set (?rr (?fp ! j2))" unfolding x_def by auto from all(1)[OF mem(1)] x1 have x1: "ipoly (?fp ! j1) x = 0" by auto from all(1)[OF mem(2)] x2 have x2: "ipoly (?fp ! j2) x = 0" by auto from False factors_of_int_poly(4)[OF refl, of p] have neq: "?fp ! j1 \ ?fp ! j2" using * unfolding distinct_conv_nth by auto have "poly (complex_of_int_poly p) x = 0" by (meson fip(1) mem(2) x2) from fip(2)[OF this] mem x1 x2 neq show False by blast qed qed qed qed qed thus "p \ 0 \ ?l = ?r" "distinct (complex_roots_of_int_poly p)" by auto qed lemma complex_roots_of_rat_poly: "p \ 0 \ set (complex_roots_of_rat_poly p) = {x. rpoly p x = 0}" (is "_ \ ?l = ?r") "distinct (complex_roots_of_rat_poly p)" proof - obtain c q where cq: "rat_to_int_poly p = (c,q)" by force from rat_to_int_poly[OF this] have pq: "p = smult (inverse (of_int c)) (of_int_poly q)" and c: "c \ 0" by auto show "distinct (complex_roots_of_rat_poly p)" unfolding complex_roots_of_rat_poly_def using complex_roots_of_int_poly(2) . assume p: "p \ 0" with pq c have q: "q \ 0" by auto have id: "{x. rpoly p x = (0 :: complex)} = {x. ipoly q x = 0}" unfolding pq by (simp add: c of_rat_of_int_poly hom_distribs) show "?l = ?r" unfolding complex_roots_of_rat_poly_def cq snd_conv id complex_roots_of_int_poly(1)[OF q] .. qed definition roots_of_complex_main :: "complex poly \ complex list" where "roots_of_complex_main p \ let n = degree p in if n = 0 then [] else if n = 1 then [roots1 p] else if n = 2 then croots2 p else (complex_roots_of_rat_poly (map_poly to_rat p))" definition roots_of_complex_poly :: "complex poly \ complex list option" where "roots_of_complex_poly p \ let (c,pis) = yun_factorization gcd p in if (c \ 0 \ (\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then Some (concat (map (roots_of_complex_main o fst) pis)) else None" lemma roots_of_complex_main: assumes p: "p \ 0" and deg: "degree p \ 2 \ set (coeffs p) \ \" shows "set (roots_of_complex_main p) = {x. poly p x = 0}" (is "set ?l = ?r") and "distinct (roots_of_complex_main p)" proof - note d = roots_of_complex_main_def Let_def have "set ?l = ?r \ distinct (roots_of_complex_main p)" proof (cases "degree p = 0") case True hence "?l = []" unfolding d by auto with roots0[OF p True] show ?thesis by auto next case False note 0 = this show ?thesis proof (cases "degree p = 1") case True hence "?l = [roots1 p]" unfolding d by auto with roots1[OF True] show ?thesis by auto next case False note 1 = this show ?thesis proof (cases "degree p = 2") case True hence "?l = croots2 p" unfolding d by auto with croots2[OF True] show ?thesis by (auto simp: croots2_def Let_def) next case False note 2 = this let ?q = "map_poly to_rat p" from 0 1 2 have l: "?l = complex_roots_of_rat_poly ?q" unfolding d by auto from deg 0 1 2 have rat: "set (coeffs p) \ \" by auto have "p = map_poly (of_rat o to_rat) p" by (rule sym, rule map_poly_idI, insert rat, auto) also have "\ = complex_of_rat_poly ?q" by (subst map_poly_map_poly, auto simp: to_rat) finally have id: "{x. poly p x = 0} = {x. poly (complex_of_rat_poly ?q) x = 0}" and q: "?q \ 0" using p by auto from complex_roots_of_rat_poly[of ?q, folded id l] q show ?thesis by auto qed qed qed thus "set ?l = ?r" "distinct ?l" by auto qed lemma roots_of_complex_poly: assumes rt: "roots_of_complex_poly p = Some xs" shows "set xs = {x. poly p x = 0}" proof - obtain c pis where yun: "yun_factorization gcd p = (c,pis)" by force from rt[unfolded roots_of_complex_poly_def yun split Let_def] have c: "c \ 0" and pis: "\ p i. (p, i)\ set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" and xs: "xs = concat (map (roots_of_complex_main \ fst) pis)" by (auto split: if_splits) note yun = square_free_factorizationD(1,2,4)[OF yun_factorization(1)[OF yun]] from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . have "{x. poly p x = 0} = {x. poly (\(a, i)\set pis. a ^ Suc i) x = 0}" unfolding p using c by auto also have "\ = \ ((\ p. {x. poly p x = 0}) ` fst ` set pis)" (is "_ = ?r") by (subst poly_prod_0, force+) finally have r: "{x. poly p x = 0} = ?r" . { fix p i assume p: "(p,i) \ set pis" have "set (roots_of_complex_main p) = {x. poly p x = 0}" by (rule roots_of_complex_main, insert yun(2)[OF p] pis[OF p], auto) } note main = this have "set xs = \ ((\ (p, i). set (roots_of_complex_main p)) ` set pis)" unfolding xs o_def by auto also have "\ = ?r" using main by auto finally show ?thesis unfolding r by simp qed subsection \Factorization of Complex Polynomials\ definition factorize_complex_main :: "complex poly \ (complex \ (complex \ nat) list) option" where "factorize_complex_main p \ let (c,pis) = yun_factorization gcd p in if ((\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then Some (c, concat (map (\ (p,i). map (\ r. (r,i)) (roots_of_complex_main p)) pis)) else None" definition factorize_complex_poly :: "complex poly \ (complex \ (complex poly \ nat) list) option" where "factorize_complex_poly p \ map_option (\ (c,ris). (c, map (\ (r,i). ([:-r,1:],Suc i)) ris)) (factorize_complex_main p)" lemma factorize_complex_main: assumes rt: "factorize_complex_main p = Some (c,xis)" shows "p = smult c (\(x, i)\xis. [:- x, 1:] ^ Suc i)" proof - obtain d pis where yun: "yun_factorization gcd p = (d,pis)" by force from rt[unfolded factorize_complex_main_def yun split Let_def] have pis: "\ p i. (p, i)\set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" and xis: "xis = concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis)" and d: "d = c" by (auto split: if_splits) note yun = yun_factorization[OF yun[unfolded d]] note yun = square_free_factorizationD[OF yun(1)] yun(2)[unfolded snd_conv] let ?exp = "\ pis. \(x, i)\concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis). [:- x, 1:] ^ Suc i" from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . also have "(\(a, i)\set pis. a ^ Suc i) = (\(a, i)\pis. a ^ Suc i)" by (rule prod.distinct_set_conv_list[OF yun(5)]) also have "\ = ?exp pis" using pis yun(2,6) proof (induct pis) case (Cons pi pis) obtain p i where pi: "pi = (p,i)" by force let ?rts = "roots_of_complex_main p" note Cons = Cons[unfolded pi] have IH: "(\(a, i)\pis. a ^ Suc i) = (?exp pis)" by (rule Cons(1)[OF Cons(2-4)], auto) from Cons(2-4)[of p i] have deg: "degree p \ 2 \ (\x\set (coeffs p). x \ \)" and p: "square_free p" "degree p \ 0" "p \ 0" "monic p" by auto have "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * (\(a, i)\pis. a ^ Suc i)" unfolding pi by simp also have "(\(a, i)\pis. a ^ Suc i) = ?exp pis" by (rule IH) finally have id: "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * ?exp pis" by simp have "?exp (pi # pis) = ?exp [(p,i)] * ?exp pis" unfolding pi by simp also have "?exp [(p,i)] = (\(x, i)\ (map (\r. (r, i)) ?rts). [:- x, 1:] ^ Suc i)" by simp also have "\ = (\ x \ ?rts. [:- x, 1:])^Suc i" unfolding prod_list_power by (rule arg_cong[of _ _ prod_list], auto) also have "(\ x \ ?rts. [:- x, 1:]) = p" proof - from fundamental_theorem_algebra_factorized[of p, unfolded \monic p\] obtain as where as: "p = (\a\as. [:- a, 1:])" by auto also have "\ = (\a\set as. [:- a, 1:])" proof (rule sym, rule prod.distinct_set_conv_list, rule ccontr) assume "\ distinct as" from not_distinct_decomp[OF this] obtain as1 as2 as3 a where a: "as = as1 @ [a] @ as2 @ [a] @ as3" by blast define q where "q = (\a\as1 @ as2 @ as3. [:- a, 1:])" have "p = (\a\as. [:- a, 1:])" by fact also have "\ = (\a\([a] @ [a]). [:- a, 1:]) * q" unfolding q_def a map_append prod_list.append by (simp only: ac_simps) also have "\ = [:-a,1:] * [:-a,1:] * q" by simp finally have "p = ([:-a,1:] * [:-a,1:]) * q" by simp hence "[:-a,1:] * [:-a,1:] dvd p" unfolding dvd_def .. with \square_free p\[unfolded square_free_def, THEN conjunct2, rule_format, of "[:-a,1:]"] show False by auto qed also have "set as = {x. poly p x = 0}" unfolding as poly_prod_list by (simp add: o_def, induct as, auto) also have "\ = set ?rts" by (rule roots_of_complex_main(1)[symmetric], insert p deg, auto) also have "(\a\set ?rts. [:- a, 1:]) = (\a\?rts. [:- a, 1:])" by (rule prod.distinct_set_conv_list[OF roots_of_complex_main(2)], insert deg p, auto) finally show ?thesis by simp qed finally have id2: "?exp (pi # pis) = p ^ Suc i * ?exp pis" by simp show ?case unfolding id id2 .. qed simp also have "?exp pis = (\(x, i)\xis. [:- x, 1:] ^ Suc i)" unfolding xis .. finally show ?thesis unfolding p xis by simp qed lemma distinct_factorize_complex_main: assumes "factorize_complex_main p = Some fctrs" shows "distinct (map fst (snd fctrs))" proof - from assms have solvable: "\x\set (snd (yun_factorization gcd p)). degree (fst x) \ 2 \ (\x\set (coeffs (fst x)). x \ \)" by (auto simp add: factorize_complex_main_def case_prod_unfold Let_def map_concat o_def split: if_splits) have sqf: "square_free_factorization p (fst (yun_factorization gcd p), snd (yun_factorization gcd p))" by (rule yun_factorization) simp have "map fst (snd fctrs) = concat (map (\x. (roots_of_complex_main (fst x))) (snd (yun_factorization gcd p)))" using assms by (auto simp add: factorize_complex_main_def case_prod_unfold Let_def map_concat o_def split: if_splits) also have "distinct \" proof (rule distinct_concat, goal_cases) case 1 show ?case proof (subst distinct_map, safe) from square_free_factorizationD(5)[OF sqf] show "distinct (snd (yun_factorization gcd p))" . show "inj_on (\x. (roots_of_complex_main (fst x))) (set (snd (yun_factorization gcd p)))" proof (rule inj_onI, clarify, goal_cases) case (1 a1 b1 a2 b2) { assume neq: "(a1, b1) \ (a2, b2)" from 1(1,2)[THEN square_free_factorizationD(2)[OF sqf]] have "degree a1 \ 0" "degree a2 \ 0" by blast+ hence [simp]: "a1 \ 0" "a2 \ 0" by auto from square_free_factorizationD(3)[OF sqf 1(1,2) neq] have "coprime a1 a2" by simp from solvable 1(1) have "{z. poly a1 z = 0} = set (roots_of_complex_main a1)" by (intro roots_of_complex_main(1) [symmetric]) auto also have "set (roots_of_complex_main a1) = set (roots_of_complex_main a2)" using 1(3) by (subst (1 2) set_remdups [symmetric]) (simp only: fst_conv) also from solvable 1(2) have "\ = {z. poly a2 z = 0}" by (intro roots_of_complex_main) auto finally have "{z. poly a1 z = 0} = {z. poly a2 z = 0}" . with coprime_imp_no_common_roots \coprime a1 a2\ have "{z. poly a1 z = 0} = {}" by auto with fundamental_theorem_of_algebra constant_degree have "degree a1 = 0" by auto with \degree a1 \ 0\ have False by contradiction } thus ?case by blast qed qed next case (2 ys) then obtain f b where fb: "(f, b) \ set (snd (yun_factorization gcd p))" and ys: "ys = roots_of_complex_main f" by auto from square_free_factorizationD(2)[OF sqf fb] have 0: "f \ 0" by auto from solvable[rule_format, OF fb] have f: "degree f \ 2 \ (set (coeffs f) \ \)" by auto show ?case unfolding ys by (rule roots_of_complex_main[OF 0 f]) next case (3 ys zs) then obtain a1 b1 a2 b2 where ab: "(a1, b1) \ set (snd (yun_factorization gcd p))" "(a2, b2) \ set (snd (yun_factorization gcd p))" "ys = roots_of_complex_main a1" "zs = roots_of_complex_main a2" by auto with 3 have neq: "(a1,b1) \ (a2,b2)" by auto from ab(1,2)[THEN square_free_factorizationD(2)[OF sqf]] have [simp]: "a1 \ 0" "a2 \ 0" by auto from square_free_factorizationD(3)[OF sqf ab(1,2) neq] have "coprime a1 a2" by simp have "set ys = {z. poly a1 z = 0}" "set zs = {z. poly a2 z = 0}" by (insert solvable ab(1,2), subst ab, rule roots_of_complex_main; (auto) [])+ with coprime_imp_no_common_roots \coprime a1 a2\ show ?case by auto qed finally show ?thesis . qed lemma factorize_complex_poly: assumes fp: "factorize_complex_poly p = Some (c,qis)" shows "p = smult c (\(q, i)\qis. q ^ i)" "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" proof - from fp[unfolded factorize_complex_poly_def] obtain pis where fp: "factorize_complex_main p = Some (c,pis)" and qis: "qis = map (\(r, i). ([:- r, 1:], Suc i)) pis" by auto from factorize_complex_main[OF fp] have p: "p = smult c (\(x, i)\pis. [:- x, 1:] ^ Suc i)" . show "p = smult c (\(q, i)\qis. q ^ i)" unfolding p qis by (rule arg_cong[of _ _ "\ p. smult c (prod_list p)"], auto) show "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" using linear_irreducible_field[of q] unfolding qis by auto qed end diff --git a/thys/Algebraic_Numbers/Factors_of_Int_Poly.thy b/thys/Algebraic_Numbers/Factors_of_Int_Poly.thy new file mode 100644 --- /dev/null +++ b/thys/Algebraic_Numbers/Factors_of_Int_Poly.thy @@ -0,0 +1,293 @@ +section \Getting a Small Representative via Factorization\ + +text \In this theory we import a factorization algorithm for + integer polynomials to turn a representing polynomial of some algebraic number into + a list of irreducible polynomials where exactly one list element + represents the same number. Moreover, we prove + that the certain polynomial operations preserve irreducibility, so that + no factorization is required.\ + +theory Factors_of_Int_Poly + imports + Berlekamp_Zassenhaus.Factorize_Int_Poly + Algebraic_Numbers_Prelim +begin + +lemma degree_of_gcd: "degree (gcd q 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 + +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 factors_of_int_poly: + 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,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 + +corollary irreducible_represents_imp_degree: + 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 irreducible_preservation: + 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 + +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,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 + +interpretation coeff_lift_hom: + factor_preserving_hom "coeff_lift :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} \ _" + by (unfold_locales, auto) + + + +end \ No newline at end of file diff --git a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy @@ -1,3540 +1,3546 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Real Algebraic Numbers\ text \Whereas we previously only proved the closure properties of algebraic numbers, this theory adds the numeric computations that are required to separate the roots, and to pick unique representatives of algebraic numbers. The development is split into three major parts. First, an ambiguous representation of algebraic numbers is used, afterwards another layer is used with special treatment of rational numbers which still does not admit unique representatives, and finally, a quotient type is created modulo the equivalence. The theory also contains a code-setup to implement real numbers via real algebraic numbers.\ text \The results are taken from the textbook \cite[pages 329ff]{AlgNumbers}.\ theory Real_Algebraic_Numbers imports "Abstract-Rewriting.SN_Order_Carrier" Deriving.Compare_Rat Deriving.Compare_Real Jordan_Normal_Form.Gauss_Jordan_IArray_Impl Algebraic_Numbers Sturm_Rat + Factors_of_Int_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] + (*TODO: move *) lemma ex1_imp_Collect_singleton: "(\!x. P x) \ P x \ Collect P = {x}" proof(intro iffI conjI, unfold conj_imp_eq_imp_imp) assume "Ex1 P" "P x" then show "Collect P = {x}" by blast next assume Px: "Collect P = {x}" then have "P y \ x = y" for y by auto then show "Ex1 P" by auto from Px show "P x" by auto qed lemma ex1_Collect_singleton[consumes 2]: assumes "\!x. P x" and "P x" and "Collect P = {x} \ thesis" shows thesis by (rule assms(3), subst ex1_imp_Collect_singleton[symmetric], insert assms(1,2), auto) lemma ex1_iff_Collect_singleton: "P x \ (\!x. P x) \ Collect P = {x}" by (subst ex1_imp_Collect_singleton[symmetric], auto) context fixes f assumes bij: "bij f" begin lemma bij_imp_ex1_iff: "(\!x. P (f x)) \ (\!y. P y)" (is "?l = ?r") proof (intro iffI) assume l: ?l then obtain x where "P (f x)" by auto with l have *: "{x} = Collect (P o f)" by auto also have "f ` \ = {y. P (f (Hilbert_Choice.inv f y))}" using bij_image_Collect_eq[OF bij] by auto also have "\ = {y. P y}" proof- have "f (Hilbert_Choice.inv f y) = y" for y by (meson bij bij_inv_eq_iff) then show ?thesis by simp qed finally have "Collect P = {f x}" by auto then show ?r by (fold ex1_imp_Collect_singleton, auto) next assume r: ?r then obtain y where "P y" by auto with r have "{y} = Collect P" by auto also have "Hilbert_Choice.inv f ` \ = Collect (P \ f)" using bij_image_Collect_eq[OF bij_imp_bij_inv[OF bij]] bij by (auto simp: inv_inv_eq) finally have "Collect (P o f) = {Hilbert_Choice.inv f y}" by (simp add: o_def) then show ?l by (fold ex1_imp_Collect_singleton, auto) qed lemma bij_ex1_imp_the_shift: assumes ex1: "\!y. P y" shows "(THE x. P (f x)) = Hilbert_Choice.inv f (THE y. P y)" (is "?l = ?r") proof- from ex1 have "P (THE y. P y)" by (rule the1I2) moreover from ex1[folded bij_imp_ex1_iff] have "P (f (THE x. P (f x)))" by (rule the1I2) ultimately have "(THE y. P y) = f (THE x. P (f x))" using ex1 by auto also have "Hilbert_Choice.inv f \ = (THE x. P (f x))" using bij by (simp add: bij_is_inj) finally show "?l = ?r" by auto qed lemma bij_imp_Collect_image: "{x. P (f x)} = Hilbert_Choice.inv f ` {y. P y}" (is "?l = ?g ` _") proof- have "?l = ?g ` f ` ?l" by (simp add: image_comp inv_o_cancel[OF bij_is_inj[OF bij]]) also have "f ` ?l = {f x | x. P (f x)}" by auto also have "\ = {y. P y}" by (metis bij bij_iff) finally show ?thesis. qed lemma bij_imp_card_image: "card (f ` X) = card X" by (metis bij bij_iff card.infinite finite_imageD inj_onI inj_on_iff_eq_card) end lemma bij_imp_card: assumes bij: "bij f" shows "card {x. P (f x)} = card {x. P x}" unfolding bij_imp_Collect_image[OF bij] bij_imp_card_image[OF bij_imp_bij_inv[OF bij]].. lemma bij_add: "bij (\x. x + y :: 'a :: group_add)" (is ?g1) and bij_minus: "bij (\x. x - y :: 'a)" (is ?g2) and inv_add[simp]: "Hilbert_Choice.inv (\x. x + y) = (\x. x - y)" (is ?g3) and inv_minus[simp]: "Hilbert_Choice.inv (\x. x - y) = (\x. x + y)" (is ?g4) proof- have 1: "(\x. x - y) \ (\x. x + y) = id" and 2: "(\x. x + y) \ (\x. x - y) = id" by auto from o_bij[OF 1 2] show ?g1. from o_bij[OF 2 1] show ?g2. from inv_unique_comp[OF 2 1] show ?g3. from inv_unique_comp[OF 1 2] show ?g4. qed lemmas ex1_shift[simp] = bij_imp_ex1_iff[OF bij_add] bij_imp_ex1_iff[OF bij_minus] lemma ex1_the_shift: assumes ex1: "\!y :: 'a :: group_add. P y" shows "(THE x. P (x + d)) = (THE y. P y) - d" and "(THE x. P (x - d)) = (THE y. P y) + d" unfolding bij_ex1_imp_the_shift[OF bij_add ex1] bij_ex1_imp_the_shift[OF bij_minus ex1] by auto lemma card_shift_image[simp]: shows "card ((\x :: 'a :: group_add. x + d) ` X) = card X" and "card ((\x. x - d) ` X) = card X" by (auto simp: bij_imp_card_image[OF bij_add] bij_imp_card_image[OF bij_minus]) lemma irreducible_root_free: fixes p :: "'a :: {idom,comm_ring_1} poly" assumes irr: "irreducible p" shows "root_free p" proof (cases "degree p" "1::nat" rule: linorder_cases) case greater { fix x assume "poly p x = 0" hence "[:-x,1:] dvd p" using poly_eq_0_iff_dvd by blast then obtain r where p: "p = r * [:-x,1:]" by (elim dvdE, auto) have deg: "degree [:-x,1:] = 1" by simp have dvd: "\ [:-x,1:] dvd 1" by (auto simp: poly_dvd_1) from greater have "degree r \ 0" using degree_mult_le[of r "[:-x,1:]", unfolded deg, folded p] by auto then have "\ r dvd 1" by (auto simp: poly_dvd_1) with p irr irreducibleD[OF irr p] dvd have False by auto } thus ?thesis unfolding root_free_def by auto next case less then have deg: "degree p = 0" by auto from deg obtain p0 where p: "p = [:p0:]" using degree0_coeffs by auto with irr have "p \ 0" by auto with p have "poly p x \ 0" for x by auto thus ?thesis by (auto simp: root_free_def) qed (auto simp: root_free_def) (* **************************************************************** *) subsection \Real Algebraic Numbers -- Innermost Layer\ text \We represent a real algebraic number \\\ by a tuple (p,l,r): \\\ is the unique root in the interval [l,r] and l and r have the same sign. We always assume that p is normalized, i.e., p is the unique irreducible and positive content-free polynomial which represents the algebraic number. This representation clearly admits duplicate representations for the same number, e.g. (...,x-3, 3,3) is equivalent to (...,x-3,2,10).\ subsubsection \Basic Definitions\ type_synonym real_alg_1 = "int poly \ rat \ rat" fun poly_real_alg_1 :: "real_alg_1 \ int poly" where "poly_real_alg_1 (p,_,_) = p" fun rai_ub :: "real_alg_1 \ rat" where "rai_ub (_,_,r) = r" fun rai_lb :: "real_alg_1 \ rat" where "rai_lb (_,l,_) = l" abbreviation "roots_below p x \ {y :: real. y \ x \ ipoly p y = 0}" abbreviation(input) unique_root :: "real_alg_1 \ bool" where "unique_root plr \ (\! x. root_cond plr x)" abbreviation the_unique_root :: "real_alg_1 \ real" where "the_unique_root plr \ (THE x. root_cond plr x)" abbreviation real_of_1 where "real_of_1 \ the_unique_root" lemma root_condI[intro]: assumes "of_rat (rai_lb plr) \ x" and "x \ of_rat (rai_ub plr)" and "ipoly (poly_real_alg_1 plr) x = 0" shows "root_cond plr x" using assms by (auto simp: root_cond_def) lemma root_condE[elim]: assumes "root_cond plr x" and "of_rat (rai_lb plr) \ x \ x \ of_rat (rai_ub plr) \ ipoly (poly_real_alg_1 plr) x = 0 \ thesis" shows thesis using assms by (auto simp: root_cond_def) lemma assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows unique_rootD: "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" "root_cond plr x" "x = y \ root_cond plr y" "y = x \ root_cond plr y" and the_unique_root_eqI: "root_cond plr y \ y = x" "root_cond plr y \ x = y" proof - from ur show x: "root_cond plr x" unfolding x_def by (rule theI') have "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from x[unfolded this] show "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" by auto from x ur show "root_cond plr y \ y = x" and "root_cond plr y \ x = y" and "x = y \ root_cond plr y" and "y = x \ root_cond plr y" by auto qed lemma unique_rootE: assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "of_rat l \ x \ x \ of_rat r \ ipoly p x = 0 \ root_cond plr x \ (\y. x = y \ root_cond plr y) \ (\y. y = x \ root_cond plr y) \ thesis" shows thesis by (rule main, unfold x_def p_def l_def r_def; rule unique_rootD[OF ur]) lemma unique_rootI: assumes "\ y. root_cond plr y \ x = y" "root_cond plr x" shows "unique_root plr" using assms by blast definition poly_cond :: "int poly \ bool" where "poly_cond p = (lead_coeff p > 0 \ irreducible p)" lemma poly_condI[intro]: assumes "lead_coeff p > 0" and "irreducible p" shows "poly_cond p" using assms by (auto simp: poly_cond_def) lemma poly_condD: assumes "poly_cond p" shows "irreducible p" and "lead_coeff p > 0" and "root_free p" and "square_free p" and "p \ 0" using assms unfolding poly_cond_def using irreducible_root_free irreducible_imp_square_free cf_pos_def by auto lemma poly_condE[elim]: assumes "poly_cond p" and "irreducible p \ lead_coeff p > 0 \ root_free p \ square_free p \ p \ 0 \ thesis" shows thesis using assms by (auto dest:poly_condD) definition invariant_1 :: "real_alg_1 \ bool" where "invariant_1 tup \ case tup of (p,l,r) \ unique_root (p,l,r) \ sgn l = sgn r \ poly_cond p" lemma invariant_1I: assumes "unique_root plr" and "sgn (rai_lb plr) = sgn (rai_ub plr)" and "poly_cond (poly_real_alg_1 plr)" shows "invariant_1 plr" using assms by (auto simp: invariant_1_def) lemma assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows invariant_1D: "root_cond plr x" "sgn l = sgn r" "sgn x = of_rat (sgn r)" "unique_root plr" "poly_cond p" "degree p > 0" "primitive p" and invariant_1_root_cond: "\ y. root_cond plr y \ y = x" proof - let ?l = "of_rat l :: real" let ?r = "of_rat r :: real" have plr: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from assms show ur: "unique_root plr" and sgn: "sgn l = sgn r" and pc: "poly_cond p" by (auto simp: invariant_1_def) from ur show rc: "root_cond plr x" by (auto simp add: x_def plr intro: theI') from this[unfolded plr] have x: "ipoly p x = 0" and bnd: "?l \ x" "x \ ?r" by auto show "sgn x = of_rat (sgn r)" proof (cases "0::real" "x" rule:linorder_cases) case less with bnd(2) have "0 < ?r" by arith thus ?thesis using less by simp next case equal with bnd have "?l \ 0" "?r \ 0" by auto hence "l \ 0" "r \ 0" by auto with \sgn l = sgn r\ have "l = 0" "r = 0" unfolding sgn_rat_def by (auto split: if_splits) with rc[unfolded plr] show ?thesis by auto next case greater with bnd(1) have "?l < 0" by arith thus ?thesis unfolding \sgn l = sgn r\[symmetric] using greater by simp qed from the_unique_root_eqI[OF ur] rc show "\ y. root_cond plr y \ y = x" by metis { assume "degree p = 0" with poly_zero[OF x, simplified] sgn bnd have "p = 0" by auto with pc have "False" by auto } then show "degree p > 0" by auto with pc show "primitive p" by (intro irreducible_imp_primitive, auto) qed lemma invariant_1E[elim]: assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "root_cond plr x \ sgn l = sgn r \ sgn x = of_rat (sgn r) \ unique_root plr \ poly_cond p \ degree p > 0 \ primitive p \ thesis" shows thesis apply (rule main) using assms(1) unfolding x_def p_def l_def r_def by (auto dest: invariant_1D) lemma invariant_1_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and "sgn l = sgn r" and ur: "unique_root plr" and "poly_cond p" shows "invariant_1 plr \ real_of_1 plr = x" using the_unique_root_eqI[OF ur x] assms by (cases plr, auto intro: invariant_1I) lemma real_of_1_0: assumes "invariant_1 (p,l,r)" shows [simp]: "the_unique_root (p,l,r) = 0 \ r = 0" and [dest]: "l = 0 \ r = 0" and [intro]: "r = 0 \ l = 0" using assms by (auto simp: sgn_0_0) lemma invariant_1_pos: assumes rc: "invariant_1 (p,l,r)" shows [simp]:"the_unique_root (p,l,r) > 0 \ r > 0" (is "?x > 0 \ _") and [simp]:"the_unique_root (p,l,r) < 0 \ r < 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [intro]: "r > 0 \ l > 0" and [dest]: "l > 0 \ r > 0" and [intro]: "r < 0 \ l < 0" and [dest]: "l < 0 \ r < 0" proof(atomize(full),goal_cases) case 1 let ?r = "real_of_rat" from assms[unfolded invariant_1_def] have ur: "unique_root (p,l,r)" and sgn: "sgn l = sgn r" by auto from unique_rootD(1-2)[OF ur] have le: "?r l \ ?x" "?x \ ?r r" by auto from rc show ?case proof (cases r "0::rat" rule:linorder_cases) case greater with sgn have "sgn l = 1" by simp hence l0: "l > 0" by (auto simp: sgn_1_pos) hence "?r l > 0" by auto hence "?x > 0" using le(1) by arith with greater l0 show ?thesis by auto next case equal with real_of_1_0[OF rc] show ?thesis by auto next case less hence "?r r < 0" by auto with le(2) have "?x < 0" by arith with less sgn show ?thesis by (auto simp: sgn_1_neg) qed qed definition invariant_1_2 where "invariant_1_2 rai \ invariant_1 rai \ degree (poly_real_alg_1 rai) > 1" definition poly_cond2 where "poly_cond2 p \ poly_cond p \ degree p > 1" lemma poly_cond2I[intro!]: "poly_cond p \ degree p > 1 \ poly_cond2 p" by (simp add: poly_cond2_def) lemma poly_cond2D: assumes "poly_cond2 p" shows "poly_cond p" and "degree p > 1" using assms by (auto simp: poly_cond2_def) lemma poly_cond2E[elim!]: assumes "poly_cond2 p" and "poly_cond p \ degree p > 1 \ thesis" shows thesis using assms by (auto simp: poly_cond2_def) lemma invariant_1_2_poly_cond2: "invariant_1_2 rai \ poly_cond2 (poly_real_alg_1 rai)" unfolding invariant_1_def invariant_1_2_def poly_cond2_def by auto lemma invariant_1_2I[intro!]: assumes "invariant_1 rai" and "degree (poly_real_alg_1 rai) > 1" shows "invariant_1_2 rai" using assms by (auto simp: invariant_1_2_def) lemma invariant_1_2E[elim!]: assumes "invariant_1_2 rai" and "invariant_1 rai \ degree (poly_real_alg_1 rai) > 1 \ thesis" shows thesis using assms[unfolded invariant_1_2_def] by auto lemma invariant_1_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond2 p" shows "invariant_1_2 plr \ real_of_1 plr = x" using invariant_1_realI[OF x] p sgn ur unfolding p_def l_def r_def by auto subsection \Real Algebraic Numbers = Rational + Irrational Real Algebraic Numbers\ text \In the next representation of real algebraic numbers, we distinguish between rational and irrational numbers. The advantage is that whenever we only work on rational numbers, there is not much overhead involved in comparison to the existing implementation of real numbers which just supports the rational numbers. For irrational numbers we additionally store the number of the root, counting from left to right. For instance $-\sqrt{2}$ and $\sqrt{2}$ would be root number 1 and 2 of $x^2 - 2$.\ subsubsection \Definitions and Algorithms on Raw Type\ datatype real_alg_2 = Rational rat | Irrational nat real_alg_1 fun invariant_2 :: "real_alg_2 \ bool" where "invariant_2 (Irrational n rai) = (invariant_1_2 rai \ n = card(roots_below (poly_real_alg_1 rai) (real_of_1 rai)))" | "invariant_2 (Rational r) = True" fun real_of_2 :: "real_alg_2 \ real" where "real_of_2 (Rational r) = of_rat r" | "real_of_2 (Irrational n rai) = real_of_1 rai" definition of_rat_2 :: "rat \ real_alg_2" where [code_unfold]: "of_rat_2 = Rational" lemma of_rat_2: "real_of_2 (of_rat_2 x) = of_rat x" "invariant_2 (of_rat_2 x)" by (auto simp: of_rat_2_def) (* Invariant type *) typedef real_alg_3 = "Collect invariant_2" morphisms rep_real_alg_3 Real_Alg_Invariant by (rule exI[of _ "Rational 0"], auto) setup_lifting type_definition_real_alg_3 lift_definition real_of_3 :: "real_alg_3 \ real" is real_of_2 . (* *************** *) subsubsection \Definitions and Algorithms on Quotient Type\ quotient_type real_alg = real_alg_3 / "\ x y. real_of_3 x = real_of_3 y" morphisms rep_real_alg Real_Alg_Quotient by (auto simp: equivp_def) metis (* real_of *) lift_definition real_of :: "real_alg \ real" is real_of_3 . lemma real_of_inj: "(real_of x = real_of y) = (x = y)" by (transfer, simp) (* ********************** *) subsubsection \Sign\ definition sgn_1 :: "real_alg_1 \ rat" where "sgn_1 x = sgn (rai_ub x)" lemma sgn_1: "invariant_1 x \ real_of_rat (sgn_1 x) = sgn (real_of_1 x)" unfolding sgn_1_def by auto lemma sgn_1_inj: "invariant_1 x \ invariant_1 y \ real_of_1 x = real_of_1 y \ sgn_1 x = sgn_1 y" by (auto simp: sgn_1_def elim!: invariant_1E) (* ********************** *) subsubsection \Normalization: Bounds Close Together\ lemma unique_root_lr: assumes ur: "unique_root plr" shows "rai_lb plr \ rai_ub plr" (is "?l \ ?r") proof - let ?p = "poly_real_alg_1 plr" from ur[unfolded root_cond_def] have ex1: "\! x :: real. of_rat ?l \ x \ x \ of_rat ?r \ ipoly ?p x = 0" by (cases plr, simp) then obtain x :: real where bnd: "of_rat ?l \ x" "x \ of_rat ?r" and rt: "ipoly ?p x = 0" by auto from bnd have "real_of_rat ?l \ of_rat ?r" by linarith thus "?l \ ?r" by (simp add: of_rat_less_eq) qed locale map_poly_zero_hom_0 = base: zero_hom_0 begin sublocale zero_hom_0 "map_poly hom" by (unfold_locales,auto) end interpretation of_int_poly_hom: map_poly_zero_hom_0 "of_int :: int \ 'a :: {ring_1, ring_char_0}" .. lemma ipoly_roots_finite: "p \ 0 \ finite {x :: 'a :: {idom, ring_char_0}. ipoly p x = 0}" by (rule poly_roots_finite, simp) lemma roots_below_the_unique_root: assumes ur: "unique_root (p,l,r)" shows "roots_below p (the_unique_root (p,l,r)) = roots_below p (of_rat r)" (is "roots_below p ?x = _") proof- from ur have rc: "root_cond (p,l,r) ?x" by (auto dest!: unique_rootD) with ur have x: "{x. root_cond (p,l,r) x} = {?x}" by (auto intro: the_unique_root_eqI) from rc have "?x \ {y. ?x \ y \ y \ of_rat r \ ipoly p y = 0}" by auto with rc have l1x: "... = {?x}" by (intro equalityI, fold x(1), force, simp add: x) have rb:"roots_below p (of_rat r) = roots_below p ?x \ {y. ?x < y \ y \ of_rat r \ ipoly p y = 0}" using rc by auto have emp: "\x. the_unique_root (p, l, r) < x \ x \ {ra. ?x \ ra \ ra \ real_of_rat r \ ipoly p ra = 0}" using l1x by auto with rb show ?thesis by auto qed lemma unique_root_sub_interval: assumes ur: "unique_root (p,l,r)" and rc: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "unique_root (p,l',r')" and "the_unique_root (p,l',r') = the_unique_root (p,l,r)" proof - from between have ord: "real_of_rat l \ of_rat l'" "real_of_rat r' \ of_rat r" by (auto simp: of_rat_less_eq) from rc have lr': "real_of_rat l' \ of_rat r'" by auto with ord have lr: "real_of_rat l \ real_of_rat r" by auto show "\!x. root_cond (p, l', r') x" proof (rule, rule rc) fix y assume "root_cond (p,l',r') y" with ord have "root_cond (p,l,r) y" by (auto intro!:root_condI) from the_unique_root_eqI[OF ur this] show "y = the_unique_root (p,l,r)" by simp qed from the_unique_root_eqI[OF this rc] show "the_unique_root (p,l',r') = the_unique_root (p,l,r)" by simp qed lemma invariant_1_sub_interval: assumes rc: "invariant_1 (p,l,r)" and sub: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "invariant_1 (p,l',r')" and "real_of_1 (p,l',r') = real_of_1 (p,l,r)" proof - let ?r = real_of_rat note rcD = invariant_1D[OF rc] from rc have ur: "unique_root (p, l', r')" and id: "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by (atomize(full), intro conjI unique_root_sub_interval[OF _ sub between], auto) show "real_of_1 (p,l',r') = real_of_1 (p,l,r)" using id by simp from rcD(1)[unfolded split] have "?r l \ ?r r" by auto hence lr: "l \ r" by (auto simp: of_rat_less_eq) from unique_rootD[OF ur] have "?r l' \ ?r r'" by auto hence lr': "l' \ r'" by (auto simp: of_rat_less_eq) have "sgn l' = sgn r'" proof (cases "r" "0::rat" rule: linorder_cases) case less with lr lr' between have "l < 0" "l' < 0" "r' < 0" "r < 0" by auto thus ?thesis unfolding sgn_rat_def by auto next case equal with rcD(2) have "l = 0" using sgn_0_0 by auto with equal between lr' have "l' = 0" "r' = 0" by auto then show ?thesis by auto next case greater with rcD(4) have "sgn r = 1" unfolding sgn_rat_def by (cases "r = 0", auto) with rcD(2) have "sgn l = 1" by simp hence l: "l > 0" unfolding sgn_rat_def by (cases "l = 0"; cases "l < 0"; auto) with lr lr' between have "l > 0" "l' > 0" "r' > 0" "r > 0" by auto thus ?thesis unfolding sgn_rat_def by auto qed with between ur rc show "invariant_1 (p,l',r')" by (auto simp add: invariant_1_def id) qed lemma root_sign_change: assumes p0: "poly (p::real poly) x = 0" and pd_ne0: "poly (pderiv p) x \ 0" obtains d where "0 < d" "sgn (poly p (x - d)) \ sgn (poly p (x + d))" "sgn (poly p (x - d)) \ 0" "0 \ sgn (poly p (x + d))" "\ d' > 0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d))" proof - assume a:"(\d. 0 < d \ sgn (poly p (x - d)) \ sgn (poly p (x + d)) \ sgn (poly p (x - d)) \ 0 \ 0 \ sgn (poly p (x + d)) \ \d'>0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d)) \ thesis)" from pd_ne0 consider "poly (pderiv p) x > 0" | "poly (pderiv p) x < 0" by linarith thus ?thesis proof(cases) case 1 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) < 0" "d1 > 0" using DERIV_pos_inc_left[OF poly_DERIV 1] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) > 0" "d2 > 0" using DERIV_pos_inc_right[OF poly_DERIV 1] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = -1" "sgn (poly p (x + d)) = 1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) next case 2 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) > 0" "d1 > 0" using DERIV_neg_dec_left[OF poly_DERIV 2] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) < 0" "d2 > 0" using DERIV_neg_dec_right[OF poly_DERIV 2] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = 1" "sgn (poly p (x + d)) = -1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) qed qed lemma rational_root_free_degree_iff: assumes rf: "root_free (map_poly rat_of_int p)" and rt: "ipoly p x = 0" shows "(x \ \) = (degree p = 1)" proof assume "x \ \" then obtain y where x: "x = of_rat y" (is "_ = ?x") unfolding Rats_def by blast from rt[unfolded x] have "poly (map_poly rat_of_int p) y = 0" by simp with rf show "degree p = 1" unfolding root_free_def by auto next assume "degree p = 1" from degree1_coeffs[OF this] obtain a b where p: "p = [:a,b:]" and b: "b \ 0" by auto from rt[unfolded p hom_distribs] have "of_int a + x * of_int b = 0" by auto from arg_cong[OF this, of "\ x. (x - of_int a) / of_int b"] have "x = - of_rat (of_int a) / of_rat (of_int b)" using b by auto also have "\ = of_rat (- of_int a / of_int b)" unfolding of_rat_minus of_rat_divide .. finally show "x \ \" by auto qed lemma rational_poly_cond_iff: assumes "poly_cond p" and "ipoly p x = 0" and "degree p > 1" shows "(x \ \) = (degree p = 1)" proof (rule rational_root_free_degree_iff[OF _ assms(2)]) from poly_condD[OF assms(1)] irreducible_connect_rev[of p] assms(3) have p: "irreducible\<^sub>d p" by auto from irreducible\<^sub>d_int_rat[OF this] have "irreducible (map_poly rat_of_int p)" by simp thus "root_free (map_poly rat_of_int p)" by (rule irreducible_root_free) qed lemma poly_cond_degree_gt_1: assumes "poly_cond p" "degree p > 1" "ipoly p x = 0" shows "x \ \" using rational_poly_cond_iff[OF assms(1,3)] assms(2) by simp lemma poly_cond2_no_rat_root: assumes "poly_cond2 p" shows "ipoly p (real_of_rat x) \ 0" using poly_cond_degree_gt_1[of p "real_of_rat x"] assms by auto context fixes p :: "int poly" and x :: "rat" begin lemma gt_rat_sign_change: assumes ur: "unique_root plr" defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes p: "poly_cond2 p" and in_interval: "l \ y" "y \ r" shows "(sgn (ipoly p y) = sgn (ipoly p r)) = (of_rat y > the_unique_root plr)" (is "?gt = _") proof (rule ccontr) have plr[simp]: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) assume "?gt \ (real_of_rat y > the_unique_root plr)" note a = this[unfolded plr] from p have "irreducible p" by auto note nz = poly_cond2_no_rat_root[OF p] hence "p \ 0" unfolding irreducible_def by auto hence p0_real: "real_of_int_poly p \ (0::real poly)" by auto let ?p = "real_of_int_poly p" note urD = unique_rootD[OF ur, simplified] let ?ur = "the_unique_root (p, l, r)" let ?r = real_of_rat from poly_cond2_no_rat_root p have ru:"ipoly p y \ 0" by auto from in_interval have in':"?r l \ ?r y" "?r y \ ?r r" unfolding of_rat_less_eq by auto from p square_free_of_int_poly[of p] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto have ur3:"poly ?p ?ur = 0" using urD(3) by simp from ur have "?ur \ of_rat r" by (auto elim!: unique_rootE) moreover from nz have "ipoly p (real_of_rat r) \ 0" by auto with ur3 have "real_of_rat r \ real_of_1 (p,l,r)" by force ultimately have "?ur < ?r r" by auto hence ur2: "0 < ?r r - ?ur" by linarith from rsquarefree_roots rsf ur3 have pd_nonz:"poly (pderiv ?p) ?ur \ 0" by auto obtain d where d':"\d'. d'>0 \ d' \ d \ sgn (poly ?p (?ur + d')) = sgn (poly ?p (?ur + d)) \ sgn (poly ?p (?ur - d')) = sgn (poly ?p (?ur - d))" "sgn (poly ?p (?ur - d)) \ sgn (poly ?p (?ur + d))" "sgn (poly ?p (?ur + d)) \ 0" and d_ge_0:"d > 0" by (metis root_sign_change[OF ur3 pd_nonz]) have sr:"sgn (poly ?p (?ur + d)) = sgn (poly ?p (?r r))" proof (cases "?r r - ?ur \ d") case True show ?thesis using d'(1)[OF ur2 True] by auto next case False hence less:"?ur + d < ?r r" by auto show ?thesis proof(rule no_roots_inbetween_imp_same_sign[OF less,rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed qed consider "?r l < ?ur - d" "?r l < ?ur" | "0 < ?ur - ?r l" "?ur - ?r l \ d" | "?ur = ?r l" using urD by argo hence sl:"sgn (poly ?p (?ur - d)) = sgn (poly ?p (?r l)) \ 0 = sgn (poly ?p (?r l))" proof (cases) case 1 have "sgn (poly ?p (?r l)) = sgn (poly ?p (?ur - d))" proof(rule no_roots_inbetween_imp_same_sign[OF 1(1),rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed thus ?thesis by auto next case 2 show ?thesis using d'(1)[OF 2] by simp qed (insert ur3,simp) have diff_sign: "sgn (ipoly p l) \ sgn (ipoly p r)" using d'(2-) sr sl real_of_rat_sgn by auto have ur':"\x. real_of_rat l \ x \ x \ real_of_rat y \ ipoly p x = 0 \ \ (?r y \ the_unique_root (p,l,r))" proof(standard+,goal_cases) case (1 x) { assume id: "the_unique_root (p,l,r) = ?r y" from nz[of y] id ur have False by (auto elim!: unique_rootE) } note neq = this have "root_cond (p, l, r) x" unfolding root_cond_def using 1 a ur by (auto elim!: unique_rootE) with conjunct2[OF 1(1)] 1(2-) the_unique_root_eqI[OF ur] show ?case by (auto intro!: neq) qed hence ur'':"\x. real_of_rat y \ x \ x \ real_of_rat r \ poly (real_of_int_poly p) x \ 0 \ \ (?r y \ the_unique_root (p,l,r))" using urD(2,3) by auto have "(sgn (ipoly p y) = sgn (ipoly p r)) = (?r y > the_unique_root (p,l,r))" proof(cases "sgn (ipoly p r) = sgn (ipoly p y)") case True have sgn:"sgn (poly ?p (real_of_rat l)) \ sgn (poly ?p (real_of_rat y))" using True diff_sign by (simp add: real_of_rat_sgn) have ly:"of_rat l < (of_rat y::real)" using in_interval True diff_sign less_eq_rat_def of_rat_less by auto with no_roots_inbetween_imp_same_sign[OF ly,of ?p] sgn ur' True show ?thesis by force next case False hence ne:"sgn (ipoly p (real_of_rat y)) \ sgn (ipoly p (real_of_rat r))" by (simp add: real_of_rat_sgn) have ry:"of_rat y < (of_rat r::real)" using in_interval False diff_sign less_eq_rat_def of_rat_less by auto obtain x where x:"real_of_rat y \ x" "x \ real_of_rat r" "ipoly p x = 0" using no_roots_inbetween_imp_same_sign[OF ry,of ?p] ne by auto hence lx:"real_of_rat l \ x" using in_interval using False a urD by auto have "?ur = x" using x lx ur by (intro the_unique_root_eqI, auto) then show ?thesis using False x by auto qed thus False using diff_sign(1) a ru by(cases "ipoly p r = 0";auto simp:sgn_0_0) qed definition tighten_poly_bounds :: "rat \ rat \ rat \ rat \ rat \ rat" where "tighten_poly_bounds l r sr = (let m = (l + r) / 2; sm = sgn (ipoly p m) in if sm = sr then (l,m,sm) else (m,r,sr))" lemma tighten_poly_bounds: assumes res: "tighten_poly_bounds l r sr = (l',r',sr')" and ur: "unique_root (p,l,r)" and p: "poly_cond2 p" and sr: "sr = sgn (ipoly p r)" shows "root_cond (p,l',r') (the_unique_root (p,l,r))" "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" proof - let ?x = "the_unique_root (p,l,r)" let ?x' = "the_unique_root (p,l',r')" let ?m = "(l + r) / 2" note d = tighten_poly_bounds_def Let_def from unique_root_lr[OF ur] have lr: "l \ r" by auto thus "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" using res sr unfolding d by (auto split: if_splits) hence "l \ ?m" "?m \ r" by auto note le = gt_rat_sign_change[OF ur,simplified,OF p this] note urD = unique_rootD[OF ur] show "root_cond (p,l',r') ?x" proof (cases "sgn (ipoly p ?m) = sgn (ipoly p r)") case *: False with res sr have id: "l' = ?m" "r' = r" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto next case *: True with res sr have id: "l' = l" "r' = ?m" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto qed qed partial_function (tailrec) tighten_poly_bounds_epsilon :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_epsilon l r sr = (if r - l \ x then (l,r,sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_epsilon l' r' sr'))" partial_function (tailrec) tighten_poly_bounds_for_x :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_for_x l r sr = (if x < l \ r < x then (l, r, sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_for_x l' r' sr'))" lemma tighten_poly_bounds_epsilon: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_epsilon l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" and x: "x > 0" shows "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" proof - let ?u = "the_unique_root (p,l,r)" define delta where "delta = x / 2" have delta: "delta > 0" unfolding delta_def using x by auto let ?dist = "\ (l,r,sr). r - l" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_epsilon l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_epsilon l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_epsilon.simps[of l r sr] show "l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "r - l \ x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "r - l > x" by auto let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight[unfolded sr] ur p] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False sr have "tighten_poly_bounds_epsilon L R SR = (l',r',sr')" by auto note IH = 1[of "(L,R,SR)", unfolded tight split lr, rule_format, OF _ ur' this] have "L \ l' \ r' \ R \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert tighten False, auto simp: delta_gt_def delta_def) thus ?thesis using tighten by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" using u by auto qed lemma tighten_poly_bounds_for_x: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_for_x l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" shows "l \ l'" "l' \ r'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" "unique_root (p,l',r')" proof - let ?u = "the_unique_root (p,l,r)" let ?x = "real_of_rat x" define delta where "delta = abs ((u - ?x) / 2)" let ?p = "real_of_int_poly p" note ru = unique_rootD[OF ur] { assume "u = ?x" note u = this[unfolded u] from poly_cond2_no_rat_root[OF p] ur have False by (elim unique_rootE, auto simp: u) } hence delta: "delta > 0" unfolding delta_def by auto let ?dist = "\ (l,r,sr). real_of_rat (r - l)" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_for_x l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) let ?l = "real_of_rat l" let ?r = "real_of_rat r" show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_for_x l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_for_x.simps[of l r] show "l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "x < l \ r < x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "?l \ ?x" "?x \ ?r" by (auto simp: of_rat_less_eq) let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight ur p sr] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False have "tighten_poly_bounds_for_x L R SR = (l',r',sr')" by auto note IH = 1[of ?tight, unfolded tight split lr, rule_format, OF _ ur' this] let ?DIFF = "real_of_rat (R - L)" let ?diff = "real_of_rat (r - l)" have diff0: "0 \ ?DIFF" using tighten(3) by (metis cancel_comm_monoid_add_class.diff_cancel diff_right_mono of_rat_less_eq of_rat_hom.hom_zero) have *: "r - l - (r - l) / 2 = (r - l) / 2" by (auto simp: field_simps) have "delta_gt delta ?diff ?DIFF = (abs (u - of_rat x) \ real_of_rat (r - l) * 1)" unfolding delta_gt_def tighten(5) delta_def of_rat_diff[symmetric] * by (simp add: hom_distribs) also have "real_of_rat (r - l) * 1 = ?r - ?l" unfolding of_rat_divide of_rat_mult of_rat_diff by auto also have "abs (u - of_rat x) \ ?r - ?l" using x ur by (elim unique_rootE, auto simp: u) finally have delta: "delta_gt delta ?diff ?DIFF" . have "L \ l' \ r' \ R \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert delta diff0 tighten(6), auto) with \l \ L\ \R \ r\ show ?thesis by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show *: "l \ l'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" unfolding u by auto from *(3)[unfolded split] have "real_of_rat l' \ of_rat r'" by auto thus "l' \ r'" unfolding of_rat_less_eq . show "unique_root (p,l',r')" using ur *(1-3) p poly_condD(5) u unique_root_sub_interval(1) by blast qed end definition real_alg_precision :: rat where "real_alg_precision \ Rat.Fract 1 2" lemma real_alg_precision: "real_alg_precision > 0" by eval definition normalize_bounds_1_main :: "rat \ real_alg_1 \ real_alg_1" where "normalize_bounds_1_main eps rai = (case rai of (p,l,r) \ let (l',r',sr') = tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)); fr = rat_of_int (floor r'); (l'',r'',_) = tighten_poly_bounds_for_x p fr l' r' sr' in (p,l'',r''))" definition normalize_bounds_1 :: "real_alg_1 \ real_alg_1" where "normalize_bounds_1 = (normalize_bounds_1_main real_alg_precision)" context fixes p q and l r :: rat assumes cong: "\ x. real_of_rat l \ x \ x \ of_rat r \ (ipoly p x = (0 :: real)) = (ipoly q x = 0)" begin lemma root_cond_cong: "root_cond (p,l,r) = root_cond (q,l,r)" by (intro ext, insert cong, auto simp: root_cond_def) lemma the_unique_root_cong: "the_unique_root (p,l,r) = the_unique_root (q,l,r)" unfolding root_cond_cong .. lemma unique_root_cong: "unique_root (p,l,r) = unique_root (q,l,r)" unfolding root_cond_cong .. end lemma normalize_bounds_1_main: assumes eps: "eps > 0" and rc: "invariant_1_2 x" defines y: "y \ normalize_bounds_1_main eps x" shows "invariant_1_2 y \ (real_of_1 y = real_of_1 x)" proof - obtain p l r where x: "x = (p,l,r)" by (cases x) auto note rc = rc[unfolded x] obtain l' r' sr' where tb: "tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "rat_of_int (floor r')" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) from y[unfolded normalize_bounds_1_main_def x] tb tbx have y: "y = (p, l'', r'')" by (auto simp: Let_def) from rc have "unique_root (p, l, r)" and p2: "poly_cond2 p" by auto from tighten_poly_bounds_epsilon[OF this tb refl eps] have bnd: "l \ l'" "r' \ r" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and eps: "r' - l' \ eps" (* currently not relevant for lemma *) and sr': "sr' = sgn (ipoly p r')" by auto from invariant_1_sub_interval[OF _ rc' bnd] rc have inv': "invariant_1 (p, l', r')" and eq: "real_of_1 (p, l', r') = real_of_1 (p, l, r)" by auto have bnd: "l' \ l''" "r'' \ r'" and rc': "root_cond (p, l'', r'') (the_unique_root (p, l', r'))" by (rule tighten_poly_bounds_for_x[OF _ p2 tbx sr'], fact invariant_1D[OF inv'])+ from invariant_1_sub_interval[OF inv' rc' bnd] p2 eq show ?thesis unfolding y x by auto qed lemma normalize_bounds_1: assumes x: "invariant_1_2 x" shows "invariant_1_2 (normalize_bounds_1 x) \ (real_of_1 (normalize_bounds_1 x) = real_of_1 x)" proof(cases x) case xx:(fields p l r) let ?res = "(p,l,r)" have norm: "normalize_bounds_1 x = (normalize_bounds_1_main real_alg_precision ?res)" unfolding normalize_bounds_1_def by (simp add: xx) from x have x: "invariant_1_2 ?res" "real_of_1 ?res = real_of_1 x" unfolding xx by auto from normalize_bounds_1_main[OF real_alg_precision x(1)] x(2-) show ?thesis unfolding normalize_bounds_1_def xx by auto qed lemma normalize_bound_1_poly: "poly_real_alg_1 (normalize_bounds_1 rai) = poly_real_alg_1 rai" unfolding normalize_bounds_1_def normalize_bounds_1_main_def Let_def by (auto split: prod.splits) definition real_alg_2_main :: "root_info \ real_alg_1 \ real_alg_2" where "real_alg_2_main ri rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root ri r) (p',l,r)))" definition real_alg_2 :: "real_alg_1 \ real_alg_2" where "real_alg_2 rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root (root_info p) r) (p',l,r)))" lemma degree_1_ipoly: assumes "degree p = Suc 0" shows "ipoly p x = 0 \ (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" proof - from roots1[of "map_poly real_of_int p"] assms have "ipoly p x = 0 \ x \ {roots1 (real_of_int_poly p)}" by auto also have "\ = (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" unfolding Fract_of_int_quotient roots1_def hom_distribs by auto finally show ?thesis . qed lemma invariant_1_degree_0: assumes inv: "invariant_1 rai" shows "degree (poly_real_alg_1 rai) \ 0" (is "degree ?p \ 0") proof (rule notI) assume deg: "degree ?p = 0" from inv have "ipoly ?p (real_of_1 rai) = 0" by auto with deg have "?p = 0" by (meson less_Suc0 representsI represents_degree) with inv show False by auto qed lemma real_alg_2_main: assumes inv: "invariant_1 rai" defines [simp]: "p \ poly_real_alg_1 rai" assumes ric: "irreducible (poly_real_alg_1 rai) \ root_info_cond ri (poly_real_alg_1 rai)" shows "invariant_2 (real_alg_2_main ri rai)" "real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" proof (atomize(full)) define l r where [simp]: "l \ rai_lb rai" and [simp]: "r \ rai_ub rai" show "invariant_2 (real_alg_2_main ri rai) \ real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" unfolding id using invariant_1D proof (cases "degree p" "Suc 0" rule: linorder_cases) case deg: equal hence id: "real_alg_2_main ri rai = Rational (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding real_alg_2_main_def Let_def by auto note rc = invariant_1D[OF inv] from degree_1_ipoly[OF deg, of "the_unique_root rai"] rc(1) show ?thesis unfolding id by auto next case deg: greater with inv have inv: "invariant_1_2 rai" unfolding p_def by auto define rai' where "rai' = normalize_bounds_1 rai" have rai': "real_of_1 rai = real_of_1 rai'" and inv': "invariant_1_2 rai'" unfolding rai'_def using normalize_bounds_1[OF inv] by auto obtain p' l' r' where "rai' = (p',l',r')" by (cases rai') with arg_cong[OF rai'_def, of poly_real_alg_1, unfolded normalize_bound_1_poly] split have split: "rai' = (p,l',r')" by auto from inv'[unfolded split] have "poly_cond p" by auto from poly_condD[OF this] have irr: "irreducible p" by simp from ric irr have ric: "root_info_cond ri p" by auto have id: "real_alg_2_main ri rai = (Irrational (root_info.number_root ri r') rai')" unfolding real_alg_2_main_def Let_def using deg split rai'_def by (auto simp: rai'_def rai') show ?thesis unfolding id using rai' root_info_condD(2)[OF ric] inv'[unfolded split] apply (elim invariant_1_2E invariant_1E) using inv' by(auto simp: split roots_below_the_unique_root) next case deg: less then have "degree p = 0" by auto from this invariant_1_degree_0[OF inv] have "p = 0" by simp with inv show ?thesis by auto qed qed lemma real_alg_2: assumes "invariant_1 rai" shows "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" proof - have deg: "0 < degree (poly_real_alg_1 rai)" using assms by auto have "real_alg_2 rai = real_alg_2_main (root_info (poly_real_alg_1 rai)) rai" unfolding real_alg_2_def real_alg_2_main_def Let_def by auto from real_alg_2_main[OF assms root_info, folded this, simplified] deg show "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" by auto qed lemma invariant_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond p" shows "invariant_2 (real_alg_2 plr) \ real_of_2 (real_alg_2 plr) = x" using invariant_1_realI[OF x,folded p_def l_def r_def] sgn ur p real_alg_2[of plr] by auto (* ********************* *) subsubsection \Comparisons\ fun compare_rat_1 :: "rat \ real_alg_1 \ order" where "compare_rat_1 x (p,l,r) = (if x < l then Lt else if x > r then Gt else if sgn (ipoly p x) = sgn(ipoly p r) then Gt else Lt)" lemma compare_rat_1: assumes rai: "invariant_1_2 y" shows "compare_rat_1 x y = compare (of_rat x) (real_of_1 y)" proof- define p l r where "p \ poly_real_alg_1 y" "l \ rai_lb y" "r \ rai_ub y" then have y [simp]: "y = (p,l,r)" by (cases y, auto) from rai have ur: "unique_root y" by auto show ?thesis proof (cases "x < l \ x > r") case True { assume xl: "x < l" hence "real_of_rat x < of_rat l" unfolding of_rat_less by auto with rai have "of_rat x < the_unique_root y" by (auto elim!: invariant_1E) with xl rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } moreover { assume xr: "\ x < l" "x > r" hence "real_of_rat x > of_rat r" unfolding of_rat_less by auto with rai have "of_rat x > the_unique_root y" by (auto elim!: invariant_1E) with xr rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } ultimately show ?thesis using True by auto next case False have 0: "ipoly p (real_of_rat x) \ 0" by (rule poly_cond2_no_rat_root, insert rai, auto) with rai have diff: "real_of_1 y \ of_rat x" by (auto elim!: invariant_1E) have "\ P. (1 < degree (poly_real_alg_1 y) \ \!x. root_cond y x \ poly_cond p \ P) \ P" using poly_real_alg_1.simps y rai invariant_1_2E invariant_1E by metis from this[OF gt_rat_sign_change] False have left: "compare_rat_1 x y = (if real_of_rat x \ the_unique_root y then Lt else Gt)" by (auto simp:poly_cond2_def) also have "\ = compare (real_of_rat x) (real_of_1 y)" using diff by (auto simp: compare_real_def comparator_of_def) finally show ?thesis . qed qed lemma cf_pos_0[simp]: "\ cf_pos 0" unfolding cf_pos_def by auto (* ********************* *) subsubsection\Negation\ fun uminus_1 :: "real_alg_1 \ real_alg_1" where "uminus_1 (p,l,r) = (abs_int_poly (poly_uminus p), -r, -l)" lemma uminus_1: assumes x: "invariant_1 x" defines y: "y \ uminus_1 x" shows "invariant_1 y \ (real_of_1 y = - real_of_1 x)" proof (cases x) case plr: (fields p l r) from x plr have inv: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] from plr have x: "x = (p,l,r)" by simp let ?p = "poly_uminus p" let ?mp = "abs_int_poly ?p" have y: "y = (?mp, -r , -l)" unfolding y plr by (simp add: Let_def) { fix y assume "root_cond (?mp, - r, - l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "- of_rat r \ y" "y \ - of_rat l" unfolding root_cond_def by (auto simp: of_rat_minus) from mpy have id: "ipoly p (- y) = 0" by auto from bnd have bnd: "of_rat l \ - y" "-y \ of_rat r" by auto from id bnd have "root_cond (p, l, r) (-y)" unfolding root_cond_def by auto with inv x have "real_of_1 x = -y" by (auto intro!: the_unique_root_eqI) then have "-real_of_1 x = y" by auto } note inj = this have rc: "root_cond (?mp, - r, - l) (- real_of_1 x)" using * unfolding root_cond_def y x by (auto simp: of_rat_minus sgn_minus_rat) from inj rc have ur': "unique_root (?mp, -r, -l)" by (auto intro: unique_rootI) with rc have the: "- real_of_1 x = the_unique_root (?mp, -r, -l)" by (auto intro: the_unique_root_eqI) have xp: "p represents (real_of_1 x)" using * unfolding root_cond_def split represents_def x by auto from * have mon: "lead_coeff ?mp > 0" by (unfold pos_poly_abs_poly, auto) from poly_uminus_irreducible * have mi: "irreducible ?mp" by auto from mi mon have pc': "poly_cond ?mp" by (auto simp: cf_pos_def) from poly_condD[OF pc'] have irr: "irreducible ?mp" by auto show ?thesis unfolding y apply (intro invariant_1_realI ur' rc) using pc' inv by auto qed lemma uminus_1_2: assumes x: "invariant_1_2 x" defines y: "y \ uminus_1 x" shows "invariant_1_2 y \ (real_of_1 y = - real_of_1 x)" proof - from x have "invariant_1 x" by auto from uminus_1[OF this] have *: "real_of_1 y = - real_of_1 x" "invariant_1 y" unfolding y by auto obtain p l r where id: "x = (p,l,r)" by (cases x) from x[unfolded id] have "degree p > 1" by auto moreover have "poly_real_alg_1 y = abs_int_poly (poly_uminus p)" unfolding y id uminus_1.simps split Let_def by auto ultimately have "degree (poly_real_alg_1 y) > 1" by simp with * show ?thesis by auto qed fun uminus_2 :: "real_alg_2 \ real_alg_2" where "uminus_2 (Rational r) = Rational (-r)" | "uminus_2 (Irrational n x) = real_alg_2 (uminus_1 x)" lemma uminus_2: assumes "invariant_2 x" shows "real_of_2 (uminus_2 x) = uminus (real_of_2 x)" "invariant_2 (uminus_2 x)" using assms real_alg_2 uminus_1 by (atomize(full), cases x, auto simp: hom_distribs) declare uminus_1.simps[simp del] lift_definition uminus_3 :: "real_alg_3 \ real_alg_3" is uminus_2 by (auto simp: uminus_2) lemma uminus_3: "real_of_3 (uminus_3 x) = - real_of_3 x" by (transfer, auto simp: uminus_2) instantiation real_alg :: uminus begin lift_definition uminus_real_alg :: "real_alg \ real_alg" is uminus_3 by (simp add: uminus_3) instance .. end lemma uminus_real_alg: "- (real_of x) = real_of (- x)" by (transfer, rule uminus_3[symmetric]) (* ********************* *) subsubsection\Inverse\ fun inverse_1 :: "real_alg_1 \ real_alg_2" where "inverse_1 (p,l,r) = real_alg_2 (abs_int_poly (reflect_poly p), inverse r, inverse l)" lemma invariant_1_2_of_rat: assumes rc: "invariant_1_2 rai" shows "real_of_1 rai \ of_rat x" proof - obtain p l r where rai: "rai = (p, l, r)" by (cases rai, auto) from rc[unfolded rai] have "poly_cond2 p" "ipoly p (the_unique_root (p, l, r)) = 0" by (auto elim!: invariant_1E) from poly_cond2_no_rat_root[OF this(1), of x] this(2) show ?thesis unfolding rai by auto qed lemma inverse_1: assumes rcx: "invariant_1_2 x" defines y: "y \ inverse_1 x" shows "invariant_2 y \ (real_of_2 y = inverse (real_of_1 x))" proof (cases x) case x: (fields p l r) from x rcx have rcx: "invariant_1_2 (p,l,r)" by auto from invariant_1_2_poly_cond2[OF rcx] have pc2: "poly_cond2 p" by simp have x0: "real_of_1 (p,l,r) \ 0" using invariant_1_2_of_rat[OF rcx, of 0] x by auto let ?x = "real_of_1 (p,l,r)" let ?mp = "abs_int_poly (reflect_poly p)" from x0 rcx have lr0: "l \ 0" and "r \ 0" by auto from x0 rcx have y: "y = real_alg_2 (?mp, inverse r, inverse l)" unfolding y x Let_def inverse_1.simps by auto from rcx have mon: "lead_coeff ?mp > 0" by (unfold lead_coeff_abs_int_poly, auto) { fix y assume "root_cond (?mp, inverse r, inverse l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "inverse (of_rat r) \ y" "y \ inverse (of_rat l)" unfolding root_cond_def by (auto simp: of_rat_inverse) from sgn_real_mono[OF bnd(1)] sgn_real_mono[OF bnd(2)] have "sgn (of_rat r) \ sgn y" "sgn y \ sgn (of_rat l)" by (simp_all add: algebra_simps) with rcx have sgn: "sgn (inverse (of_rat r)) = sgn y" "sgn y = sgn (inverse (of_rat l))" unfolding sgn_inverse inverse_sgn by (auto simp add: real_of_rat_sgn intro: order_antisym) from sgn[simplified, unfolded real_of_rat_sgn] lr0 have "y \ 0" by (auto simp:sgn_0_0) with mpy have id: "ipoly p (inverse y) = 0" by (auto simp: ipoly_reflect_poly) from inverse_le_sgn[OF sgn(1) bnd(1)] inverse_le_sgn[OF sgn(2) bnd(2)] have bnd: "of_rat l \ inverse y" "inverse y \ of_rat r" by auto from id bnd have "root_cond (p,l,r) (inverse y)" unfolding root_cond_def by auto from rcx this x0 have "?x = inverse y" by auto then have "inverse ?x = y" by auto } note inj = this have rc: "root_cond (?mp, inverse r, inverse l) (inverse ?x)" using rcx x0 apply (elim invariant_1_2E invariant_1E) by (simp add: root_cond_def of_rat_inverse real_of_rat_sgn inverse_le_iff_sgn ipoly_reflect_poly) from inj rc have ur: "unique_root (?mp, inverse r, inverse l)" by (auto intro: unique_rootI) with rc have the: "the_unique_root (?mp, inverse r, inverse l) = inverse ?x" by (auto intro: the_unique_root_eqI) have xp: "p represents ?x" unfolding split represents_def using rcx by (auto elim!: invariant_1E) from reflect_poly_irreducible[OF _ xp x0] poly_condD rcx have mi: "irreducible ?mp" by auto from mi mon have un: "poly_cond ?mp" by (auto simp: poly_cond_def) show ?thesis using rcx rc ur unfolding y by (intro invariant_2_realI, auto simp: x y un) qed fun inverse_2 :: "real_alg_2 \ real_alg_2" where "inverse_2 (Rational r) = Rational (inverse r)" | "inverse_2 (Irrational n x) = inverse_1 x" lemma inverse_2: assumes "invariant_2 x" shows "real_of_2 (inverse_2 x) = inverse (real_of_2 x)" "invariant_2 (inverse_2 x)" using assms by (atomize(full), cases x, auto simp: real_alg_2 inverse_1 hom_distribs) lift_definition inverse_3 :: "real_alg_3 \ real_alg_3" is inverse_2 by (auto simp: inverse_2) lemma inverse_3: "real_of_3 (inverse_3 x) = inverse (real_of_3 x)" by (transfer, auto simp: inverse_2) (* ********************* *) subsubsection\Floor\ fun floor_1 :: "real_alg_1 \ int" where "floor_1 (p,l,r) = (let (l',r',sr') = tighten_poly_bounds_epsilon p (1/2) l r (sgn (ipoly p r)); fr = floor r'; fl = floor l'; fr' = rat_of_int fr in (if fr = fl then fr else let (l'',r'',sr'') = tighten_poly_bounds_for_x p fr' l' r' sr' in if fr' < l'' then fr else fl))" lemma floor_1: assumes "invariant_1_2 x" shows "floor (real_of_1 x) = floor_1 x" proof (cases x) case (fields p l r) obtain l' r' sr' where tbe: "tighten_poly_bounds_epsilon p (1 / 2) l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "floor r'" let ?fl = "floor l'" let ?fr' = "rat_of_int ?fr" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr' l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) note rc = assms[unfolded fields] hence rc1: "invariant_1 (p,l,r)" by auto have id: "floor_1 x = ((if ?fr = ?fl then ?fr else if ?fr' < l'' then ?fr else ?fl))" unfolding fields floor_1.simps tbe Let_def split tbx by simp let ?x = "real_of_1 x" have x: "?x = the_unique_root (p,l,r)" unfolding fields by simp have bnd: "l \ l'" "r' \ r" "r' - l' \ 1 / 2" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and sr': "sr' = sgn (ipoly p r')" by (atomize(full), intro conjI tighten_poly_bounds_epsilon[OF _ _ tbe refl],insert rc,auto elim!: invariant_1E) let ?r = real_of_rat from rc'[folded x, unfolded split] have ineq: "?r l' \ ?x" "?x \ ?r r'" "?r l' \ ?r r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq by simp have flr: "?fl \ ?fr" by (rule floor_mono[OF lr']) from invariant_1_sub_interval[OF rc1 rc' bnd(1,2)] have rc': "invariant_1 (p, l', r')" and id': "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by auto with rc have rc2': "invariant_1_2 (p, l', r')" by auto have x: "?x = the_unique_root (p,l',r')" unfolding fields using id' by simp { assume "?fr \ ?fl" with flr have flr: "?fl \ ?fr - 1" by simp have "?fr' \ r'" "l' \ ?fr'" using flr bnd by linarith+ } note fl_diff = this show ?thesis proof (cases "?fr = ?fl") case True hence id1: "floor_1 x = ?fr" unfolding id by auto from True have id: "floor (?r l') = floor (?r r')" by simp have "floor ?x \ floor (?r r')" by (rule floor_mono[OF ineq(2)]) moreover have "floor (?r l') \ floor ?x" by (rule floor_mono[OF ineq(1)]) ultimately have "floor ?x = floor (?r r')" unfolding id by (simp add: id) then show ?thesis by (simp add: id1) next case False with id have id: "floor_1 x = (if ?fr' < l'' then ?fr else ?fl)" by simp from rc2' have "unique_root (p,l',r')" "poly_cond2 p" by auto from tighten_poly_bounds_for_x[OF this tbx sr'] have ineq': "l' \ l''" "r'' \ r'" and lr'': "l'' \ r''" and rc'': "root_cond (p,l'',r'') ?x" and fr': "\ (l'' \ ?fr' \ ?fr' \ r'')" unfolding x by auto from rc''[unfolded split] have ineq'': "?r l'' \ ?x" "?x \ ?r r''" by auto from False have "?fr \ ?fl" by auto note fr = fl_diff[OF this] show ?thesis proof (cases "?fr' < l''") case True with id have id: "floor_1 x = ?fr" by simp have "floor ?x \ ?fr" using floor_mono[OF ineq(2)] by simp moreover from True have "?r ?fr' < ?r l''" unfolding of_rat_less . with ineq''(1) have "?r ?fr' \ ?x" by simp from floor_mono[OF this] have "?fr \ floor ?x" by simp ultimately show ?thesis unfolding id by auto next case False with id have id: "floor_1 x = ?fl" by simp from False have "l'' \ ?fr'" by auto from floor_mono[OF ineq(1)] have "?fl \ floor ?x" by simp moreover have "floor ?x \ ?fl" proof - from False fr' have fr': "r'' < ?fr'" by auto hence "floor r'' < ?fr" by linarith with floor_mono[OF ineq''(2)] have "floor ?x \ ?fr - 1" by auto also have "?fr - 1 = floor (r' - 1)" by simp also have "\ \ ?fl" by (rule floor_mono, insert bnd, auto) finally show ?thesis . qed ultimately show ?thesis unfolding id by auto qed qed qed (* ********************* *) subsubsection\Generic Factorization and Bisection Framework\ lemma card_1_Collect_ex1: assumes "card (Collect P) = 1" shows "\! x. P x" proof - from assms[unfolded card_eq_1_iff] obtain x where "Collect P = {x}" by auto thus ?thesis by (intro ex1I[of _ x], auto) qed fun sub_interval :: "rat \ rat \ rat \ rat \ bool" where "sub_interval (l,r) (l',r') = (l' \ l \ r \ r')" fun in_interval :: "rat \ rat \ real \ bool" where "in_interval (l,r) x = (of_rat l \ x \ x \ of_rat r)" definition converges_to :: "(nat \ rat \ rat) \ real \ bool" where "converges_to f x \ (\ n. in_interval (f n) x \ sub_interval (f (Suc n)) (f n)) \ (\ (eps :: real) > 0. \ n l r. f n = (l,r) \ of_rat r - of_rat l \ eps)" context fixes bnd_update :: "'a \ 'a" and bnd_get :: "'a \ rat \ rat" begin definition at_step :: "(nat \ rat \ rat) \ nat \ 'a \ bool" where "at_step f n a \ \ i. bnd_get ((bnd_update ^^ i) a) = f (n + i)" partial_function (tailrec) select_correct_factor_main :: "'a \ (int poly \ root_info)list \ (int poly \ root_info)list \ rat \ rat \ nat \ (int poly \ root_info) \ rat \ rat" where [code]: "select_correct_factor_main bnd todo old l r n = (case todo of Nil \ if n = 1 then (hd old, l, r) else let bnd' = bnd_update bnd in (case bnd_get bnd' of (l,r) \ select_correct_factor_main bnd' old [] l r 0) | Cons (p,ri) todo \ let m = root_info.l_r ri l r in if m = 0 then select_correct_factor_main bnd todo old l r n else select_correct_factor_main bnd todo ((p,ri) # old) l r (n + m))" definition select_correct_factor :: "'a \ (int poly \ root_info)list \ (int poly \ root_info) \ rat \ rat" where "select_correct_factor init polys = (case bnd_get init of (l,r) \ select_correct_factor_main init polys [] l r 0)" lemma select_correct_factor_main: assumes conv: "converges_to f x" and at: "at_step f i a" and res: "select_correct_factor_main a todo old l r n = ((q,ri_fin),(l_fin,r_fin))" and bnd: "bnd_get a = (l,r)" and ri: "\ q ri. (q,ri) \ set todo \ set old \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set todo \ set old \ q \ 0" and ex: "\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and dist: "distinct (map fst (todo @ old))" and old: "\ q ri. (q,ri) \ set old \ root_info.l_r ri l r \ 0" and un: "\ x :: real. (\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0) \ \!q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and n: "n = sum_list (map (\ (q,ri). root_info.l_r ri l r) old)" shows "unique_root (q,l_fin,r_fin) \ (q,ri_fin) \ set todo \ set old \ x = the_unique_root (q,l_fin,r_fin)" proof - define orig where "orig = set todo \ set old" have orig: "set todo \ set old \ orig" unfolding orig_def by auto let ?rts = "{x :: real. \ q ri. (q,ri) \ orig \ ipoly q x = 0}" define rts where "rts = ?rts" let ?h = "\ (x,y). abs (x - y)" let ?r = real_of_rat have rts: "?rts = (\ ((\ (q,ri). {x. ipoly q x = 0}) ` set (todo @ old)))" unfolding orig_def by auto have "finite rts" unfolding rts rts_def using finite_ipoly_roots[OF q0] finite_set[of "todo @ old"] by auto hence fin: "finite (rts \ rts - Id)" by auto define diffs where "diffs = insert 1 {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" have "finite {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" by (rule subst[of _ _ finite, OF _ finite_imageI[OF fin, of ?h]], auto) hence diffs: "finite diffs" "diffs \ {}" unfolding diffs_def by auto define eps where "eps = Min diffs / 2" have "\ x. x \ diffs \ x > 0" unfolding diffs_def by auto with Min_gr_iff[OF diffs] have eps: "eps > 0" unfolding eps_def by auto note conv = conv[unfolded converges_to_def] from conv eps obtain N L R where N: "f N = (L,R)" "?r R - ?r L \ eps" by auto obtain pair where pair: "pair = (todo,i)" by auto define rel where "rel = measures [ \ (t,i). N - i, \ (t :: (int poly \ root_info) list,i). length t]" have wf: "wf rel" unfolding rel_def by simp show ?thesis using at res bnd ri q0 ex dist old un n pair orig proof (induct pair arbitrary: todo i old a l r n rule: wf_induct[OF wf]) case (1 pair todo i old a l r n) note IH = 1(1)[rule_format] note at = 1(2) note res = 1(3)[unfolded select_correct_factor_main.simps[of _ todo]] note bnd = 1(4) note ri = 1(5) note q0 = 1(6) note ex = 1(7) note dist = 1(8) note old = 1(9) note un = 1(10) note n = 1(11) note pair = 1(12) note orig = 1(13) from at[unfolded at_step_def, rule_format, of 0] bnd have fi: "f i = (l,r)" by auto with conv have inx: "in_interval (f i) x" by blast hence lxr: "?r l \ x" "x \ ?r r" unfolding fi by auto from order.trans[OF this] have lr: "l \ r" unfolding of_rat_less_eq . show ?case proof (cases todo) case (Cons rri tod) obtain s ri where rri: "rri = (s,ri)" by force with Cons have todo: "todo = (s,ri) # tod" by simp note res = res[unfolded todo list.simps split Let_def] from root_info_condD(1)[OF ri[of s ri, unfolded todo] lr] have ri': "root_info.l_r ri l r = card {x. root_cond (s, l, r) x}" by auto from q0 have s0: "s \ 0" unfolding todo by auto from finite_ipoly_roots[OF s0] have fins: "finite {x. root_cond (s, l, r) x}" unfolding root_cond_def by auto have rel: "((tod,i), pair) \ rel" unfolding rel_def pair todo by simp show ?thesis proof (cases "root_info.l_r ri l r = 0") case True with res have res: "select_correct_factor_main a tod old l r n = ((q, ri_fin), l_fin, r_fin)" by auto from ri'[symmetric, unfolded True] fins have empty: "{x. root_cond (s, l, r) x} = {}" by simp from ex lxr empty have ex': "(\q. q \ fst ` set tod \ fst ` set old \ ipoly q x = 0)" unfolding todo root_cond_def split by auto have "unique_root (q, l_fin, r_fin) \ (q, ri_fin) \ set tod \ set old \ x = the_unique_root (q, l_fin, r_fin)" proof (rule IH[OF rel at res bnd ri _ ex' _ _ _ n refl], goal_cases) case (5 y) thus ?case using un[of y] unfolding todo by auto next case 2 thus ?case using q0 unfolding todo by auto qed (insert dist old orig, auto simp: todo) thus ?thesis unfolding todo by auto next case False with res have res: "select_correct_factor_main a tod ((s, ri) # old) l r (n + root_info.l_r ri l r) = ((q, ri_fin), l_fin, r_fin)" by auto from ex have ex': "\q. q \ fst ` set tod \ fst ` set ((s, ri) # old) \ ipoly q x = 0" unfolding todo by auto from dist have dist: "distinct (map fst (tod @ (s, ri) # old))" unfolding todo by auto have id: "set todo \ set old = set tod \ set ((s, ri) # old)" unfolding todo by simp show ?thesis unfolding id proof (rule IH[OF rel at res bnd ri _ ex' dist], goal_cases) case 4 thus ?case using un unfolding todo by auto qed (insert old False orig, auto simp: q0 todo n) qed next case Nil note res = res[unfolded Nil list.simps Let_def] from ex[unfolded Nil] lxr obtain s where "s \ fst ` set old \ root_cond (s,l,r) x" unfolding root_cond_def by auto then obtain q1 ri1 old' where old': "old = (q1,ri1) # old'" using id by (cases old, auto) let ?ri = "root_info.l_r ri1 l r" from old[unfolded old'] have 0: "?ri \ 0" by auto from n[unfolded old'] 0 have n0: "n \ 0" by auto from ri[unfolded old'] have ri': "root_info_cond ri1 q1" by auto show ?thesis proof (cases "n = 1") case False with n0 have n1: "n > 1" by auto obtain l' r' where bnd': "bnd_get (bnd_update a) = (l',r')" by force with res False have res: "select_correct_factor_main (bnd_update a) old [] l' r' 0 = ((q, ri_fin), l_fin, r_fin)" by auto have at': "at_step f (Suc i) (bnd_update a)" unfolding at_step_def proof (intro allI, goal_cases) case (1 n) have id: "(bnd_update ^^ Suc n) a = (bnd_update ^^ n) (bnd_update a)" by (induct n, auto) from at[unfolded at_step_def, rule_format, of "Suc n"] show ?case unfolding id by simp qed from 0[unfolded root_info_condD(1)[OF ri' lr]] obtain y1 where y1: "root_cond (q1,l,r) y1" by (cases "Collect (root_cond (q1, l, r)) = {}", auto) from n1[unfolded n old'] have "?ri > 1 \ sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" by (cases "sum_list (map (\ (q,ri). root_info.l_r ri l r) old')", auto) hence "\ q2 ri2 y2. (q2,ri2) \ set old \ root_cond (q2,l,r) y2 \ y1 \ y2" proof assume "?ri > 1" with root_info_condD(1)[OF ri' lr] have "card {x. root_cond (q1, l, r) x} > 1" by simp from card_gt_1D[OF this] y1 obtain y2 where "root_cond (q1,l,r) y2" and "y1 \ y2" by auto thus ?thesis unfolding old' by auto next assume "sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" then obtain q2 ri2 where mem: "(q2,ri2) \ set old'" and ri2: "root_info.l_r ri2 l r \ 0" by auto with q0 ri have "root_info_cond ri2 q2" unfolding old' by auto from ri2[unfolded root_info_condD(1)[OF this lr]] obtain y2 where y2: "root_cond (q2,l,r) y2" by (cases "Collect (root_cond (q2, l, r)) = {}", auto) from dist[unfolded old'] split_list[OF mem] have diff: "q1 \ q2" by auto from y1 have q1: "q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" unfolding old' root_cond_def by auto from y2 have q2: "q2 \ fst ` set todo \ fst ` set old \ ipoly q2 y2 = 0" unfolding old' root_cond_def using mem by force have "y1 \ y2" proof assume id: "y1 = y2" from q1 have "\ q1. q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" by blast from un[OF this] q1 q2[folded id] have "q1 = q2" by auto with diff show False by simp qed with mem y2 show ?thesis unfolding old' by auto qed then obtain q2 ri2 y2 where mem2: "(q2,ri2) \ set old" and y2: "root_cond (q2,l,r) y2" and diff: "y1 \ y2" by auto from mem2 orig have "(q1,ri1) \ orig" "(q2,ri2) \ orig" unfolding old' by auto with y1 y2 diff have "abs (y1 - y2) \ diffs" unfolding diffs_def rts_def root_cond_def by auto from Min_le[OF diffs(1) this] have "abs (y1 - y2) \ 2 * eps" unfolding eps_def by auto with eps have eps: "abs (y1 - y2) > eps" by auto from y1 y2 have l: "of_rat l \ min y1 y2" unfolding root_cond_def by auto from y1 y2 have r: "of_rat r \ max y1 y2" unfolding root_cond_def by auto from l r eps have eps: "of_rat r - of_rat l > eps" by auto have "i < N" proof (rule ccontr) assume "\ i < N" hence "\ k. i = N + k" by presburger then obtain k where i: "i = N + k" by auto { fix k l r assume "f (N + k) = (l,r)" hence "of_rat r - of_rat l \ eps" proof (induct k arbitrary: l r) case 0 with N show ?case by auto next case (Suc k l r) obtain l' r' where f: "f (N + k) = (l',r')" by force from Suc(1)[OF this] have IH: "?r r' - ?r l' \ eps" by auto from f Suc(2) conv[THEN conjunct1, rule_format, of "N + k"] have "?r l \ ?r l'" "?r r \ ?r r'" by (auto simp: of_rat_less_eq) thus ?case using IH by auto qed } note * = this from at[unfolded at_step_def i, rule_format, of 0] bnd have "f (N + k) = (l,r)" by auto from *[OF this] eps show False by auto qed hence rel: "((old, Suc i), pair) \ rel" unfolding pair rel_def by auto from dist have dist: "distinct (map fst (old @ []))" unfolding Nil by auto have id: "set todo \ set old = set old \ set []" unfolding Nil by auto show ?thesis unfolding id proof (rule IH[OF rel at' res bnd' ri _ _ dist _ _ _ refl], goal_cases) case 2 thus ?case using q0 by auto qed (insert ex un orig Nil, auto) next case True with res old' have id: "q = q1" "ri_fin = ri1" "l_fin = l" "r_fin = r" by auto from n[unfolded True old'] 0 have 1: "?ri = 1" by (cases ?ri; cases "?ri - 1", auto) from root_info_condD(1)[OF ri' lr] 1 have "card {x. root_cond (q1,l,r) x} = 1" by auto from card_1_Collect_ex1[OF this] have unique: "unique_root (q1,l,r)" . from ex[unfolded Nil old'] consider (A) "ipoly q1 x = 0" | (B) q where "q \ fst ` set old'" "ipoly q x = 0" by auto hence "x = the_unique_root (q1,l,r)" proof (cases) case A with lxr have "root_cond (q1,l,r) x" unfolding root_cond_def by auto from the_unique_root_eqI[OF unique this] show ?thesis by simp next case (B q) with lxr have "root_cond (q,l,r) x" unfolding root_cond_def by auto hence empty: "{x. root_cond (q,l,r) x} \ {}" by auto from B(1) obtain ri' where mem: "(q,ri') \ set old'" by force from q0[unfolded old'] mem have q0: "q \ 0" by auto from finite_ipoly_roots[OF this] have "finite {x. root_cond (q,l,r) x}" unfolding root_cond_def by auto with empty have card: "card {x. root_cond (q,l,r) x} \ 0" by simp from ri[unfolded old'] mem have "root_info_cond ri' q" by auto from root_info_condD(1)[OF this lr] card have "root_info.l_r ri' l r \ 0" by auto with n[unfolded True old'] 1 split_list[OF mem] have False by auto thus ?thesis by simp qed thus ?thesis unfolding id using unique ri' unfolding old' by auto qed qed qed qed lemma select_correct_factor: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and res: "select_correct_factor init polys = ((q,ri),(l,r))" and ri: "\ q ri. (q,ri) \ set polys \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set polys \ q \ 0" and ex: "\q. q \ fst ` set polys \ ipoly q x = 0" and dist: "distinct (map fst polys)" and un: "\ x :: real. (\q. q \ fst ` set polys \ ipoly q x = 0) \ \!q. q \ fst ` set polys \ ipoly q x = 0" shows "unique_root (q,l,r) \ (q,ri) \ set polys \ x = the_unique_root (q,l,r)" proof - obtain l' r' where init: "bnd_get init = (l',r')" by force from res[unfolded select_correct_factor_def init split] have res: "select_correct_factor_main init polys [] l' r' 0 = ((q, ri), l, r)" by auto have at: "at_step (\ i. bnd_get ((bnd_update ^^ i) init)) 0 init" unfolding at_step_def by auto have "unique_root (q,l,r) \ (q,ri) \ set polys \ set [] \ x = the_unique_root (q,l,r)" by (rule select_correct_factor_main[OF conv at res init ri], insert dist un ex q0, auto) thus ?thesis by auto qed definition real_alg_2' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where [code del]: "real_alg_2' ri p l r = ( if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else real_alg_2_main ri (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l',r',sr') \ (p, l', r')))" lemma real_alg_2'_code[code]: "real_alg_2' ri p l r = (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" unfolding real_alg_2'_def real_alg_2_main_def by (cases "tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r))", simp add: Let_def) definition real_alg_2'' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where "real_alg_2'' ri p l r = (case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" lemma real_alg_2'': "degree p \ 1 \ real_alg_2'' ri p l r = real_alg_2' ri p l r" unfolding real_alg_2'_code real_alg_2''_def by auto lemma poly_cond_degree_0_imp_no_root: fixes x :: "'b :: {comm_ring_1,ring_char_0}" assumes pc: "poly_cond p" and deg: "degree p = 0" shows "ipoly p x \ 0" proof from pc have "p \ 0" by auto moreover assume "ipoly p x = 0" note poly_zero[OF this] ultimately show False using deg by auto qed lemma real_alg_2': assumes ur: "unique_root (q,l,r)" and pc: "poly_cond q" and ri: "root_info_cond ri q" shows "invariant_2 (real_alg_2' ri q l r) \ real_of_2 (real_alg_2' ri q l r) = the_unique_root (q,l,r)" (is "_ \ _ = ?x") proof (cases "degree q" "Suc 0" rule: linorder_cases) case deg: less then have "degree q = 0" by auto from poly_cond_degree_0_imp_no_root[OF pc this] ur have False by force then show ?thesis by auto next case deg: equal hence id: "real_alg_2' ri q l r = Rational (Rat.Fract (- coeff q 0) (coeff q 1))" unfolding real_alg_2'_def by auto show ?thesis unfolding id using degree_1_ipoly[OF deg] using unique_rootD(4)[OF ur] by auto next case deg: greater with pc have pc2: "poly_cond2 q" by auto let ?rai = "real_alg_2' ri q l r" let ?r = real_of_rat obtain l' r' sr' where tight: "tighten_poly_bounds_for_x q 0 l r (sgn (ipoly q r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?rai' = "(q, l', r')" have rai': "?rai = real_alg_2_main ri ?rai'" unfolding real_alg_2'_def using deg tight by auto hence rai: "real_of_1 ?rai' = the_unique_root (q,l',r')" by auto note tight = tighten_poly_bounds_for_x[OF ur pc2 tight refl] let ?x = "the_unique_root (q, l, r)" from tight have tight: "root_cond (q,l',r') ?x" "l \ l'" "l' \ r'" "r' \ r" "l' > 0 \ r' < 0" by auto from unique_root_sub_interval[OF ur tight(1) tight(2,4)] poly_condD[OF pc] have ur': "unique_root (q, l', r')" and x: "?x = the_unique_root (q,l',r')" by auto from tight(2-) have sgn: "sgn l' = sgn r'" by auto show ?thesis unfolding rai' using real_alg_2_main[of ?rai' ri] invariant_1_realI[of ?rai' ?x] by (auto simp: tight(1) sgn pc ri ur') qed definition select_correct_factor_int_poly :: "'a \ int poly \ real_alg_2" where "select_correct_factor_int_poly init p \ let qs = factors_of_int_poly p; polys = map (\ q. (q, root_info q)) qs; ((q,ri),(l,r)) = select_correct_factor init polys in real_alg_2' ri q l r" lemma select_correct_factor_int_poly: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and rai: "select_correct_factor_int_poly init p = rai" and x: "ipoly p x = 0" and p: "p \ 0" shows "invariant_2 rai \ real_of_2 rai = x" proof - obtain qs where fact: "factors_of_int_poly p = qs" by auto define polys where "polys = map (\ q. (q, root_info q)) qs" obtain q ri l r where res: "select_correct_factor init polys = ((q,ri),(l,r))" by (cases "select_correct_factor init polys", auto) have fst: "map fst polys = qs" "fst ` set polys = set qs" unfolding polys_def map_map o_def by force+ note fact' = factors_of_int_poly[OF fact] note rai = rai[unfolded select_correct_factor_int_poly_def Let_def fact, folded polys_def, unfolded res split] from fact' fst have dist: "distinct (map fst polys)" by auto from fact'(2)[OF p, of x] x fst have ex: "\q. q \ fst ` set polys \ ipoly q x = 0" by auto { fix q ri assume "(q,ri) \ set polys" hence ri: "ri = root_info q" and q: "q \ set qs" unfolding polys_def by auto from fact'(1)[OF q] have *: "lead_coeff q > 0" "irreducible q" "degree q > 0" by auto from * have q0: "q \ 0" by auto from root_info[OF *(2-3)] ri have ri: "root_info_cond ri q" by auto note ri q0 * } note polys = this have "unique_root (q, l, r) \ (q, ri) \ set polys \ x = the_unique_root (q, l, r)" by (rule select_correct_factor[OF conv res polys(1) _ ex dist, unfolded fst, OF _ _ fact'(3)[OF p]], insert fact'(2)[OF p] polys(2), auto) hence ur: "unique_root (q,l,r)" and mem: "(q,ri) \ set polys" and x: "x = the_unique_root (q,l,r)" by auto note polys = polys[OF mem] from polys(3-4) have ty: "poly_cond q" by (simp add: poly_cond_def) show ?thesis unfolding x rai[symmetric] by (intro real_alg_2' ur ty polys(1)) qed end (* ********************* *) subsubsection\Addition\ lemma ipoly_0_0[simp]: "ipoly f (0::'a::{comm_ring_1,ring_char_0}) = 0 \ poly f 0 = 0" unfolding poly_0_coeff_0 by simp lemma add_rat_roots_below[simp]: "roots_below (poly_add_rat r p) x = (\y. y + of_rat r) ` roots_below p (x - of_rat r)" proof (unfold add_rat_roots image_def, intro Collect_eqI, goal_cases) case (1 y) then show ?case by (auto intro: exI[of _ "y - real_of_rat r"]) qed lemma add_rat_root_cond: shows "root_cond (cf_pos_poly (poly_add_rat m p),l,r) x = root_cond (p, l - m, r - m) (x - of_rat m)" by (unfold root_cond_def, auto simp add: add_rat_roots hom_distribs) lemma add_rat_unique_root: "unique_root (cf_pos_poly (poly_add_rat m p), l, r) = unique_root (p, l-m, r-m)" by (auto simp: add_rat_root_cond) fun add_rat_1 :: "rat \ real_alg_1 \ real_alg_1" where "add_rat_1 r1 (p2,l2,r2) = ( let p = cf_pos_poly (poly_add_rat r1 p2); (l,r,sr) = tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) in (p,l,r))" lemma poly_real_alg_1_add_rat[simp]: "poly_real_alg_1 (add_rat_1 r y) = cf_pos_poly (poly_add_rat r (poly_real_alg_1 y))" by (cases y, auto simp: Let_def split: prod.split) lemma sgn_cf_pos: assumes "lead_coeff p > 0" shows "sgn (ipoly (cf_pos_poly p) (x::'a::linordered_field)) = sgn (ipoly p x)" proof (cases "p = 0") case True with assms show ?thesis by auto next case False from cf_pos_poly_main False obtain d where p': "Polynomial.smult d (cf_pos_poly p) = p" by auto have "d > 0" proof (rule zero_less_mult_pos2) from False assms have "0 < lead_coeff p" by (auto simp: cf_pos_def) also from p' have "\ = d * lead_coeff (cf_pos_poly p)" by (metis lead_coeff_smult) finally show "0 < \". show "lead_coeff (cf_pos_poly p) > 0" using False by (unfold lead_coeff_cf_pos_poly) qed moreover from p' have "ipoly p x = of_int d * ipoly (cf_pos_poly p) x" by (fold poly_smult of_int_hom.map_poly_hom_smult, auto) ultimately show ?thesis by (auto simp: sgn_mult[where 'a='a]) qed lemma add_rat_1: fixes r1 :: rat assumes inv_y: "invariant_1_2 y" defines "z \ add_rat_1 r1 y" shows "invariant_1_2 z \ (real_of_1 z = of_rat r1 + real_of_1 y)" proof (cases y) case y_def: (fields p2 l2 r2) define p where "p \ cf_pos_poly (poly_add_rat r1 p2)" obtain l r sr where lr: "tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) = (l,r,sr)" by (metis surj_pair) from lr have z: "z = (p,l,r)" by (auto simp: y_def z_def p_def Let_def) from inv_y have ur: "unique_root (p, l2 + r1, r2 + r1)" by (auto simp: p_def add_rat_root_cond y_def add_rat_unique_root) from inv_y[unfolded y_def invariant_1_2_def,simplified] have pc2: "poly_cond2 p" unfolding p_def apply (intro poly_cond2I poly_add_rat_irreducible poly_condI, unfold lead_coeff_cf_pos_poly) apply (auto elim!: invariant_1E) done note main = tighten_poly_bounds_for_x[OF ur pc2 lr refl, simplified] then have "sgn l = sgn r" unfolding sgn_if apply simp apply linarith done from invariant_1_2_realI[OF main(4) _ main(7), simplified, OF this pc2] main(1-3) ur show ?thesis by (auto simp: z p_def y_def add_rat_root_cond ex1_the_shift) qed fun tighten_poly_bounds_binary :: "int poly \ int poly \ (rat \ rat \ rat) \ rat \ rat \ rat \ (rat \ rat \ rat) \ rat \ rat \ rat" where "tighten_poly_bounds_binary cr1 cr2 ((l1,r1,sr1),(l2,r2,sr2)) = (tighten_poly_bounds cr1 l1 r1 sr1, tighten_poly_bounds cr2 l2 r2 sr2)" lemma tighten_poly_bounds_binary: assumes ur: "unique_root (p1,l1,r1)" "unique_root (p2,l2,r2)" and pt: "poly_cond2 p1" "poly_cond2 p2" defines "x \ the_unique_root (p1,l1,r1)" and "y \ the_unique_root (p2,l2,r2)" assumes bnd: "\ l1 r1 l2 r2 l r sr1 sr2. I l1 \ I l2 \ root_cond (p1,l1,r1) x \ root_cond (p2,l2,r2) y \ bnd ((l1,r1,sr1),(l2,r2,sr2)) = (l,r) \ of_rat l \ f x y \ f x y \ of_rat r" and approx: "\ l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' sr1 sr2 sr1' sr2'. I l1 \ I l2 \ l1 \ r1 \ l2 \ r2 \ (l,r) = bnd ((l1,r1,sr1), (l2,r2,sr2)) \ (l',r') = bnd ((l1',r1',sr1'), (l2',r2',sr2')) \ (l1',r1') \ {(l1,(l1+r1)/2),((l1+r1)/2,r1)} \ (l2',r2') \ {(l2,(l2+r2)/2),((l2+r2)/2,r2)} \ (r' - l') \ 3/4 * (r - l) \ l \ l' \ r' \ r" and I_mono: "\ l l'. I l \ l \ l' \ I l'" and I: "I l1" "I l2" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "converges_to (\ i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sr1),(l2,r2,sr2)))) (f x y)" proof - let ?upd = "tighten_poly_bounds_binary p1 p2" define upd where "upd = ?upd" define init where "init = ((l1, r1, sr1), l2, r2, sr2)" let ?g = "(\i. bnd ((upd ^^ i) init))" obtain l r where bnd_init: "bnd init = (l,r)" by force note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from ur1(4) ur2(4) x_def y_def have rc1: "root_cond (p1,l1,r1) x" and rc2: "root_cond (p2,l2,r2) y" by auto define g where "g = ?g" { fix i L1 R1 L2 R2 L R j SR1 SR2 assume "((upd ^^ i)) init = ((L1,R1,SR1),(L2,R2,SR2))" "g i = (L,R)" hence "I L1 \ I L2 \ root_cond (p1,L1,R1) x \ root_cond (p2,L2,R2) y \ unique_root (p1, L1, R1) \ unique_root (p2, L2, R2) \ in_interval (L,R) (f x y) \ (i = Suc j \ sub_interval (g i) (g j) \ (R - L \ 3/4 * (snd (g j) - fst (g j)))) \ SR1 = sgn (ipoly p1 R1) \ SR2 = sgn (ipoly p2 R2)" proof (induct i arbitrary: L1 R1 L2 R2 L R j SR1 SR2) case 0 thus ?case using I rc1 rc2 ur bnd[of l1 l2 r1 r2 sr1 sr2 L R] g_def sr unfolding init_def by auto next case (Suc i) obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bndi: "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force hence gi: "g i = (l,r)" using updi unfolding g_def by auto have "(upd ^^ Suc i) init = upd ((l1, r1, sr1), l2, r2, sr2)" using updi by simp from Suc(2)[unfolded this] have upd: "upd ((l1, r1, sr1), l2, r2, sr2) = ((L1, R1, SR1), L2, R2, SR2)" . from upd updi Suc(3) have bndsi: "bnd ((L1, R1, SR1), L2, R2, SR2) = (L,R)" by (auto simp: g_def) from Suc(1)[OF updi gi] have I: "I l1" "I l2" and rc: "root_cond (p1,l1,r1) x" "root_cond (p2,l2,r2) y" and ur: "unique_root (p1, l1, r1)" "unique_root (p2, l2, r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" by auto from upd[unfolded upd_def] have tight: "tighten_poly_bounds p1 l1 r1 sr1 = (L1, R1, SR1)" "tighten_poly_bounds p2 l2 r2 sr2 = (L2, R2, SR2)" by auto note tight1 = tighten_poly_bounds[OF tight(1) ur(1) pt(1) sr(1)] note tight2 = tighten_poly_bounds[OF tight(2) ur(2) pt(2) sr(2)] from tight1 have lr1: "l1 \ r1" by auto from tight2 have lr2: "l2 \ r2" by auto note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from tight1 I_mono[OF I(1)] have I1: "I L1" by auto from tight2 I_mono[OF I(2)] have I2: "I L2" by auto note ur1 = unique_root_sub_interval[OF ur(1) tight1(1,2,4)] note ur2 = unique_root_sub_interval[OF ur(2) tight2(1,2,4)] from rc(1) ur ur1 have x: "x = the_unique_root (p1,L1,R1)" by (auto intro!:the_unique_root_eqI) from rc(2) ur ur2 have y: "y = the_unique_root (p2,L2,R2)" by (auto intro!:the_unique_root_eqI) from unique_rootD[OF ur1(1)] x have x: "root_cond (p1,L1,R1) x" by auto from unique_rootD[OF ur2(1)] y have y: "root_cond (p2,L2,R2) y" by auto from tight(1) have half1: "(L1, R1) \ {(l1, (l1 + r1) / 2), ((l1 + r1) / 2, r1)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from tight(2) have half2: "(L2, R2) \ {(l2, (l2 + r2) / 2), ((l2 + r2) / 2, r2)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from approx[OF I lr1 lr2 bndi[symmetric] bndsi[symmetric] half1 half2] have "R - L \ 3 / 4 * (r - l) \ l \ L \ R \ r" . hence "sub_interval (g (Suc i)) (g i)" "R - L \ 3/4 * (snd (g i) - fst (g i))" unfolding gi Suc(3) by auto with bnd[OF I1 I2 x y bndsi] show ?case using I1 I2 x y ur1 ur2 tight1(6) tight2(6) by auto qed } note invariants = this define L where "L = (\ i. fst (g i))" define R where "R = (\ i. snd (g i))" { fix i obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bnd': "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force have gi: "g i = (l,r)" unfolding g_def updi bnd' by auto hence id: "l = L i" "r = R i" unfolding L_def R_def by auto from invariants[OF updi gi[unfolded id]] have "in_interval (L i, R i) (f x y)" "\ j. i = Suc j \ sub_interval (g i) (g j) \ R i - L i \ 3 / 4 * (R j - L j)" unfolding L_def R_def by auto } note * = this { fix i from *(1)[of i] *(2)[of "Suc i", OF refl] have "in_interval (g i) (f x y)" "sub_interval (g (Suc i)) (g i)" "R (Suc i) - L (Suc i) \ 3 / 4 * (R i - L i)" unfolding L_def R_def by auto } note * = this show ?thesis unfolding upd_def[symmetric] init_def[symmetric] g_def[symmetric] unfolding converges_to_def proof (intro conjI allI impI, rule *(1), rule *(2)) fix eps :: real assume eps: "0 < eps" let ?r = real_of_rat define r where "r = (\ n. ?r (R n))" define l where "l = (\ n. ?r (L n))" define diff where "diff = (\ n. r n - l n)" { fix n from *(3)[of n] have "?r (R (Suc n) - L (Suc n)) \ ?r (3 / 4 * (R n - L n))" unfolding of_rat_less_eq by simp also have "?r (R (Suc n) - L (Suc n)) = (r (Suc n) - l (Suc n))" unfolding of_rat_diff r_def l_def by simp also have "?r (3 / 4 * (R n - L n)) = 3 / 4 * (r n - l n)" unfolding r_def l_def by (simp add: hom_distribs) finally have "diff (Suc n) \ 3 / 4 * diff n" unfolding diff_def . } note * = this { fix i have "diff i \ (3/4)^i * diff 0" proof (induct i) case (Suc i) from Suc *[of i] show ?case by auto qed auto } then obtain c where *: "\ i. diff i \ (3/4)^i * c" by auto have "\ n. diff n \ eps" proof (cases "c \ 0") case True with *[of 0] eps show ?thesis by (intro exI[of _ 0], auto) next case False hence c: "c > 0" by auto with eps have "inverse c * eps > 0" by auto from exp_tends_to_zero[of "3/4 :: real", OF _ _ this] obtain n where "(3/4) ^ n \ inverse c * eps" by auto from mult_right_mono[OF this, of c] c have "(3/4) ^ n * c \ eps" by (auto simp: field_simps) with *[of n] show ?thesis by (intro exI[of _ n], auto) qed then obtain n where "?r (R n) - ?r (L n) \ eps" unfolding l_def r_def diff_def by blast thus "\n l r. g n = (l, r) \ ?r r - ?r l \ eps" unfolding L_def R_def by (intro exI[of _ n], force) qed qed fun add_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "add_1 (p1,l1,r1) (p2,l2,r2) = ( select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 + l2, r1 + r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2))" lemma add_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ add_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x + real_of_1 y)" proof (cases x) case xt: (fields p1 l1 r1) show ?thesis proof (cases y) case yt: (fields p2 l2 r2) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_add p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!: invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!: invariant_1E) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 + l2, r1 + r2))" define bnd where "bnd = ?bnd" have "invariant_2 z \ real_of_2 z = ?x + ?y" proof (intro select_correct_factor_int_poly) from represents_add[OF ax ay] show "?p \ 0" "ipoly ?p (?x + ?y) = 0" by auto from z[unfolded xt yt] show sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2) = z" by (auto simp: bnd_def) have ur1: "unique_root (p1,l1,r1)" "poly_cond2 p1" using x by auto have ur2: "unique_root (p2,l2,r2)" "poly_cond2 p2" using y by auto show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x + ?y)" by (intro tighten_poly_bounds_binary ur1 ur2; force simp: bnd_def hom_distribs) qed thus ?thesis unfolding xt yt . qed qed declare add_rat_1.simps[simp del] declare add_1.simps[simp del] (* ********************* *) subsubsection\Multiplication\ context begin private fun mult_rat_1_pos :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1_pos r1 (p2,l2,r2) = real_alg_2 (cf_pos_poly (poly_mult_rat r1 p2), l2*r1, r2*r1)" private fun mult_1_pos :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1_pos (p1,l1,r1) (p2,l2,r2) = select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 * l2, r1 * r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_mult p1 p2)" fun mult_rat_1 :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1 x y = (if x < 0 then uminus_2 (mult_rat_1_pos (-x) y) else if x = 0 then Rational 0 else (mult_rat_1_pos x y))" fun mult_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1 x y = (case (x,y) of ((p1,l1,r1),(p2,l2,r2)) \ if r1 > 0 then if r2 > 0 then mult_1_pos x y else uminus_2 (mult_1_pos x (uminus_1 y)) else if r2 > 0 then uminus_2 (mult_1_pos (uminus_1 x) y) else mult_1_pos (uminus_1 x) (uminus_1 y))" lemma mult_rat_1_pos: fixes r1 :: rat assumes r1: "r1 > 0" and y: "invariant_1 y" defines z: "z \ mult_rat_1_pos r1 y" shows "invariant_2 z \ (real_of_2 z = of_rat r1 * real_of_1 y)" proof - obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_rat r1" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_mult_rat r1 p2" let ?mp = "cf_pos_poly ?p" note y = y[unfolded yt] note yD = invariant_1D[OF y] from yD r1 have p: "?p \ 0" and r10: "r1 \ 0" by auto hence mp: "?mp \ 0" by simp from yD(1) have rt: "ipoly p2 ?y = 0" and bnd: "of_rat l2 \ ?y" "?y \ of_rat r2" by auto from rt r1 have rt: "ipoly ?mp (?x * ?y) = 0" by (auto simp add: field_simps ipoly_mult_rat[OF r10]) from yD(5) have irr: "irreducible p2" unfolding represents_def using y unfolding root_cond_def split by auto from poly_mult_rat_irreducible[OF this _ r10] yD have irr: "irreducible ?mp" by simp from p have mon: "cf_pos ?mp" by auto obtain l r where lr: "l = l2 * r1" "r = r2 * r1" by force from bnd r1 have bnd: "of_rat l \ ?x * ?y" "?x * ?y \ of_rat r" unfolding lr of_rat_mult by auto with rt have rc: "root_cond (?mp,l,r) (?x * ?y)" unfolding root_cond_def by auto have ur: "unique_root (?mp,l,r)" proof (rule ex1I, rule rc) fix z assume "root_cond (?mp,l,r) z" from this[unfolded root_cond_def split] have bndz: "of_rat l \ z" "z \ of_rat r" and rt: "ipoly ?mp z = 0" by auto have "fst (quotient_of r1) \ 0" using quotient_of_div[of r1] r10 by (cases "quotient_of r1", auto) with rt have rt: "ipoly p2 (z * inverse ?x) = 0" by (auto simp: ipoly_mult_rat[OF r10]) from bndz r1 have "of_rat l2 \ z * inverse ?x" "z * inverse ?x \ of_rat r2" unfolding lr of_rat_mult by (auto simp: field_simps) with rt have "root_cond (p2,l2,r2) (z * inverse ?x)" unfolding root_cond_def by auto also note invariant_1_root_cond[OF y] finally have "?y = z * inverse ?x" by auto thus "z = ?x * ?y" using r1 by auto qed from r1 have sgnr: "sgn r = sgn r2" unfolding lr by (cases "r2 = 0"; cases "r2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from r1 have sgnl: "sgn l = sgn l2" unfolding lr by (cases "l2 = 0"; cases "l2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from the_unique_root_eqI[OF ur rc] have xy: "?x * ?y = the_unique_root (?mp,l,r)" by auto from z[unfolded yt, simplified, unfolded Let_def lr[symmetric] split] have z: "z = real_alg_2 (?mp, l, r)" by simp have yp2: "p2 represents ?y" using yD unfolding root_cond_def split represents_def by auto with irr mon have pc: "poly_cond ?mp" by (auto simp: poly_cond_def cf_pos_def) have rc: "invariant_1 (?mp, l, r)" unfolding z using yD(2) pc ur by (auto simp add: invariant_1_def ur mp sgnr sgnl) show ?thesis unfolding z using real_alg_2[OF rc] unfolding yt xy unfolding z by simp qed lemma mult_1_pos: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ mult_1_pos x y" assumes pos: "real_of_1 x > 0" "real_of_1 y > 0" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt: "x = (p1,l1,r1)" by (cases x, auto) obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?r = "real_of_rat" let ?p = "poly_mult p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x y have basic: "unique_root (p1, l1, r1)" "poly_cond2 p1" "unique_root (p2, l2, r2)" "poly_cond2 p2" by auto from basic have irr1: "irreducible p1" and irr2: "irreducible p2" by auto from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!:invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!:invariant_1E) from ax ay pos[unfolded xt yt] have axy: "?p represents (?x * ?y)" by (intro represents_mult represents_irr_non_0[OF irr2], auto) from representsD[OF this] have p: "?p \ 0" and rt: "ipoly ?p (?x * ?y) = 0" . from x pos(1)[unfolded xt] have "?r r1 > 0" unfolding split by auto hence "sgn r1 = 1" unfolding sgn_rat_def by (auto split: if_splits) with x have "sgn l1 = 1" by auto hence l1_pos: "l1 > 0" unfolding sgn_rat_def by (cases "l1 = 0"; cases "l1 < 0"; auto) from y pos(2)[unfolded yt] have "?r r2 > 0" unfolding split by auto hence "sgn r2 = 1" unfolding sgn_rat_def by (auto split: if_splits) with y have "sgn l2 = 1" by auto hence l2_pos: "l2 > 0" unfolding sgn_rat_def by (cases "l2 = 0"; cases "l2 < 0"; auto) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 * l2, r1 * r2))" define bnd where "bnd = ?bnd" obtain z' where sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) ?p = z'" by auto have main: "invariant_2 z' \ real_of_2 z' = ?x * ?y" proof (rule select_correct_factor_int_poly[OF _ sel rt p]) { fix l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' :: rat let ?m1 = "(l1+r1)/2" let ?m2 = "(l2+r2)/2" define d1 where "d1 = r1 - l1" define d2 where "d2 = r2 - l2" let ?M1 = "l1 + d1/2" let ?M2 = "l2 + d2/2" assume le: "l1 > 0" "l2 > 0" "l1 \ r1" "l2 \ r2" and id: "(l, r) = (l1 * l2, r1 * r2)" "(l', r') = (l1' * l2', r1' * r2')" and mem: "(l1', r1') \ {(l1, ?m1), (?m1, r1)}" "(l2', r2') \ {(l2, ?m2), (?m2, r2)}" hence id: "l = l1 * l2" "r = (l1 + d1) * (l2 + d2)" "l' = l1' * l2'" "r' = r1' * r2'" "r1 = l1 + d1" "r2 = l2 + d2" and id': "?m1 = ?M1" "?m2 = ?M2" unfolding d1_def d2_def by (auto simp: field_simps) define l1d1 where "l1d1 = l1 + d1" from le have ge0: "d1 \ 0" "d2 \ 0" "l1 \ 0" "l2 \ 0" unfolding d1_def d2_def by auto have "4 * (r' - l') \ 3 * (r - l)" proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 1 = this show ?thesis proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = ?M2 \ r2' = r2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = ?M2" "r2' = r2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 2 = this show ?thesis proof (cases "l1' = ?M1 \ r1' = r1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = ?M1" "r1' = r1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 3 = this from 1 2 3 mem have id2: "l1' = ?M1" "r1' = r1" "l2' = ?M2" "r2' = r2" unfolding id' by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp qed qed qed hence "r' - l' \ 3 / 4 * (r - l)" by simp } note decr = this show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x * ?y)" proof (intro tighten_poly_bounds_binary[where f = "(*)" and I = "\ l. l > 0"] basic l1_pos l2_pos, goal_cases) case (1 L1 R1 L2 R2 L R) hence "L = L1 * L2" "R = R1 * R2" unfolding bnd_def by auto hence id: "?r L = ?r L1 * ?r L2" "?r R = ?r R1 * ?r R2" by (auto simp: hom_distribs) from 1(3-4) have le: "?r L1 \ ?x" "?x \ ?r R1" "?r L2 \ ?y" "?y \ ?r R2" unfolding root_cond_def by auto from 1(1-2) have lt: "0 < ?r L1" "0 < ?r L2" by auto from mult_mono[OF le(1,3), folded id] lt le have L: "?r L \ ?x * ?y" by linarith have R: "?x * ?y \ ?r R" by (rule mult_mono[OF le(2,4), folded id], insert lt le, linarith+) show ?case using L R by blast next case (2 l1 r1 l2 r2 l1' r1' l2' r2' l l' r r') from 2(5-6) have lr: "l = l1 * l2" "r = r1 * r2" "l' = l1' * l2'" "r' = r1' * r2'" unfolding bnd_def by auto from 2(1-4) have le: "0 < l1" "0 < l2" "l1 \ r1" "l2 \ r2" by auto from 2(7-8) le have le': "l1 \ l1'" "r1' \ r1" "l2 \ l2'" "r2' \ r2" "0 < r2'" "0 < r2" by auto from mult_mono[OF le'(1,3), folded lr] le le' have l: "l \ l'" by auto have r: "r' \ r" by (rule mult_mono[OF le'(2,4), folded lr], insert le le', linarith+) have "r' - l' \ 3 / 4 * (r - l)" by (rule decr[OF _ _ _ _ _ _ 2(7-8)], insert le le' lr, auto) thus ?case using l r by blast qed auto qed have z': "z' = z" unfolding z[unfolded xt yt, simplified, unfolded bnd_def[symmetric] sel] by auto from main[unfolded this] show ?thesis unfolding xt yt by simp qed lemma mult_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z[simp]: "z \ mult_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt[simp]: "x = (p1,l1,r1)" by (cases x) obtain p2 l2 r2 where yt[simp]: "y = (p2,l2,r2)" by (cases y) let ?xt = "(p1, l1, r1)" let ?yt = "(p2, l2, r2)" let ?x = "real_of_1 ?xt" let ?y = "real_of_1 ?yt" let ?mxt = "uminus_1 ?xt" let ?myt = "uminus_1 ?yt" let ?mx = "real_of_1 ?mxt" let ?my = "real_of_1 ?myt" let ?r = "real_of_rat" from invariant_1_2_of_rat[OF x, of 0] have x0: "?x < 0 \ ?x > 0" by auto from invariant_1_2_of_rat[OF y, of 0] have y0: "?y < 0 \ ?y > 0" by auto from uminus_1_2[OF x] have mx: "invariant_1_2 ?mxt" and [simp]: "?mx = - ?x" by auto from uminus_1_2[OF y] have my: "invariant_1_2 ?myt" and [simp]: "?my = - ?y" by auto have id: "r1 > 0 \ ?x > 0" "r1 < 0 \ ?x < 0" "r2 > 0 \ ?y > 0" "r2 < 0 \ ?y < 0" using x y by auto show ?thesis proof (cases "?x > 0") case x0: True show ?thesis proof (cases "?y > 0") case y0: True with x y x0 mult_1_pos[OF x y] show ?thesis by auto next case False with y0 have y0: "?y < 0" by auto with x0 have z: "z = uminus_2 (mult_1_pos ?xt ?myt)" unfolding z xt yt mult_1.simps split id by simp from x0 y0 mult_1_pos[OF x my] uminus_2[of "mult_1_pos ?xt ?myt"] show ?thesis unfolding z by simp qed next case False with x0 have x0: "?x0 < 0" by simp show ?thesis proof (cases "?y > 0") case y0: True with x0 x y id have z: "z = uminus_2 (mult_1_pos ?mxt ?yt)" by simp from x0 y0 mult_1_pos[OF mx y] uminus_2[of "mult_1_pos ?mxt ?yt"] show ?thesis unfolding z by auto next case False with y0 have y0: "?y < 0" by simp with x0 x y have z: "z = mult_1_pos ?mxt ?myt" by auto with x0 y0 x y mult_1_pos[OF mx my] show ?thesis unfolding z by auto qed qed qed lemma mult_rat_1: fixes x assumes y: "invariant_1 y" defines z: "z \ mult_rat_1 x y" shows "invariant_2 z \ (real_of_2 z = of_rat x * real_of_1 y)" proof (cases y) case yt: (fields p2 l2 r2) let ?yt = "(p2, l2, r2)" let ?x = "real_of_rat x" let ?y = "real_of_1 ?yt" let ?myt = "mult_rat_1_pos (- x) ?yt" note y = y[unfolded yt] note z = z[unfolded yt] show ?thesis proof(cases x "0::rat" rule:linorder_cases) case x: greater with z have z: "z = mult_rat_1_pos x ?yt" by simp from mult_rat_1_pos[OF x y] show ?thesis unfolding yt z by auto next case less then have x: "- x > 0" by auto hence z: "z = uminus_2 ?myt" unfolding z by simp from mult_rat_1_pos[OF x y] have rc: "invariant_2 ?myt" and rr: "real_of_2 ?myt = - ?x * ?y" by (auto simp: hom_distribs) from uminus_2[OF rc] rr show ?thesis unfolding z[symmetric] unfolding yt[symmetric] by simp qed (auto simp: z) qed end declare mult_1.simps[simp del] declare mult_rat_1.simps[simp del] (* ********************* *) subsubsection\Root\ definition ipoly_root_delta :: "int poly \ real" where "ipoly_root_delta p = Min (insert 1 { abs (x - y) | x y. ipoly p x = 0 \ ipoly p y = 0 \ x \ y}) / 4" lemma ipoly_root_delta: assumes "p \ 0" shows "ipoly_root_delta p > 0" "2 \ card (Collect (root_cond (p, l, r))) \ ipoly_root_delta p \ real_of_rat (r - l) / 4" proof - let ?z = "0 :: real" let ?R = "{x. ipoly p x = ?z}" let ?set = "{ abs (x - y) | x y. ipoly p x = ?z \ ipoly p y = 0 \ x \ y}" define S where "S = insert 1 ?set" from finite_ipoly_roots[OF assms] have finR: "finite ?R" and fin: "finite (?R \ ?R)" by auto have "finite ?set" by (rule finite_subset[OF _ finite_imageI[OF fin, of "\ (x,y). abs (x - y)"]], force) hence fin: "finite S" and ne: "S \ {}" and pos: "\ x. x \ S \ x > 0" unfolding S_def by auto have delta: "ipoly_root_delta p = Min S / 4" unfolding ipoly_root_delta_def S_def .. have pos: "Min S > 0" using fin ne pos by auto show "ipoly_root_delta p > 0" unfolding delta using pos by auto let ?S = "Collect (root_cond (p, l, r))" assume "2 \ card ?S" hence 2: "Suc (Suc 0) \ card ?S" by simp from 2[unfolded card_le_Suc_iff[of _ ?S]] obtain x T where ST: "?S = insert x T" and xT: "x \ T" and 1: "Suc 0 \ card T" by auto from 1[unfolded card_le_Suc_iff[of _ T]] obtain y where yT: "y \ T" by auto from ST xT yT have x: "x \ ?S" and y: "y \ ?S" and xy: "x \ y" by auto hence "abs (x - y) \ S" unfolding S_def root_cond_def[abs_def] by auto with fin have "Min S \ abs (x - y)" by auto with pos have le: "Min S / 2 \ abs (x - y) / 2" by auto from x y have "abs (x - y) \ of_rat r - of_rat l" unfolding root_cond_def[abs_def] by auto also have "\ = of_rat (r - l)" by (auto simp: of_rat_diff) finally have "abs (x - y) / 2 \ of_rat (r - l) / 2" by auto with le show "ipoly_root_delta p \ real_of_rat (r - l) / 4" unfolding delta by auto qed lemma sgn_less_eq_1_rat: fixes a b :: rat shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, hide_lams) not_less one_neq_neg_one one_neq_zero order_trans sgn_rat_def) lemma sgn_less_eq_1_real: fixes a b :: real shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, hide_lams) not_less one_neq_neg_one one_neq_zero order_trans sgn_real_def) definition compare_1_rat :: "real_alg_1 \ rat \ order" where "compare_1_rat rai = (let p = poly_real_alg_1 rai in if degree p = 1 then let x = Rat.Fract (- coeff p 0) (coeff p 1) in (\ y. compare y x) else (\ y. compare_rat_1 y rai))" lemma compare_real_of_rat: "compare (real_of_rat x) (of_rat y) = compare x y" unfolding compare_rat_def compare_real_def comparator_of_def of_rat_less by auto lemma compare_1_rat: assumes rc: "invariant_1 y" shows "compare_1_rat y x = compare (of_rat x) (real_of_1 y)" proof (cases "degree (poly_real_alg_1 y)" "Suc 0" rule: linorder_cases) case less with invariant_1_degree_0[OF rc] show ?thesis by auto next case deg: greater with rc have rc: "invariant_1_2 y" by auto from deg compare_rat_1[OF rc, of x] show ?thesis unfolding compare_1_rat_def by auto next case deg: equal obtain p l r where y: "y = (p,l,r)" by (cases y) note rc = invariant_1D[OF rc[unfolded y]] from deg have p: "degree p = Suc 0" and id: "compare_1_rat y x = compare x (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding compare_1_rat_def by (auto simp: Let_def y) from rc(1)[unfolded split] have "ipoly p (real_of_1 y) = 0" unfolding y by auto with degree_1_ipoly[OF p, of "real_of_1 y"] have id': "real_of_1 y = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1))" by simp show ?thesis unfolding id id' compare_real_of_rat .. qed context fixes n :: nat begin private definition initial_lower_bound :: "rat \ rat" where "initial_lower_bound l = (if l \ 1 then l else of_int (root_rat_floor n l))" private definition initial_upper_bound :: "rat \ rat" where "initial_upper_bound r = (of_int (root_rat_ceiling n r))" context fixes cmpx :: "rat \ order" begin fun tighten_bound_root :: "rat \ rat \ rat \ rat" where "tighten_bound_root (l',r') = (let m' = (l' + r') / 2; m = m' ^ n in case cmpx m of Eq \ (m',m') | Lt \ (m',r') | Gt \ (l',m'))" lemma tighten_bound_root: assumes sgn: "sgn il = 1" "real_of_1 x \ 0" and il: "real_of_rat il \ root n (real_of_1 x)" and ir: "root n (real_of_1 x) \ real_of_rat ir" and rai: "invariant_1 x" and cmpx: "cmpx = compare_1_rat x" and n: "n \ 0" shows "converges_to (\ i. (tighten_bound_root ^^ i) (il, ir)) (root n (real_of_1 x))" (is "converges_to ?f ?x") unfolding converges_to_def proof (intro conjI impI allI) { fix x :: real have "x \ 0 \ (root n x) ^ n = x" using n by simp } note root_exp_cancel = this { fix x :: real have "x \ 0 \ root n (x ^ n) = x" using n using real_root_pos_unique by blast } note root_exp_cancel' = this from il ir have "real_of_rat il \ of_rat ir" by auto hence ir_il: "il \ ir" by (auto simp: of_rat_less_eq) from n have n': "n > 0" by auto { fix i have "in_interval (?f i) ?x \ sub_interval (?f i) (il,ir) \ (i \ 0 \ sub_interval (?f i) (?f (i - 1))) \ snd (?f i) - fst (?f i) \ (ir - il) / 2^i" proof (induct i) case 0 show ?case using il ir by auto next case (Suc i) obtain l' r' where id: "(tighten_bound_root ^^ i) (il, ir) = (l',r')" by (cases "(tighten_bound_root ^^ i) (il, ir)", auto) let ?m' = "(l' + r') / 2" let ?m = "?m' ^ n" define m where "m = ?m" note IH = Suc[unfolded id split snd_conv fst_conv] from IH have "sub_interval (l', r') (il, ir)" by auto hence ill': "il \ l'" "r' \ ir" by auto with sgn have l'0: "l' > 0" using sgn_1_pos sgn_less_eq_1_rat by blast from IH have lr'x: "in_interval (l', r') ?x" by auto hence lr'': "real_of_rat l' \ of_rat r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq . with l'0 have r'0: "r' > 0" by auto note compare = compare_1_rat[OF rai, of ?m, folded cmpx] from IH have *: "r' - l' \ (ir - il) / 2 ^ i" by auto have "r' - (l' + r') / 2 = (r' - l') / 2" by (simp add: field_simps) also have "\ \ (ir - il) / 2 ^ i / 2" using * by (rule divide_right_mono, auto) finally have size: "r' - (l' + r') / 2 \ (ir - il) / (2 * 2 ^ i)" by simp also have "r' - (l' + r') / 2 = (l' + r') / 2 - l'" by auto finally have size': "(l' + r') / 2 - l' \ (ir - il) / (2 * 2 ^ i)" by simp have "root n (real_of_rat ?m) = root n ((real_of_rat ?m') ^ n)" by (simp add: hom_distribs) also have "\ = real_of_rat ?m'" by (rule root_exp_cancel', insert l'0 lr', auto) finally have root: "root n (of_rat ?m) = of_rat ?m'" . show ?case proof (cases "cmpx ?m") case Eq from compare[unfolded Eq] have "real_of_1 x = of_rat ?m" unfolding compare_real_def comparator_of_def by (auto split: if_splits) from arg_cong[OF this, of "root n"] have "?x = root n (of_rat ?m)" . also have "\ = root n (real_of_rat ?m') ^ n" using n real_root_power by (auto simp: hom_distribs) also have "\ = of_rat ?m'" by (rule root_exp_cancel, insert IH sgn(2) l'0 r'0, auto) finally have x: "?x = of_rat ?m'" . show ?thesis using x id Eq lr' ill' ir_il by (auto simp: Let_def) next case Lt from compare[unfolded Lt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (?m',r')" "?f (Suc i - 1) = (l',r')" using Lt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "of_rat ?m' \ ?x" unfolding root by simp with lr'x lr'' have ineq': "real_of_rat l' + real_of_rat r' \ ?x * 2" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size ineq' lr' ill' lr'x ir_il, auto) next case Gt from compare[unfolded Gt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (l',?m')" "?f (Suc i - 1) = (l',r')" using Gt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "?x \ of_rat ?m'" unfolding root by simp with lr'x lr'' have ineq': "?x * 2 \ real_of_rat l' + real_of_rat r'" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size' ineq' lr' ill' lr'x ir_il, auto) qed qed } note main = this fix i from main[of i] show "in_interval (?f i) ?x" by auto from main[of "Suc i"] show "sub_interval (?f (Suc i)) (?f i)" by auto fix eps :: real assume eps: "0 < eps" define c where "c = eps / (max (real_of_rat (ir - il)) 1)" have c0: "c > 0" using eps unfolding c_def by auto from exp_tends_to_zero[OF _ _ this, of "1/2"] obtain i where c: "(1/2)^i \ c" by auto obtain l' r' where fi: "?f i = (l',r')" by force from main[of i, unfolded fi] have le: "r' - l' \ (ir - il) / 2 ^ i" by auto have iril: "real_of_rat (ir - il) \ 0" using ir_il by (auto simp: of_rat_less_eq) show "\n la ra. ?f n = (la, ra) \ real_of_rat ra - real_of_rat la \ eps" proof (intro conjI exI, rule fi) have "real_of_rat r' - of_rat l' = real_of_rat (r' - l')" by (auto simp: hom_distribs) also have "\ \ real_of_rat ((ir - il) / 2 ^ i)" using le unfolding of_rat_less_eq . also have "\ = (real_of_rat (ir - il)) * ((1/2) ^ i)" by (simp add: field_simps hom_distribs) also have "\ \ (real_of_rat (ir - il)) * c" by (rule mult_left_mono[OF c iril]) also have "\ \ eps" proof (cases "real_of_rat (ir - il) \ 1") case True hence "c = eps" unfolding c_def by (auto simp: hom_distribs) thus ?thesis using eps True by auto next case False hence "max (real_of_rat (ir - il)) 1 = real_of_rat (ir - il)" "real_of_rat (ir - il) \ 0" by (auto simp: hom_distribs) hence "(real_of_rat (ir - il)) * c = eps" unfolding c_def by auto thus ?thesis by simp qed finally show "real_of_rat r' - of_rat l' \ eps" . qed qed end private fun root_pos_1 :: "real_alg_1 \ real_alg_2" where "root_pos_1 (p,l,r) = ( (select_correct_factor_int_poly (tighten_bound_root (compare_1_rat (p,l,r))) (\ x. x) (initial_lower_bound l, initial_upper_bound r) (poly_nth_root n p)))" fun root_1 :: "real_alg_1 \ real_alg_2" where "root_1 (p,l,r) = ( if n = 0 \ r = 0 then Rational 0 else if r > 0 then root_pos_1 (p,l,r) else uminus_2 (root_pos_1 (uminus_1 (p,l,r))))" context assumes n: "n \ 0" begin lemma initial_upper_bound: assumes x: "x > 0" and xr: "x \ of_rat r" shows "sgn (initial_upper_bound r) = 1" "root n x \ of_rat (initial_upper_bound r)" proof - have n: "n > 0" using n by auto note d = initial_upper_bound_def let ?r = "initial_upper_bound r" from x xr have r0: "r > 0" by (meson not_less of_rat_le_0_iff order_trans) hence "of_rat r > (0 :: real)" by auto hence "root n (of_rat r) > 0" using n by simp hence "1 \ ceiling (root n (of_rat r))" by auto hence "(1 :: rat) \ of_int (ceiling (root n (of_rat r)))" by linarith also have "\ = ?r" unfolding d by simp finally show "sgn ?r = 1" unfolding sgn_rat_def by auto have "root n x \ root n (of_rat r)" unfolding real_root_le_iff[OF n] by (rule xr) also have "\ \ of_rat ?r" unfolding d by simp finally show "root n x \ of_rat ?r" . qed lemma initial_lower_bound: assumes l: "l > 0" and lx: "of_rat l \ x" shows "sgn (initial_lower_bound l) = 1" "of_rat (initial_lower_bound l) \ root n x" proof - have n: "n > 0" using n by auto note d = initial_lower_bound_def let ?l = "initial_lower_bound l" from l lx have x0: "x > 0" by (meson not_less of_rat_le_0_iff order_trans) have "sgn ?l = 1 \ of_rat ?l \ root n x" proof (cases "l \ 1") case True hence ll: "?l = l" and l0: "of_rat l \ (0 :: real)" and l1: "of_rat l \ (1 :: real)" using l unfolding True d by auto have sgn: "sgn ?l = 1" using l unfolding ll by auto have "of_rat ?l = of_rat l" unfolding ll by simp also have "of_rat l \ root n (of_rat l)" using real_root_increasing[OF _ _ l0 l1, of 1 n] n by (cases "n = 1", auto) also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally show ?thesis using sgn by auto next case False hence l: "(1 :: real) \ of_rat l" and ll: "?l = of_int (floor (root n (of_rat l)))" unfolding d by auto hence "root n 1 \ root n (of_rat l)" unfolding real_root_le_iff[OF n] by auto hence "1 \ root n (of_rat l)" using n by auto from floor_mono[OF this] have "1 \ ?l" using one_le_floor unfolding ll by fastforce hence sgn: "sgn ?l = 1" by simp have "of_rat ?l \ root n (of_rat l)" unfolding ll by simp also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally have "of_rat ?l \ root n x" . with sgn show ?thesis by auto qed thus "sgn ?l = 1" "of_rat ?l \ root n x" by auto qed lemma root_pos_1: assumes x: "invariant_1 x" and pos: "rai_ub x > 0" defines y: "y \ root_pos_1 x" shows "invariant_2 y \ real_of_2 y = root n (real_of_1 x)" proof (cases x) case (fields p l r) let ?l = "initial_lower_bound l" let ?r = "initial_upper_bound r" from x fields have rai: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] let ?x = "the_unique_root (p,l,r)" from pos[unfolded fields] * have sgnl: "sgn l = 1" by auto from sgnl have l0: "l > 0" by (unfold sgn_1_pos) hence ll0: "real_of_rat l > 0" by auto from * have lx: "of_rat l \ ?x" by auto with ll0 have x0: "?x > 0" by linarith note il = initial_lower_bound[OF l0 lx] from * have "?x \ of_rat r" by auto note iu = initial_upper_bound[OF x0 this] let ?p = "poly_nth_root n p" from x0 have id: "root n ?x ^ n = ?x" using n real_root_pow_pos by blast have rc: "root_cond (?p, ?l, ?r) (root n ?x)" using il iu * by (intro root_condI, auto simp: ipoly_nth_root id) hence root: "ipoly ?p (root n (real_of_1 x)) = 0" unfolding root_cond_def fields by auto from * have "p \ 0" by auto hence p': "?p \ 0" using poly_nth_root_0[of n p] n by auto have tbr: "0 \ real_of_1 x" "real_of_rat (initial_lower_bound l) \ root n (real_of_1 x)" "root n (real_of_1 x) \ real_of_rat (initial_upper_bound r)" using x0 il(2) iu(2) fields by auto from select_correct_factor_int_poly[OF tighten_bound_root[OF il(1)[folded fields] tbr x refl n] refl root p'] show ?thesis by (simp add: y fields) qed end lemma root_1: assumes x: "invariant_1 x" defines y: "y \ root_1 x" shows "invariant_2 y \ (real_of_2 y = root n (real_of_1 x))" proof (cases "n = 0 \ rai_ub x = 0") case True with x have "n = 0 \ real_of_1 x = 0" by (cases x, auto) then have "root n (real_of_1 x) = 0" by auto then show ?thesis unfolding y root_1.simps using x by (cases x, auto) next case False with x have n: "n \ 0" and x0: "real_of_1 x \ 0" by (simp, cases x, auto) note rt = root_pos_1 show ?thesis proof (cases "rai_ub x" "0::rat" rule:linorder_cases) case greater with rt[OF n x this] n show ?thesis by (unfold y, cases x, simp) next case less let ?um = "uminus_1" let ?rt = "root_pos_1" from n less y x0 have y: "y = uminus_2 (?rt (?um x))" by (cases x, auto) from uminus_1[OF x] have umx: "invariant_1 (?um x)" and umx2: "real_of_1 (?um x) = - real_of_1 x" by auto with x less have "0 < rai_ub (uminus_1 x)" by (cases x, auto simp: uminus_1.simps Let_def) from rt[OF n umx this] umx2 have rumx: "invariant_2 (?rt (?um x))" and rumx2: "real_of_2 (?rt (?um x)) = root n (- real_of_1 x)" by auto from uminus_2[OF rumx] rumx2 y real_root_minus show ?thesis by auto next case equal with x0 x show ?thesis by (cases x, auto) qed qed end declare root_1.simps[simp del] (* ********************** *) subsubsection \Embedding of Rational Numbers\ definition of_rat_1 :: "rat \ real_alg_1" where "of_rat_1 x \ (poly_rat x,x,x)" lemma of_rat_1: shows "invariant_1 (of_rat_1 x)" and "real_of_1 (of_rat_1 x) = of_rat x" unfolding of_rat_1_def by (atomize(full), intro invariant_1_realI unique_rootI poly_condI, auto ) fun info_2 :: "real_alg_2 \ rat + int poly \ nat" where "info_2 (Rational x) = Inl x" | "info_2 (Irrational n (p,l,r)) = Inr (p,n)" lemma info_2_card: assumes rc: "invariant_2 x" shows "info_2 x = Inr (p,n) \ poly_cond p \ ipoly p (real_of_2 x) = 0 \ degree p \ 2 \ card (roots_below p (real_of_2 x)) = n" "info_2 x = Inl y \ real_of_2 x = of_rat y" proof (atomize(full), goal_cases) case 1 show ?case proof (cases x) case (Irrational m rai) then obtain q l r where x: "x = Irrational m (q,l,r)" by (cases rai, auto) show ?thesis proof (cases "q = p \ m = n") case False thus ?thesis using x by auto next case True with x have x: "x = Irrational n (p,l,r)" by auto from rc[unfolded x, simplified] have inv: "invariant_1_2 (p,l,r)" and n: "card (roots_below p (real_of_2 x)) = n" and 1: "degree p \ 1" by (auto simp: x) from inv have "degree p \ 0" unfolding irreducible_def by auto with 1 have "degree p \ 2" by linarith thus ?thesis unfolding n using inv x by (auto elim!: invariant_1E) qed qed auto qed lemma real_of_2_Irrational: "invariant_2 (Irrational n rai) \ real_of_2 (Irrational n rai) \ of_rat x" proof assume "invariant_2 (Irrational n rai)" and rat: "real_of_2 (Irrational n rai) = real_of_rat x" hence "real_of_1 rai \ \" "invariant_1_2 rai" by auto from invariant_1_2_of_rat[OF this(2)] rat show False by auto qed lemma info_2: assumes ix: "invariant_2 x" and iy: "invariant_2 y" shows "info_2 x = info_2 y \ real_of_2 x = real_of_2 y" proof (cases x) case x: (Irrational n1 rai1) note ix = ix[unfolded x] show ?thesis proof (cases y) case (Rational y) with real_of_2_Irrational[OF ix, of y] show ?thesis unfolding x by (cases rai1, auto) next case y: (Irrational n2 rai2) obtain p1 l1 r1 where rai1: "rai1 = (p1,l1,r1)" by (cases rai1) obtain p2 l2 r2 where rai2: "rai2 = (p2,l2,r2)" by (cases rai2) let ?rx = "the_unique_root (p1,l1,r1)" let ?ry = "the_unique_root (p2,l2,r2)" have id: "(info_2 x = info_2 y) = (p1 = p2 \ n1 = n2)" "(real_of_2 x = real_of_2 y) = (?rx = ?ry)" unfolding x y rai1 rai2 by auto from ix[unfolded x rai1] have ix: "invariant_1 (p1, l1, r1)" and deg1: "degree p1 > 1" and n1: "n1 = card (roots_below p1 ?rx)" by auto note Ix = invariant_1D[OF ix] from deg1 have p1_0: "p1 \ 0" by auto from iy[unfolded y rai2] have iy: "invariant_1 (p2, l2, r2)" and "degree p2 > 1" and n2: "n2 = card (roots_below p2 ?ry)" by auto note Iy = invariant_1D[OF iy] show ?thesis unfolding id proof assume eq: "?rx = ?ry" from Ix have algx: "p1 represents ?rx \ irreducible p1 \ lead_coeff p1 > 0" unfolding represents_def by auto from iy have algy: "p2 represents ?rx \ irreducible p2 \ lead_coeff p2 > 0" unfolding represents_def eq by (auto elim!: invariant_1E) from algx have "algebraic ?rx" unfolding algebraic_altdef_ipoly by auto note unique = algebraic_imp_represents_unique[OF this] with algx algy have id: "p2 = p1" by auto from eq id n1 n2 show "p1 = p2 \ n1 = n2" by auto next assume "p1 = p2 \ n1 = n2" hence id: "p1 = p2" "n1 = n2" by auto hence card: "card (roots_below p1 ?rx) = card (roots_below p1 ?ry)" unfolding n1 n2 by auto show "?rx = ?ry" proof (cases ?rx ?ry rule: linorder_cases) case less have "roots_below p1 ?rx = roots_below p1 ?ry" proof (intro card_subset_eq finite_subset[OF _ ipoly_roots_finite] card) from less show "roots_below p1 ?rx \ roots_below p1 ?ry" by auto qed (insert p1_0, auto) then show ?thesis using id less unique_rootD(3)[OF Iy(4)] by (auto simp: less_eq_real_def) next case equal then show ?thesis by (simp add: id) next case greater have "roots_below p1 ?ry = roots_below p1 ?rx" proof (intro card_subset_eq card[symmetric] finite_subset[OF _ ipoly_roots_finite[OF p1_0]]) from greater show "roots_below p1 ?ry \ roots_below p1 ?rx" by auto qed auto hence "roots_below p2 ?ry = roots_below p2 ?rx" unfolding id by auto thus ?thesis using id greater unique_rootD(3)[OF Ix(4)] by (auto simp: less_eq_real_def) qed qed qed next case x: (Rational x) show ?thesis proof (cases y) case (Rational y) thus ?thesis using x by auto next case y: (Irrational n rai) with real_of_2_Irrational[OF iy[unfolded y], of x] show ?thesis unfolding x by (cases rai, auto) qed qed lemma info_2_unique: "invariant_2 x \ invariant_2 y \ real_of_2 x = real_of_2 y \ info_2 x = info_2 y" using info_2 by blast lemma info_2_inj: "invariant_2 x \ invariant_2 y \ info_2 x = info_2 y \ real_of_2 x = real_of_2 y" using info_2 by blast context fixes cr1 cr2 :: "rat \ rat \ nat" begin partial_function (tailrec) compare_1 :: "int poly \ int poly \ rat \ rat \ rat \ rat \ rat \ rat \ order" where [code]: "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = (if r1 < l2 then Lt else if r2 < l1 then Gt else let (l1',r1',sr1') = tighten_poly_bounds p1 l1 r1 sr1; (l2',r2',sr2') = tighten_poly_bounds p2 l2 r2 sr2 in compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2') " lemma compare_1: assumes ur1: "unique_root (p1,l1,r1)" and ur2: "unique_root (p2,l2,r2)" and pc: "poly_cond2 p1" "poly_cond2 p2" and diff: "the_unique_root (p1,l1,r1) \ the_unique_root (p2,l2,r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = compare (the_unique_root (p1,l1,r1)) (the_unique_root (p2,l2,r2))" proof - let ?r = real_of_rat { fix d x y assume d: "d = (r1 - l1) + (r2 - l2)" and xy: "x = the_unique_root (p1,l1,r1)" "y = the_unique_root (p2,l2,r2)" define delta where "delta = abs (x - y) / 4" have delta: "delta > 0" and diff: "x \ y" unfolding delta_def using diff xy by auto let ?rel' = "{(x, y). 0 \ y \ delta_gt delta x y}" let ?rel = "inv_image ?rel' ?r" have SN: "SN ?rel" by (rule SN_inv_image[OF delta_gt_SN[OF delta]]) from d ur1 ur2 have ?thesis unfolding xy[symmetric] using xy sr proof (induct d arbitrary: l1 r1 l2 r2 sr1 sr2 rule: SN_induct[OF SN]) case (1 d l1 r1 l2 r2) note IH = 1(1) note d = 1(2) note ur = 1(3-4) note xy = 1(5-6) note sr = 1(7-8) note simps = compare_1.simps[of p1 p2 l1 r1 sr1 l2 r2 sr2] note urx = unique_rootD[OF ur(1), folded xy] note ury = unique_rootD[OF ur(2), folded xy] show ?case (is "?l = _") proof (cases "r1 < l2") case True hence l: "?l = Lt" and lt: "?r r1 < ?r l2" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True urx(2) ury(1) by (auto simp: compare_real_def comparator_of_def) next case False note le = this show ?thesis proof (cases "r2 < l1") case True with le have l: "?l = Gt" and lt: "?r r2 < ?r l1" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True ury(2) urx(1) by (auto simp: compare_real_def comparator_of_def) next case False obtain l1' r1' sr1' where tb1: "tighten_poly_bounds p1 l1 r1 sr1 = (l1',r1',sr1')" by (cases rule: prod_cases3, auto) obtain l2' r2' sr2' where tb2: "tighten_poly_bounds p2 l2 r2 sr2 = (l2',r2',sr2')" by (cases rule: prod_cases3, auto) from False le tb1 tb2 have l: "?l = compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2'" unfolding simps by auto from tighten_poly_bounds[OF tb1 ur(1) pc(1) sr(1)] have rc1: "root_cond (p1, l1', r1') (the_unique_root (p1, l1, r1))" and bnd1: "l1 \ l1'" "l1' \ r1'" "r1' \ r1" and d1: "r1' - l1' = (r1 - l1) / 2" and sr1: "sr1' = sgn (ipoly p1 r1')" by auto from pc have "p1 \ 0" "p2 \ 0" by auto from unique_root_sub_interval[OF ur(1) rc1 bnd1(1,3)] xy ur this have ur1: "unique_root (p1, l1', r1')" and x: "x = the_unique_root (p1, l1', r1')" by (auto intro!: the_unique_root_eqI) from tighten_poly_bounds[OF tb2 ur(2) pc(2) sr(2)] have rc2: "root_cond (p2, l2', r2') (the_unique_root (p2, l2, r2))" and bnd2: "l2 \ l2'" "l2' \ r2'" "r2' \ r2" and d2: "r2' - l2' = (r2 - l2) / 2" and sr2: "sr2' = sgn (ipoly p2 r2')" by auto from unique_root_sub_interval[OF ur(2) rc2 bnd2(1,3)] xy ur pc have ur2: "unique_root (p2, l2', r2')" and y: "y = the_unique_root (p2, l2', r2')" by auto define d' where "d' = d/2" have d': "d' = r1' - l1' + (r2' - l2')" unfolding d'_def d d1 d2 by (simp add: field_simps) have d'0: "d' \ 0" using bnd1 bnd2 unfolding d' by auto have dd: "d - d' = d/2" unfolding d'_def by simp have "abs (x - y) \ 2 * ?r d" proof (rule ccontr) assume "\ ?thesis" hence lt: "2 * ?r d < abs (x - y)" by auto have "r1 - l1 \ d" "r2 - l2 \ d" unfolding d using bnd1 bnd2 by auto from this[folded of_rat_less_eq[where 'a = real]] lt have "?r (r1 - l1) < abs (x - y) / 2" "?r (r2 - l2) < abs (x - y) / 2" and dd: "?r r1 - ?r l1 \ ?r d" "?r r2 - ?r l2 \ ?r d" by (auto simp: of_rat_diff) from le have "r1 \ l2" by auto hence r1l2: "?r r1 \ ?r l2" unfolding of_rat_less_eq by auto from False have "r2 \ l1" by auto hence r2l1: "?r r2 \ ?r l1" unfolding of_rat_less_eq by auto show False proof (cases "x \ y") case True from urx(1-2) dd(1) have "?r r1 \ x + ?r d" by auto with r1l2 have "?r l2 \ x + ?r d" by auto with True lt ury(2) dd(2) show False by auto next case False from ury(1-2) dd(2) have "?r r2 \ y + ?r d" by auto with r2l1 have "?r l1 \ y + ?r d" by auto with False lt urx(2) dd(1) show False by auto qed qed hence dd': "delta_gt delta (?r d) (?r d')" unfolding delta_gt_def delta_def using dd by (auto simp: hom_distribs) show ?thesis unfolding l by (rule IH[OF _ d' ur1 ur2 x y sr1 sr2], insert d'0 dd', auto) qed qed qed } thus ?thesis by auto qed end (* **************************************************************** *) fun real_alg_1 :: "real_alg_2 \ real_alg_1" where "real_alg_1 (Rational r) = of_rat_1 r" | "real_alg_1 (Irrational n rai) = rai" lemma real_alg_1: "real_of_1 (real_alg_1 x) = real_of_2 x" by (cases x, auto simp: of_rat_1) definition root_2 :: "nat \ real_alg_2 \ real_alg_2" where "root_2 n x = root_1 n (real_alg_1 x)" lemma root_2: assumes "invariant_2 x" shows "real_of_2 (root_2 n x) = root n (real_of_2 x)" "invariant_2 (root_2 n x)" proof (atomize(full), cases x, goal_cases) case (1 y) from of_rat_1[of y] root_1[of "of_rat_1 y" n] assms 1 real_alg_2 show ?case by (simp add: root_2_def) next case (2 i rai) from root_1[of rai n] assms 2 real_alg_2 show ?case by (auto simp: root_2_def) qed fun add_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "add_2 (Rational r) (Rational q) = Rational (r + q)" | "add_2 (Rational r) (Irrational n x) = Irrational n (add_rat_1 r x)" | "add_2 (Irrational n x) (Rational q) = Irrational n (add_rat_1 q x)" | "add_2 (Irrational n x) (Irrational m y) = add_1 x y" lemma add_2: assumes x: "invariant_2 x" and y: "invariant_2 y" shows "invariant_2 (add_2 x y)" (is ?g1) and "real_of_2 (add_2 x y) = real_of_2 x + real_of_2 y" (is ?g2) using assms add_rat_1 add_1 by (atomize (full), (cases x; cases y), auto simp: hom_distribs) fun mult_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "mult_2 (Rational r) (Rational q) = Rational (r * q)" | "mult_2 (Rational r) (Irrational n y) = mult_rat_1 r y" | "mult_2 (Irrational n x) (Rational q) = mult_rat_1 q x" | "mult_2 (Irrational n x) (Irrational m y) = mult_1 x y" lemma mult_2: assumes "invariant_2 x" "invariant_2 y" shows "real_of_2 (mult_2 x y) = real_of_2 x * real_of_2 y" "invariant_2 (mult_2 x y)" using assms by (atomize(full), (cases x; cases y; auto simp: mult_rat_1 mult_1 hom_distribs)) fun to_rat_2 :: "real_alg_2 \ rat option" where "to_rat_2 (Rational r) = Some r" | "to_rat_2 (Irrational n rai) = None" lemma to_rat_2: assumes rc: "invariant_2 x" shows "to_rat_2 x = (if real_of_2 x \ \ then Some (THE q. real_of_2 x = of_rat q) else None)" proof (cases x) case (Irrational n rai) from real_of_2_Irrational[OF rc[unfolded this]] show ?thesis unfolding Irrational Rats_def by auto qed simp fun equal_2 :: "real_alg_2 \ real_alg_2 \ bool" where "equal_2 (Rational r) (Rational q) = (r = q)" | "equal_2 (Irrational n (p,_)) (Irrational m (q,_)) = (p = q \ n = m)" | "equal_2 (Rational r) (Irrational _ yy) = False" | "equal_2 (Irrational _ xx) (Rational q) = False" lemma equal_2[simp]: assumes rc: "invariant_2 x" "invariant_2 y" shows "equal_2 x y = (real_of_2 x = real_of_2 y)" using info_2[OF rc] by (cases x; cases y, auto) fun compare_2 :: "real_alg_2 \ real_alg_2 \ order" where "compare_2 (Rational r) (Rational q) = (compare r q)" | "compare_2 (Irrational n (p,l,r)) (Irrational m (q,l',r')) = (if p = q \ n = m then Eq else compare_1 p q l r (sgn (ipoly p r)) l' r' (sgn (ipoly q r')))" | "compare_2 (Rational r) (Irrational _ xx) = (compare_rat_1 r xx)" | "compare_2 (Irrational _ xx) (Rational r) = (invert_order (compare_rat_1 r xx))" lemma compare_2: assumes rc: "invariant_2 x" "invariant_2 y" shows "compare_2 x y = compare (real_of_2 x) (real_of_2 y)" proof (cases x) case (Rational r) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this show ?thesis unfolding xx yy by (simp add: compare_rat_def compare_real_def comparator_of_def of_rat_less) next case (Irrational n yy) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by (simp add: of_rat_1) qed next case (Irrational n xx) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by simp next case (Irrational m yy) note yy = this obtain p l r where xxx: "xx = (p,l,r)" by (cases xx) obtain q l' r' where yyy: "yy = (q,l',r')" by (cases yy) note rc = rc[unfolded xx xxx yy yyy] from rc have I: "invariant_1_2 (p,l,r)" "invariant_1_2 (q,l',r')" by auto then have "unique_root (p,l,r)" "unique_root (q,l',r')" "poly_cond2 p" "poly_cond2 q" by auto from compare_1[OF this _ refl refl] show ?thesis using equal_2[OF rc] unfolding xx xxx yy yyy by simp qed qed fun sgn_2 :: "real_alg_2 \ rat" where "sgn_2 (Rational r) = sgn r" | "sgn_2 (Irrational n rai) = sgn_1 rai" lemma sgn_2: "invariant_2 x \ real_of_rat (sgn_2 x) = sgn (real_of_2 x)" using sgn_1 by (cases x, auto simp: real_of_rat_sgn) fun floor_2 :: "real_alg_2 \ int" where "floor_2 (Rational r) = floor r" | "floor_2 (Irrational n rai) = floor_1 rai" lemma floor_2: "invariant_2 x \ floor_2 x = floor (real_of_2 x)" by (cases x, auto simp: floor_1) (* *************** *) subsubsection \Definitions and Algorithms on Type with Invariant\ lift_definition of_rat_3 :: "rat \ real_alg_3" is of_rat_2 by (auto simp: of_rat_2) lemma of_rat_3: "real_of_3 (of_rat_3 x) = of_rat x" by (transfer, auto simp: of_rat_2) lift_definition root_3 :: "nat \ real_alg_3 \ real_alg_3" is root_2 by (auto simp: root_2) lemma root_3: "real_of_3 (root_3 n x) = root n (real_of_3 x)" by (transfer, auto simp: root_2) lift_definition equal_3 :: "real_alg_3 \ real_alg_3 \ bool" is equal_2 . lemma equal_3: "equal_3 x y = (real_of_3 x = real_of_3 y)" by (transfer, auto) lift_definition compare_3 :: "real_alg_3 \ real_alg_3 \ order" is compare_2 . lemma compare_3: "compare_3 x y = (compare (real_of_3 x) (real_of_3 y))" by (transfer, auto simp: compare_2) lift_definition add_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is add_2 by (auto simp: add_2) lemma add_3: "real_of_3 (add_3 x y) = real_of_3 x + real_of_3 y" by (transfer, auto simp: add_2) lift_definition mult_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is mult_2 by (auto simp: mult_2) lemma mult_3: "real_of_3 (mult_3 x y) = real_of_3 x * real_of_3 y" by (transfer, auto simp: mult_2) lift_definition sgn_3 :: "real_alg_3 \ rat" is sgn_2 . lemma sgn_3: "real_of_rat (sgn_3 x) = sgn (real_of_3 x)" by (transfer, auto simp: sgn_2) lift_definition to_rat_3 :: "real_alg_3 \ rat option" is to_rat_2 . lemma to_rat_3: "to_rat_3 x = (if real_of_3 x \ \ then Some (THE q. real_of_3 x = of_rat q) else None)" by (transfer, simp add: to_rat_2) lift_definition floor_3 :: "real_alg_3 \ int" is floor_2 . lemma floor_3: "floor_3 x = floor (real_of_3 x)" by (transfer, auto simp: floor_2) (* *************** *) (* info *) lift_definition info_3 :: "real_alg_3 \ rat + int poly \ nat" is info_2 . lemma info_3_fun: "real_of_3 x = real_of_3 y \ info_3 x = info_3 y" by (transfer, intro info_2_unique, auto) lift_definition info_real_alg :: "real_alg \ rat + int poly \ nat" is info_3 by (metis info_3_fun) lemma info_real_alg: "info_real_alg x = Inr (p,n) \ p represents (real_of x) \ card {y. y \ real_of x \ ipoly p y = 0} = n \ irreducible p" "info_real_alg x = Inl q \ real_of x = of_rat q" proof (atomize(full), transfer, transfer, goal_cases) case (1 x p n q) from 1 have x: "invariant_2 x" by auto note info = info_2_card[OF this] show ?case proof (cases x) case irr: (Irrational m rai) from info(1)[of p n] show ?thesis unfolding irr by (cases rai, auto simp: poly_cond_def) qed (insert 1 info, auto) qed (* add *) instantiation real_alg :: plus begin lift_definition plus_real_alg :: "real_alg \ real_alg \ real_alg" is add_3 by (simp add: add_3) instance .. end lemma plus_real_alg: "(real_of x) + (real_of y) = real_of (x + y)" by (transfer, rule add_3[symmetric]) (* minus *) instantiation real_alg :: minus begin definition minus_real_alg :: "real_alg \ real_alg \ real_alg" where "minus_real_alg x y = x + (-y)" instance .. end lemma minus_real_alg: "(real_of x) - (real_of y) = real_of (x - y)" unfolding minus_real_alg_def minus_real_def uminus_real_alg plus_real_alg .. (* of_rat *) lift_definition of_rat_real_alg :: "rat \ real_alg" is of_rat_3 . lemma of_rat_real_alg: "real_of_rat x = real_of (of_rat_real_alg x)" by (transfer, rule of_rat_3[symmetric]) (* zero *) instantiation real_alg :: zero begin definition zero_real_alg :: "real_alg" where "zero_real_alg \ of_rat_real_alg 0" instance .. end lemma zero_real_alg: "0 = real_of 0" unfolding zero_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* one *) instantiation real_alg :: one begin definition one_real_alg :: "real_alg" where "one_real_alg \ of_rat_real_alg 1" instance .. end lemma one_real_alg: "1 = real_of 1" unfolding one_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* times *) instantiation real_alg :: times begin lift_definition times_real_alg :: "real_alg \ real_alg \ real_alg" is mult_3 by (simp add: mult_3) instance .. end lemma times_real_alg: "(real_of x) * (real_of y) = real_of (x * y)" by (transfer, rule mult_3[symmetric]) (* inverse *) instantiation real_alg :: inverse begin lift_definition inverse_real_alg :: "real_alg \ real_alg" is inverse_3 by (simp add: inverse_3) definition divide_real_alg :: "real_alg \ real_alg \ real_alg" where "divide_real_alg x y = x * inverse y" (* TODO: better to use poly_div *) instance .. end lemma inverse_real_alg: "inverse (real_of x) = real_of (inverse x)" by (transfer, rule inverse_3[symmetric]) lemma divide_real_alg: "(real_of x) / (real_of y) = real_of (x / y)" unfolding divide_real_alg_def times_real_alg[symmetric] divide_real_def inverse_real_alg .. (* group *) instance real_alg :: ab_group_add apply intro_classes apply (transfer, unfold add_3, force) apply (unfold zero_real_alg_def, transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 uminus_3 of_rat_3, force) apply (unfold minus_real_alg_def, force) done (* field *) instance real_alg :: field apply intro_classes apply (transfer, unfold mult_3, force) apply (transfer, unfold mult_3, force) apply (unfold one_real_alg_def, transfer, unfold mult_3 of_rat_3, force) apply (transfer, unfold mult_3 add_3, force simp: field_simps) apply (unfold zero_real_alg_def, transfer, unfold of_rat_3, force) apply (transfer, unfold mult_3 inverse_3 of_rat_3, force simp: field_simps) apply (unfold divide_real_alg_def, force) apply (transfer, unfold inverse_3 of_rat_3, force) done (* numeral *) instance real_alg :: numeral .. (* root *) lift_definition root_real_alg :: "nat \ real_alg \ real_alg" is root_3 by (simp add: root_3) lemma root_real_alg: "root n (real_of x) = real_of (root_real_alg n x)" by (transfer, rule root_3[symmetric]) (* sgn *) lift_definition sgn_real_alg_rat :: "real_alg \ rat" is sgn_3 by (insert sgn_3, metis to_rat_of_rat) lemma sgn_real_alg_rat: "real_of_rat (sgn_real_alg_rat x) = sgn (real_of x)" by (transfer, auto simp: sgn_3) instantiation real_alg :: sgn begin definition sgn_real_alg :: "real_alg \ real_alg" where "sgn_real_alg x = of_rat_real_alg (sgn_real_alg_rat x)" instance .. end lemma sgn_real_alg: "sgn (real_of x) = real_of (sgn x)" unfolding sgn_real_alg_def of_rat_real_alg[symmetric] by (transfer, simp add: sgn_3) (* equal *) instantiation real_alg :: equal begin lift_definition equal_real_alg :: "real_alg \ real_alg \ bool" is equal_3 by (simp add: equal_3) instance proof fix x y :: real_alg show "equal_class.equal x y = (x = y)" by (transfer, simp add: equal_3) qed end lemma equal_real_alg: "HOL.equal (real_of x) (real_of y) = (x = y)" unfolding equal_real_def by (transfer, auto) (* comparisons *) instantiation real_alg :: ord begin definition less_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_real_alg x y = (real_of x < real_of y)" definition less_eq_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_eq_real_alg x y = (real_of x \ real_of y)" instance .. end lemma less_real_alg: "less (real_of x) (real_of y) = (x < y)" unfolding less_real_alg_def .. lemma less_eq_real_alg: "less_eq (real_of x) (real_of y) = (x \ y)" unfolding less_eq_real_alg_def .. instantiation real_alg :: compare_order begin lift_definition compare_real_alg :: "real_alg \ real_alg \ order" is compare_3 by (simp add: compare_3) lemma compare_real_alg: "compare (real_of x) (real_of y) = (compare x y)" by (transfer, simp add: compare_3) instance proof (intro_classes, unfold compare_real_alg[symmetric, abs_def]) show "le_of_comp (\x y. compare (real_of x) (real_of y)) = (\)" by (intro ext, auto simp: compare_real_def comparator_of_def le_of_comp_def less_eq_real_alg_def) show "lt_of_comp (\x y. compare (real_of x) (real_of y)) = (<)" by (intro ext, auto simp: compare_real_def comparator_of_def lt_of_comp_def less_real_alg_def) show "comparator (\x y. compare (real_of x) (real_of y))" unfolding comparator_def proof (intro conjI impI allI) fix x y z :: "real_alg" let ?r = real_of note rc = comparator_compare[where 'a = real, unfolded comparator_def] from rc show "invert_order (compare (?r x) (?r y)) = compare (?r y) (?r x)" by blast from rc show "compare (?r x) (?r y) = Lt \ compare (?r y) (?r z) = Lt \ compare (?r x) (?r z) = Lt" by blast assume "compare (?r x) (?r y) = Eq" with rc have "?r x = ?r y" by blast thus "x = y" unfolding real_of_inj . qed qed end lemma less_eq_real_alg_code[code]: "(less_eq :: real_alg \ real_alg \ bool) = le_of_comp compare" "(less :: real_alg \ real_alg \ bool) = lt_of_comp compare" by (rule ord_defs(1)[symmetric], rule ord_defs(2)[symmetric]) instantiation real_alg :: abs begin definition abs_real_alg :: "real_alg \ real_alg" where "abs_real_alg x = (if real_of x < 0 then uminus x else x)" instance .. end lemma abs_real_alg: "abs (real_of x) = real_of (abs x)" unfolding abs_real_alg_def abs_real_def if_distrib by (auto simp: uminus_real_alg) lemma sgn_real_alg_sound: "sgn x = (if x = 0 then 0 else if 0 < real_of x then 1 else - 1)" (is "_ = ?r") proof - have "real_of (sgn x) = sgn (real_of x)" by (simp add: sgn_real_alg) also have "\ = real_of ?r" unfolding sgn_real_def if_distrib by (auto simp: less_real_alg_def zero_real_alg_def one_real_alg_def of_rat_real_alg[symmetric] equal_real_alg[symmetric] equal_real_def uminus_real_alg[symmetric]) finally show "sgn x = ?r" unfolding equal_real_alg[symmetric] equal_real_def by simp qed lemma real_of_of_int: "real_of_rat (rat_of_int z) = real_of (of_int z)" proof (cases "z \ 0") case True define n where "n = nat z" from True have z: "z = int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg hom_distribs) next case False define n where "n = nat (-z)" from False have z: "z = - int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg uminus_real_alg[symmetric] minus_real_alg[symmetric] hom_distribs) qed instance real_alg :: linordered_field apply standard apply (unfold less_eq_real_alg_def plus_real_alg[symmetric], force) apply (unfold abs_real_alg_def less_real_alg_def zero_real_alg[symmetric], rule refl) apply (unfold less_real_alg_def times_real_alg[symmetric], force) apply (rule sgn_real_alg_sound) done instantiation real_alg :: floor_ceiling begin lift_definition floor_real_alg :: "real_alg \ int" is floor_3 by (auto simp: floor_3) lemma floor_real_alg: "floor (real_of x) = floor x" by (transfer, auto simp: floor_3) instance proof fix x :: real_alg show "of_int \x\ \ x \ x < of_int (\x\ + 1)" unfolding floor_real_alg[symmetric] using floor_correct[of "real_of x"] unfolding less_eq_real_alg_def less_real_alg_def real_of_of_int[symmetric] by (auto simp: hom_distribs) hence "x \ of_int (\x\ + 1)" by auto thus "\z. x \ of_int z" by blast qed end definition real_alg_of_real :: "real \ real_alg" where "real_alg_of_real x = (if (\ y. x = real_of y) then (THE y. x = real_of y) else 0)" lemma real_alg_of_real_code[code]: "real_alg_of_real (real_of x) = x" using real_of_inj unfolding real_alg_of_real_def by auto lift_definition to_rat_real_alg_main :: "real_alg \ rat option" is to_rat_3 by (simp add: to_rat_3) lemma to_rat_real_alg_main: "to_rat_real_alg_main x = (if real_of x \ \ then Some (THE q. real_of x = of_rat q) else None)" by (transfer, simp add: to_rat_3) definition to_rat_real_alg :: "real_alg \ rat" where "to_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ q | None \ 0)" definition is_rat_real_alg :: "real_alg \ bool" where "is_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ True | None \ False)" lemma is_rat_real_alg: "is_rat (real_of x) = (is_rat_real_alg x)" unfolding is_rat_real_alg_def is_rat to_rat_real_alg_main by auto lemma to_rat_real_alg: "to_rat (real_of x) = (to_rat_real_alg x)" unfolding to_rat to_rat_real_alg_def to_rat_real_alg_main by auto subsection \Real Algebraic Numbers as Implementation for Real Numbers\ lemmas real_alg_code_eqns = one_real_alg zero_real_alg uminus_real_alg root_real_alg minus_real_alg plus_real_alg times_real_alg inverse_real_alg divide_real_alg equal_real_alg less_real_alg less_eq_real_alg compare_real_alg sgn_real_alg abs_real_alg floor_real_alg is_rat_real_alg to_rat_real_alg code_datatype real_of declare [[code drop: "plus :: real \ real \ real" "uminus :: real \ real" "minus :: real \ real \ real" "times :: real \ real \ real" "inverse :: real \ real" "divide :: real \ real \ real" "floor :: real \ int" "HOL.equal :: real \ real \ bool" "compare :: real \ real \ order" "less_eq :: real \ real \ bool" "less :: real \ real \ bool" "0 :: real" "1 :: real" "sgn :: real \ real" "abs :: real \ real" root]] declare real_alg_code_eqns [code equation] lemma [code]: "Ratreal = real_of \ of_rat_real_alg" by (transfer, transfer) (simp add: fun_eq_iff of_rat_2) end diff --git a/thys/Algebraic_Numbers/Real_Roots.thy b/thys/Algebraic_Numbers/Real_Roots.thy --- a/thys/Algebraic_Numbers/Real_Roots.thy +++ b/thys/Algebraic_Numbers/Real_Roots.thy @@ -1,743 +1,746 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Real Roots\ text \This theory contains an algorithm to determine the set of real roots of a rational polynomial. It further contains an algorithm which tries to determine the real roots of real-valued polynomial, which incorporates Yun-factorization and closed formulas for polynomials of degree 2.\ theory Real_Roots imports Real_Algebraic_Numbers begin +hide_const (open) UnivPoly.coeff +hide_const (open) Module.smult + text \Division of integers, rounding to the upper value.\ definition div_ceiling :: "int \ int \ int" where "div_ceiling x y = (let q = x div y in if q * y = x then q else q + 1)" definition root_bound :: "int poly \ rat" where "root_bound p \ let n = degree p; m = 1 + div_ceiling (max_list_non_empty (map (\i. abs (coeff p i)) [0.. \round to the next higher number \2^n\, so that bisection will\ \ \stay on integers for as long as possible\ in of_int (2 ^ (log_ceiling 2 m))" partial_function (tailrec) roots_of_2_main :: "int poly \ root_info \ (rat \ rat \ nat) \ (rat \ rat)list \ real_alg_2 list \ real_alg_2 list" where [code]: "roots_of_2_main p ri cr lrs rais = (case lrs of Nil \ rais | (l,r) # lrs \ let c = cr l r in if c = 0 then roots_of_2_main p ri cr lrs rais else if c = 1 then roots_of_2_main p ri cr lrs (real_alg_2'' ri p l r # rais) else let m = (l + r) / 2 in roots_of_2_main p ri cr ((m,r) # (l,m) # lrs) rais)" definition roots_of_2_irr :: "int poly \ real_alg_2 list" where "roots_of_2_irr p = (if degree p = 1 then [Rational (Rat.Fract (- coeff p 0) (coeff p 1)) ] else let ri = root_info p; cr = root_info.l_r ri; B = root_bound p in (roots_of_2_main p ri cr [(-B,B)] []))" lemma root_imp_deg_nonzero: assumes "p \ 0" "poly p x = 0" shows "degree p \ 0" proof assume "degree p = 0" from degree0_coeffs[OF this] assms show False by auto qed lemma cauchy_root_bound: fixes x :: "'a :: real_normed_field" assumes x: "poly p x = 0" and p: "p \ 0" shows "norm x \ 1 + max_list_non_empty (map (\ i. norm (coeff p i)) [0 ..< degree p]) / norm (lead_coeff p)" (is "_ \ _ + ?max / ?nlc") proof - let ?n = "degree p" let ?p = "coeff p" let ?lc = "lead_coeff p" define ml where "ml = ?max / ?nlc" from p have lc: "?lc \ 0" by auto hence nlc: "norm ?lc > 0" by auto from root_imp_deg_nonzero[OF p x] have *: "0 \ set [0 ..< degree p]" by auto have "0 \ norm (?p 0)" by simp also have "\ \ ?max" by (rule max_list_non_empty, insert *, auto) finally have max0: "?max \ 0" . with nlc have ml0: "ml \ 0" unfolding ml_def by auto hence easy: "norm x \ 1 \ ?thesis" unfolding ml_def[symmetric] by auto show ?thesis proof (cases "norm x \ 1") case True thus ?thesis using easy by auto next case False hence nx: "norm x > 1" by simp hence x0: "x \ 0" by auto hence xn0: "0 < norm x ^ ?n" by auto from x[unfolded poly_altdef] have "x ^ ?n * ?lc = x ^ ?n * ?lc - (\i\?n. x ^ i * ?p i)" unfolding poly_altdef by (simp add: ac_simps) also have "(\i\?n. x ^ i * ?p i) = x ^ ?n * ?lc + (\i < ?n. x ^ i * ?p i)" by (subst sum.remove[of _ ?n], auto intro: sum.cong) finally have "x ^ ?n * ?lc = - (\i < ?n. x ^ i * ?p i)" by simp with lc have "x ^ ?n = - (\i < ?n. x ^ i * ?p i) / ?lc" by (simp add: field_simps) from arg_cong[OF this, of norm] have "norm x ^ ?n = norm ((\i < ?n. x ^ i * ?p i) / ?lc)" unfolding norm_power by simp also have "(\i < ?n. x ^ i * ?p i) / ?lc = (\i < ?n. x ^ i * ?p i / ?lc)" by (rule sum_divide_distrib) also have "norm \ \ (\i < ?n. norm (x ^ i * (?p i / ?lc)))" by (simp add: field_simps, rule norm_sum) also have "\ = (\i < ?n. norm x ^ i * norm (?p i / ?lc))" unfolding norm_mult norm_power .. also have "\ \ (\i < ?n. norm x ^ i * ml)" proof (rule sum_mono) fix i assume "i \ {.. norm x ^ i * ml" proof (rule mult_left_mono) show "0 \ norm x ^ i" using nx by auto show "norm (?p i / ?lc) \ ml" unfolding norm_divide ml_def by (rule divide_right_mono[OF max_list_non_empty], insert nlc i, auto) qed qed also have "\ = ml * (\i < ?n. norm x ^ i)" unfolding sum_distrib_right[symmetric] by simp also have "(\i < ?n. norm x ^ i) = (norm x ^ ?n - 1) / (norm x - 1)" by (rule geometric_sum, insert nx, auto) finally have "norm x ^ ?n \ ml * (norm x ^ ?n - 1) / (norm x - 1)" by simp from mult_left_mono[OF this, of "norm x - 1"] have "(norm x - 1) * (norm x ^ ?n) \ ml * (norm x ^ ?n - 1)" using nx by auto also have "\ = (ml * (1 - 1 / (norm x ^ ?n))) * norm x ^ ?n" using nx False x0 by (simp add: field_simps) finally have "(norm x - 1) * (norm x ^ ?n) \ (ml * (1 - 1 / (norm x ^ ?n))) * norm x ^ ?n" . from mult_right_le_imp_le[OF this xn0] have "norm x - 1 \ ml * (1 - 1 / (norm x ^ ?n))" by simp hence "norm x \ 1 + ml - ml / (norm x ^ ?n)" by (simp add: field_simps) also have "\ \ 1 + ml" using ml0 xn0 by auto finally show ?thesis unfolding ml_def . qed qed lemma div_le_div_ceiling: "x div y \ div_ceiling x y" unfolding div_ceiling_def Let_def by auto lemma div_ceiling: assumes q: "q \ 0" shows "(of_int x :: 'a :: floor_ceiling) / of_int q \ of_int (div_ceiling x q)" proof (cases "q dvd x") case True then obtain k where xqk: "x = q * k" unfolding dvd_def by auto hence id: "div_ceiling x q = k" unfolding div_ceiling_def Let_def using q by auto show ?thesis unfolding id unfolding xqk using q by simp next case False { assume "x div q * q = x" hence "x = q * (x div q)" by (simp add: ac_simps) hence "q dvd x" unfolding dvd_def by auto with False have False by simp } hence id: "div_ceiling x q = x div q + 1" unfolding div_ceiling_def Let_def using q by auto show ?thesis unfolding id by (metis floor_divide_of_int_eq le_less add1_zle_eq floor_less_iff) qed lemma max_list_non_empty_map: assumes hom: "\ x y. max (f x) (f y) = f (max x y)" shows "xs \ [] \ max_list_non_empty (map f xs) = f (max_list_non_empty xs)" by (induct xs rule: max_list_non_empty.induct, auto simp: hom) lemma root_bound: assumes "root_bound p = B" and deg: "degree p > 0" shows "ipoly p (x :: real) = 0 \ norm x \ of_rat B" "B \ 0" proof - let ?r = real_of_rat let ?i = real_of_int let ?p = "real_of_int_poly p" define n where "n = degree p" let ?lc = "coeff p n" let ?list = "map (\i. abs (coeff p i)) [0.. 0" by auto from p0 have alc0: "abs ?lc \ 0" unfolding n_def by auto from deg have mem: "abs (coeff p 0) \ set ?list" unfolding n_def by auto from max_list_non_empty[OF this, folded m_def] have m0: "m \ 0" by auto have "div_ceiling m (abs ?lc) \ 0" by (rule order_trans[OF _ div_le_div_ceiling[of m "abs ?lc"]], subst pos_imp_zdiv_nonneg_iff, insert p0 m0, auto simp: n_def) hence mup: "m_up \ 1" unfolding m_up_def by auto have "m_up \ 2 ^ (log_ceiling 2 m_up)" using mup log_ceiling_sound(1) by auto hence Cmup: "C \ of_int m_up" unfolding C_def by linarith with mup have C: "C \ 1" by auto from assms(1)[unfolded root_bound_def Let_def] have B: "C = of_rat B" unfolding C_def m_up_def n_def m_def by auto note dc = div_le_div_ceiling[of m "abs ?lc"] with C show "B \ 0" unfolding B by auto assume "ipoly p x = 0" hence rt: "poly ?p x = 0" by simp from root_imp_deg_nonzero[OF _ this] p0 have n0: "n \ 0" unfolding n_def by auto from cauchy_root_bound[OF rt] p0 have "norm x \ 1 + max_list_non_empty ?list' / ?i (abs ?lc)" by (simp add: n_def) also have "?list' = map ?i ?list" by simp also have "max_list_non_empty \ = ?i m" unfolding m_def by (rule max_list_non_empty_map, insert mem, auto) also have "1 + m / ?i (abs ?lc) \ ?i m_up" unfolding m_up_def using div_ceiling[OF alc0, of m] by auto also have "\ \ ?r C" using Cmup using of_rat_less_eq by force finally have "norm x \ ?r C" . thus "norm x \ ?r B" unfolding B by simp qed fun pairwise_disjoint :: "'a set list \ bool" where "pairwise_disjoint [] = True" | "pairwise_disjoint (x # xs) = ((x \ (\ y \ set xs. y) = {}) \ pairwise_disjoint xs)" lemma roots_of_2_irr: assumes pc: "poly_cond p" and deg: "degree p > 0" shows "real_of_2 ` set (roots_of_2_irr p) = {x. ipoly p x = 0}" (is ?one) "Ball (set (roots_of_2_irr p)) invariant_2" (is ?two) "distinct (map real_of_2 (roots_of_2_irr p))" (is ?three) proof - note d = roots_of_2_irr_def from poly_condD[OF pc] have mon: "lead_coeff p > 0" and irr: "irreducible p" by auto let ?norm = "real_alg_2'" have "?one \ ?two \ ?three" proof (cases "degree p = 1") case True define c where "c = coeff p 0" define d where "d = coeff p 1" from True have rr: "roots_of_2_irr p = [Rational (Rat.Fract (- c) (d))]" unfolding d d_def c_def by auto from degree1_coeffs[OF True] have p: "p = [:c,d:]" and d: "d \ 0" unfolding c_def d_def by auto have *: "real_of_int c + x * real_of_int d = 0 \ x = - (real_of_int c / real_of_int d)" for x using d by (simp add: field_simps) show ?thesis unfolding rr using d * unfolding p using of_rat_1[of "Rat.Fract (- c) (d)"] by (auto simp: Fract_of_int_quotient hom_distribs) next case False let ?r = real_of_rat let ?rp = "map_poly ?r" let ?rr = "set (roots_of_2_irr p)" define ri where "ri = root_info p" define cr where "cr = root_info.l_r ri" define bnds where "bnds = [(-root_bound p, root_bound p)]" define empty where "empty = (Nil :: real_alg_2 list)" have empty: "Ball (set empty) invariant_2 \ distinct (map real_of_2 empty)" unfolding empty_def by auto from mon have p: "p \ 0" by auto from root_info[OF irr deg] have ri: "root_info_cond ri p" unfolding ri_def . from False have rr: "roots_of_2_irr p = roots_of_2_main p ri cr bnds empty" unfolding d ri_def cr_def Let_def bnds_def empty_def by auto note root_bound = root_bound[OF refl deg] from root_bound(2) have bnds: "\ l r. (l,r) \ set bnds \ l \ r" unfolding bnds_def by auto have "ipoly p x = 0 \ ?r (- root_bound p) \ x \ x \ ?r (root_bound p)" for x using root_bound(1)[of x] by (auto simp: hom_distribs) hence rts: "{x. ipoly p x = 0} = real_of_2 ` set empty \ {x. \ l r. root_cond (p,l,r) x \ (l,r) \ set bnds}" unfolding empty_def bnds_def by (force simp: root_cond_def) define rts where "rts lr = Collect (root_cond (p,lr))" for lr have disj: "pairwise_disjoint (real_of_2 ` set empty # map rts bnds)" unfolding empty_def bnds_def by auto from deg False have deg1: "degree p > 1" by auto define delta where "delta = ipoly_root_delta p" note delta = ipoly_root_delta[OF p, folded delta_def] define rel' where "rel' = ({(x, y). 0 \ y \ delta_gt delta x y})^-1" define mm where "mm = (\bnds. mset (map (\ (l,r). ?r r - ?r l) bnds))" define rel where "rel = inv_image (mult1 rel') mm" have wf: "wf rel" unfolding rel_def rel'_def by (rule wf_inv_image[OF wf_mult1[OF SN_imp_wf[OF delta_gt_SN[OF delta(1)]]]]) let ?main = "roots_of_2_main p ri cr" have "real_of_2 ` set (?main bnds empty) = real_of_2 ` set empty \ {x. \l r. root_cond (p, l, r) x \ (l, r) \ set bnds} \ Ball (set (?main bnds empty)) invariant_2 \ distinct (map real_of_2 (?main bnds empty))" (is "?one' \ ?two' \ ?three'") using empty bnds disj proof (induct bnds arbitrary: empty rule: wf_induct[OF wf]) case (1 lrss rais) note rais = 1(2)[rule_format] note lrs = 1(3) note disj = 1(4) note IH = 1(1)[rule_format] note simp = roots_of_2_main.simps[of p ri cr lrss rais] show ?case proof (cases lrss) case Nil with rais show ?thesis unfolding simp by auto next case (Cons lr lrs) obtain l r where lr': "lr = (l,r)" by force { fix lr' assume lt: "\ l' r'. (l',r') \ set lr' \ l' \ r' \ delta_gt delta (?r r - ?r l) (?r r' - ?r l')" have l: "mm (lr' @ lrs) = mm lrs + mm lr'" unfolding mm_def by (auto simp: ac_simps) have r: "mm lrss = mm lrs + {# ?r r - ?r l #}" unfolding Cons lr' rel_def mm_def by auto have "(mm (lr' @ lrs), mm lrss) \ mult1 rel'" unfolding l r mult1_def proof (rule, unfold split, intro exI conjI, unfold add_mset_add_single[symmetric], rule refl, rule refl, intro allI impI) fix d assume "d \# mm lr'" then obtain l' r' where d: "d = ?r r' - ?r l'" and lr': "(l',r') \ set lr'" unfolding mm_def in_multiset_in_set by auto from lt[OF lr'] show "(d, ?r r - ?r l) \ rel'" unfolding d rel'_def by (auto simp: of_rat_less_eq) qed hence "(lr' @ lrs, lrss) \ rel" unfolding rel_def by auto } note rel = this from rel[of Nil] have easy_rel: "(lrs,lrss) \ rel" by auto define c where "c = cr l r" from simp Cons lr' have simp: "?main lrss rais = (if c = 0 then ?main lrs rais else if c = 1 then ?main lrs (real_alg_2' ri p l r # rais) else let m = (l + r) / 2 in ?main ((m, r) # (l, m) # lrs) rais)" unfolding c_def simp Cons lr' using real_alg_2''[OF False] by auto note lrs = lrs[unfolded Cons lr'] from lrs have lr: "l \ r" by auto from root_info_condD(1)[OF ri lr, folded cr_def] have c: "c = card {x. root_cond (p,l,r) x}" unfolding c_def by auto let ?rt = "\ lrs. {x. \l r. root_cond (p, l, r) x \ (l, r) \ set lrs}" have rts: "?rt lrss = ?rt lrs \ {x. root_cond (p,l,r) x}" (is "?rt1 = ?rt2 \ ?rt3") unfolding Cons lr' by auto show ?thesis proof (cases "c = 0") case True with simp have simp: "?main lrss rais = ?main lrs rais" by simp from disj have disj: "pairwise_disjoint (real_of_2 ` set rais # map rts lrs)" unfolding Cons by auto from finite_ipoly_roots[OF p] True[unfolded c] have empty: "?rt3 = {}" unfolding root_cond_def[abs_def] split by simp with rts have rts: "?rt1 = ?rt2" by auto show ?thesis unfolding simp rts by (rule IH[OF easy_rel rais lrs disj], auto) next case False show ?thesis proof (cases "c = 1") case True let ?rai = "real_alg_2' ri p l r" from True simp have simp: "?main lrss rais = ?main lrs (?rai # rais)" by auto from card_1_Collect_ex1[OF c[symmetric, unfolded True]] have ur: "unique_root (p,l,r)" . from real_alg_2'[OF ur pc ri] have rai: "invariant_2 ?rai" "real_of_2 ?rai = the_unique_root (p, l, r)" by auto with rais have rais: "\ x. x \ set (?rai # rais) \ invariant_2 x" and dist: "distinct (map real_of_2 rais)" by auto have rt3: "?rt3 = {real_of_2 ?rai}" using rc1 ur rai by (auto intro: the_unique_root_eqI theI') have "real_of_2 ` set (roots_of_2_main p ri cr lrs (?rai # rais)) = real_of_2 ` set (?rai # rais) \ ?rt2 \ Ball (set (roots_of_2_main p ri cr lrs (?rai # rais))) invariant_2 \ distinct (map real_of_2 (roots_of_2_main p ri cr lrs (?rai # rais)))" (is "?one \ ?two \ ?three") proof (rule IH[OF easy_rel, of "?rai # rais", OF conjI lrs]) show "Ball (set (real_alg_2' ri p l r # rais)) invariant_2" using rais by auto have "real_of_2 (real_alg_2' ri p l r) \ set (map real_of_2 rais)" using disj rt3 unfolding Cons lr' rts_def by auto thus "distinct (map real_of_2 (real_alg_2' ri p l r # rais))" using dist by auto show "pairwise_disjoint (real_of_2 ` set (real_alg_2' ri p l r # rais) # map rts lrs)" using disj rt3 unfolding Cons lr' rts_def by auto qed auto hence ?one ?two ?three by blast+ show ?thesis unfolding simp rts rt3 by (rule conjI[OF _ conjI[OF \?two\ \?three\]], unfold \?one\, auto) next case False let ?m = "(l+r)/2" let ?lrs = "[(?m,r),(l,?m)] @ lrs" from False \c \ 0\ have simp: "?main lrss rais = ?main ?lrs rais" unfolding simp by (auto simp: Let_def) from False \c \ 0\ have "c \ 2" by auto from delta(2)[OF this[unfolded c]] have delta: "delta \ ?r (r - l) / 4" by auto have lrs: "\ l r. (l,r) \ set ?lrs \ l \ r" using lr lrs by (fastforce simp: field_simps) have "?r ?m \ \" unfolding Rats_def by blast with poly_cond_degree_gt_1[OF pc deg1, of "?r ?m"] have disj1: "?r ?m \ rts lr" for lr unfolding rts_def root_cond_def by auto have disj2: "rts (?m, r) \ rts (l, ?m) = {}" using disj1[of "(l,?m)"] disj1[of "(?m,r)"] unfolding rts_def root_cond_def by auto have disj3: "(rts (l,?m) \ rts (?m,r)) = rts (l,r)" unfolding rts_def root_cond_def by (auto simp: hom_distribs) have disj4: "real_of_2 ` set rais \ rts (l,r) = {}" using disj unfolding Cons lr' by auto have disj: "pairwise_disjoint (real_of_2 ` set rais # map rts ([(?m, r), (l, ?m)] @ lrs))" using disj disj2 disj3 disj4 by (auto simp: Cons lr') have "(?lrs,lrss) \ rel" proof (rule rel, intro conjI) fix l' r' assume mem: "(l', r') \ set [(?m,r),(l,?m)]" from mem lr show "l' \ r'" by auto from mem have diff: "?r r' - ?r l' = (?r r - ?r l) / 2" by auto (metis eq_diff_eq minus_diff_eq mult_2_right of_rat_add of_rat_diff, metis of_rat_add of_rat_mult of_rat_numeral_eq) show "delta_gt delta (?r r - ?r l) (?r r' - ?r l')" unfolding diff delta_gt_def by (rule order.trans[OF delta], insert lr, auto simp: field_simps of_rat_diff of_rat_less_eq) qed note IH = IH[OF this, of rais, OF rais lrs disj] have "real_of_2 ` set (?main ?lrs rais) = real_of_2 ` set rais \ ?rt ?lrs \ Ball (set (?main ?lrs rais)) invariant_2 \ distinct (map real_of_2 (?main ?lrs rais))" (is "?one \ ?two") by (rule IH) hence ?one ?two by blast+ have cong: "\ a b c. b = c \ a \ b = a \ c" by auto have id: "?rt ?lrs = ?rt lrs \ ?rt [(?m,r),(l,?m)]" by auto show ?thesis unfolding rts simp \?one\ id proof (rule conjI[OF cong[OF cong] conjI]) have "\ x. root_cond (p,l,r) x = (root_cond (p,l,?m) x \ root_cond (p,?m,r) x)" unfolding root_cond_def by (auto simp:hom_distribs) hence id: "Collect (root_cond (p,l,r)) = {x. (root_cond (p,l,?m) x \ root_cond (p,?m,r) x)}" by auto show "?rt [(?m,r),(l,?m)] = Collect (root_cond (p,l,r))" unfolding id list.simps by blast show "\ a \ set (?main ?lrs rais). invariant_2 a" using \?two\ by auto show "distinct (map real_of_2 (?main ?lrs rais))" using \?two\ by auto qed qed qed qed qed hence idd: "?one'" and cond: ?two' ?three' by blast+ define res where "res = roots_of_2_main p ri cr bnds empty" have e: "set empty = {}" unfolding empty_def by auto from idd[folded res_def] e have idd: "real_of_2 ` set res = {} \ {x. \l r. root_cond (p, l, r) x \ (l, r) \ set bnds}" by auto show ?thesis unfolding rr unfolding rts id e norm_def using cond unfolding res_def[symmetric] image_empty e idd[symmetric] by auto qed thus ?one ?two ?three by blast+ qed definition roots_of_2 :: "int poly \ real_alg_2 list" where "roots_of_2 p = concat (map roots_of_2_irr (factors_of_int_poly p))" lemma roots_of_2: shows "p \ 0 \ real_of_2 ` set (roots_of_2 p) = {x. ipoly p x = 0}" "Ball (set (roots_of_2 p)) invariant_2" "distinct (map real_of_2 (roots_of_2 p))" proof - let ?rr = "roots_of_2 p" note d = roots_of_2_def note frp1 = factors_of_int_poly { fix q r assume "q \ set ?rr" then obtain s where s: "s \ set (factors_of_int_poly p)" and q: "q \ set (roots_of_2_irr s)" unfolding d by auto from frp1(1)[OF refl s] have "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr[OF this] q have "invariant_2 q" by auto } thus "Ball (set ?rr) invariant_2" by auto { assume p: "p \ 0" have "real_of_2 ` set ?rr = (\ ((\ p. real_of_2 ` set (roots_of_2_irr p)) ` (set (factors_of_int_poly p))))" (is "_ = ?rrr") unfolding d set_concat set_map by auto also have "\ = {x. ipoly p x = 0}" proof - { fix x assume "x \ ?rrr" then obtain q s where s: "s \ set (factors_of_int_poly p)" and q: "q \ set (roots_of_2_irr s)" and x: "x = real_of_2 q" by auto from frp1(1)[OF refl s] have s0: "s \ 0" and pt: "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr[OF pt] q have rt: "ipoly s x = 0" unfolding x by auto from frp1(2)[OF refl p, of x] rt s have rt: "ipoly p x = 0" by auto } moreover { fix x :: real assume rt: "ipoly p x = 0" from rt frp1(2)[OF refl p, of x] obtain s where s: "s \ set (factors_of_int_poly p)" and rt: "ipoly s x = 0" by auto from frp1(1)[OF refl s] have s0: "s \ 0" and ty: "poly_cond s" "degree s > 0" by (auto simp: poly_cond_def) from roots_of_2_irr(1)[OF ty] rt obtain q where q: "q \ set (roots_of_2_irr s)" and x: "x = real_of_2 q" by blast have "x \ ?rrr" unfolding x using q s by auto } ultimately show ?thesis by auto qed finally show "real_of_2 ` set ?rr = {x. ipoly p x = 0}" by auto } show "distinct (map real_of_2 (roots_of_2 p))" proof (cases "p = 0") case True from factors_of_int_poly_const[of 0] True show ?thesis unfolding roots_of_2_def by auto next case p: False note frp1 = frp1[OF refl] let ?fp = "factors_of_int_poly p" let ?cc = "concat (map roots_of_2_irr ?fp)" show ?thesis unfolding roots_of_2_def distinct_conv_nth length_map proof (intro allI impI notI) fix i j assume ij: "i < length ?cc" "j < length ?cc" "i \ j" and id: "map real_of_2 ?cc ! i = map real_of_2 ?cc ! j" from ij id have id: "real_of_2 (?cc ! i) = real_of_2 (?cc ! j)" by auto from nth_concat_diff[OF ij, unfolded length_map] obtain j1 k1 j2 k2 where *: "(j1,k1) \ (j2,k2)" "j1 < length ?fp" "j2 < length ?fp" and "k1 < length (map roots_of_2_irr ?fp ! j1)" "k2 < length (map roots_of_2_irr ?fp ! j2)" "?cc ! i = map roots_of_2_irr ?fp ! j1 ! k1" "?cc ! j = map roots_of_2_irr ?fp ! j2 ! k2" by blast hence **: "k1 < length (roots_of_2_irr (?fp ! j1))" "k2 < length (roots_of_2_irr (?fp ! j2))" "?cc ! i = roots_of_2_irr (?fp ! j1) ! k1" "?cc ! j = roots_of_2_irr (?fp ! j2) ! k2" by auto from * have mem: "?fp ! j1 \ set ?fp" "?fp ! j2 \ set ?fp" by auto from frp1(1)[OF mem(1)] frp1(1)[OF mem(2)] have pc1: "poly_cond (?fp ! j1)" "degree (?fp ! j1) > 0" and pc10: "?fp ! j1 \ 0" and pc2: "poly_cond (?fp ! j2)" "degree (?fp ! j2) > 0" by (auto simp: poly_cond_def) show False proof (cases "j1 = j2") case True with * have neq: "k1 \ k2" by auto from **[unfolded True] id * have "map real_of_2 (roots_of_2_irr (?fp ! j2)) ! k1 = real_of_2 (?cc ! j)" "map real_of_2 (roots_of_2_irr (?fp ! j2)) ! k1 = real_of_2 (?cc ! j)" by auto hence "\ distinct (map real_of_2 (roots_of_2_irr (?fp ! j2)))" unfolding distinct_conv_nth using * ** True by auto with roots_of_2_irr(3)[OF pc2] show False by auto next case neq: False with frp1(4)[of p] * have neq: "?fp ! j1 \ ?fp ! j2" unfolding distinct_conv_nth by auto let ?x = "real_of_2 (?cc ! i)" define x where "x = ?x" from ** have "x \ real_of_2 ` set (roots_of_2_irr (?fp ! j1))" unfolding x_def by auto with roots_of_2_irr(1)[OF pc1] have x1: "ipoly (?fp ! j1) x = 0" by auto from ** id have "x \ real_of_2 ` set (roots_of_2_irr (?fp ! j2))" unfolding x_def by (metis image_eqI nth_mem) with roots_of_2_irr(1)[OF pc2] have x2: "ipoly (?fp ! j2) x = 0" by auto have "ipoly p x = 0" using x1 mem unfolding roots_of_2_def by (metis frp1(2) p) from frp1(3)[OF p this] x1 x2 neq mem show False by blast qed qed qed qed lift_definition roots_of_3 :: "int poly \ real_alg_3 list" is roots_of_2 by (insert roots_of_2, auto simp: list_all_iff) lemma roots_of_3: shows "p \ 0 \ real_of_3 ` set (roots_of_3 p) = {x. ipoly p x = 0}" "distinct (map real_of_3 (roots_of_3 p))" proof - show "p \ 0 \ real_of_3 ` set (roots_of_3 p) = {x. ipoly p x = 0}" by (transfer; intro roots_of_2, auto) show "distinct (map real_of_3 (roots_of_3 p))" by (transfer; insert roots_of_2, auto) qed lift_definition roots_of_real_alg :: "int poly \ real_alg list" is roots_of_3 . lemma roots_of_real_alg: "p \ 0 \ real_of ` set (roots_of_real_alg p) = {x. ipoly p x = 0}" "distinct (map real_of (roots_of_real_alg p))" proof - show "p \ 0 \ real_of ` set (roots_of_real_alg p) = {x. ipoly p x = 0}" by (transfer', insert roots_of_3, auto) show "distinct (map real_of (roots_of_real_alg p))" by (transfer, insert roots_of_3(2), auto) qed text \It follows an implementation for @{const roots_of_3}, since the current definition does not provide a code equation.\ context begin private typedef real_alg_2_list = "{xs. Ball (set xs) invariant_2}" by (intro exI[of _ Nil], auto) setup_lifting type_definition_real_alg_2_list private lift_definition roots_of_2_list :: "int poly \ real_alg_2_list" is roots_of_2 by (insert roots_of_2, auto) private lift_definition real_alg_2_list_nil :: "real_alg_2_list \ bool" is "\ xs. case xs of Nil \ True | _ \ False" . private fun real_alg_2_list_hd_intern :: "real_alg_2 list \ real_alg_2" where "real_alg_2_list_hd_intern (Cons x xs) = x" | "real_alg_2_list_hd_intern Nil = of_rat_2 0" private lift_definition real_alg_2_list_hd :: "real_alg_2_list \ real_alg_3" is real_alg_2_list_hd_intern proof (goal_cases) case (1 xs) thus ?case using of_rat_2[of 0] by (cases xs, auto) qed private lift_definition real_alg_2_list_tl :: "real_alg_2_list \ real_alg_2_list" is tl proof (goal_cases) case (1 xs) thus ?case by (cases xs, auto) qed private lift_definition real_alg_2_list_length :: "real_alg_2_list \ nat" is length . private lemma real_alg_2_list_length[simp]: "\ real_alg_2_list_nil xs \ real_alg_2_list_length (real_alg_2_list_tl xs) < real_alg_2_list_length xs" by (transfer, auto split: list.splits) private function real_alg_2_list_convert :: "real_alg_2_list \ real_alg_3 list" where "real_alg_2_list_convert xs = (if real_alg_2_list_nil xs then [] else real_alg_2_list_hd xs # real_alg_2_list_convert (real_alg_2_list_tl xs))" by pat_completeness auto termination by (relation "measure real_alg_2_list_length", auto) private definition roots_of_3_impl :: "int poly \ real_alg_3 list" where "roots_of_3_impl p = real_alg_2_list_convert (roots_of_2_list p)" private lift_definition real_alg_2_list_convert_id :: "real_alg_2_list \ real_alg_3 list" is id by (auto simp: list_all_iff) lemma real_alg_2_list_convert: "real_alg_2_list_convert xs = real_alg_2_list_convert_id xs" proof (induct xs rule: wf_induct[OF wf_measure[of real_alg_2_list_length], rule_format]) case (1 xs) show ?case proof (cases "real_alg_2_list_nil xs") case True hence "real_alg_2_list_convert xs = []" by auto also have "[] = real_alg_2_list_convert_id xs" using True by (transfer', auto split: list.splits) finally show ?thesis . next case False hence "real_alg_2_list_convert xs = real_alg_2_list_hd xs # real_alg_2_list_convert (real_alg_2_list_tl xs)" by simp also have "real_alg_2_list_convert (real_alg_2_list_tl xs) = real_alg_2_list_convert_id (real_alg_2_list_tl xs)" by (rule 1, insert False, simp) also have "real_alg_2_list_hd xs # \ = real_alg_2_list_convert_id xs" using False by (transfer', auto split: list.splits) finally show ?thesis . qed qed lemma roots_of_3_code[code]: "roots_of_3 p = roots_of_3_impl p" unfolding roots_of_3_impl_def real_alg_2_list_convert by (transfer, simp) end definition real_roots_of_int_poly :: "int poly \ real list" where "real_roots_of_int_poly p = map real_of (roots_of_real_alg p)" definition real_roots_of_rat_poly :: "rat poly \ real list" where "real_roots_of_rat_poly p = map real_of (roots_of_real_alg (snd (rat_to_int_poly p)))" abbreviation rpoly :: "rat poly \ 'a :: field_char_0 \ 'a" where "rpoly f \ poly (map_poly of_rat f)" lemma real_roots_of_int_poly: "p \ 0 \ set (real_roots_of_int_poly p) = {x. ipoly p x = 0}" "distinct (real_roots_of_int_poly p)" unfolding real_roots_of_int_poly_def using roots_of_real_alg[of p] by auto lemma real_roots_of_rat_poly: "p \ 0 \ set (real_roots_of_rat_poly p) = {x. rpoly p x = 0}" "distinct (real_roots_of_rat_poly p)" proof - obtain c q where cq: "rat_to_int_poly p = (c,q)" by force from rat_to_int_poly[OF this] have pq: "p = smult (inverse (of_int c)) (of_int_poly q)" and c: "c \ 0" by auto have id: "{x. rpoly p x = (0 :: real)} = {x. ipoly q x = 0}" unfolding pq by (simp add: c of_rat_of_int_poly hom_distribs) show "distinct (real_roots_of_rat_poly p)" unfolding real_roots_of_rat_poly_def cq snd_conv using roots_of_real_alg(2)[of q] . assume "p \ 0" with pq c have q: "q \ 0" by auto show "set (real_roots_of_rat_poly p) = {x. rpoly p x = 0}" unfolding id unfolding real_roots_of_rat_poly_def cq snd_conv using roots_of_real_alg(1)[OF q] by auto qed text \The upcoming functions no longer demand an integer or rational polynomial as input.\ definition roots_of_real_main :: "real poly \ real list" where "roots_of_real_main p \ let n = degree p in if n = 0 then [] else if n = 1 then [roots1 p] else if n = 2 then rroots2 p else (real_roots_of_rat_poly (map_poly to_rat p))" definition roots_of_real_poly :: "real poly \ real list option" where "roots_of_real_poly p \ let (c,pis) = yun_factorization gcd p in if (c \ 0 \ (\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then Some (concat (map (roots_of_real_main o fst) pis)) else None" lemma roots_of_real_main: assumes p: "p \ 0" and deg: "degree p \ 2 \ set (coeffs p) \ \" shows "set (roots_of_real_main p) = {x. poly p x = 0}" (is "?l = ?r") proof - note d = roots_of_real_main_def Let_def show ?thesis proof (cases "degree p = 0") case True hence "?l = {}" unfolding d by auto with roots0[OF p True] show ?thesis by auto next case False note 0 = this show ?thesis proof (cases "degree p = 1") case True hence "?l = {roots1 p}" unfolding d by auto with roots1[OF True] show ?thesis by auto next case False note 1 = this show ?thesis proof (cases "degree p = 2") case True hence "?l = set (rroots2 p)" unfolding d by auto with rroots2[OF True] show ?thesis by auto next case False note 2 = this let ?q = "map_poly to_rat p" from 0 1 2 have l: "?l = set (real_roots_of_rat_poly ?q)" unfolding d by auto from deg 0 1 2 have rat: "set (coeffs p) \ \" by auto have "p = map_poly (of_rat o to_rat) p" by (rule sym, rule map_poly_idI, insert rat, auto) also have "\ = real_of_rat_poly ?q" by (subst map_poly_map_poly, auto simp: to_rat) finally have id: "{x. poly p x = 0} = {x. poly (real_of_rat_poly ?q) x = 0}" and q: "?q \ 0" using p by auto from real_roots_of_rat_poly(1)[OF q, folded id l] show ?thesis by simp qed qed qed qed lemma roots_of_real_poly: assumes rt: "roots_of_real_poly p = Some xs" shows "set xs = {x. poly p x = 0}" proof - obtain c pis where yun: "yun_factorization gcd p = (c,pis)" by force from rt[unfolded roots_of_real_poly_def yun split Let_def] have c: "c \ 0" and pis: "\ p i. (p, i)\set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" and xs: "xs = concat (map (roots_of_real_main \ fst) pis)" by (auto split: if_splits) note yun = square_free_factorizationD(1,2,4)[OF yun_factorization(1)[OF yun]] from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . have "{x. poly p x = 0} = {x. poly (\(a, i)\set pis. a ^ Suc i) x = 0}" unfolding p using c by auto also have "\ = \ ((\ p. {x. poly p x = 0}) ` fst ` set pis)" (is "_ = ?r") by (subst poly_prod_0, force+) finally have r: "{x. poly p x = 0} = ?r" . { fix p i assume p: "(p,i) \ set pis" have "set (roots_of_real_main p) = {x. poly p x = 0}" by (rule roots_of_real_main, insert yun(2)[OF p] pis[OF p], auto) } note main = this have "set xs = \ ((\ (p, i). set (roots_of_real_main p)) ` set pis)" unfolding xs o_def by auto also have "\ = ?r" using main by auto finally show ?thesis unfolding r by simp qed end diff --git a/thys/Algebraic_Numbers/Resultant.thy b/thys/Algebraic_Numbers/Resultant.thy --- a/thys/Algebraic_Numbers/Resultant.thy +++ b/thys/Algebraic_Numbers/Resultant.thy @@ -1,1078 +1,1081 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) subsection \Resultant\ text \This theory contains facts about resultants which are required for addition and multiplication of algebraic numbers. The results are taken from the textbook \cite[pages 227ff and 235ff]{AlgNumbers}. \ theory Resultant imports - Polynomial_Factorization.Rational_Factorization - Subresultants.Subresultant_Gcd + "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" (* for lmpoly_base_conv *) + Subresultants.Resultant_Prelim Berlekamp_Zassenhaus.Unique_Factorization_Poly Bivariate_Polynomials begin subsubsection \Sylvester matrices and vector representation of polynomials\ definition vec_of_poly_rev_shifted where "vec_of_poly_rev_shifted p n j \ vec n (\i. if i \ j \ j \ degree p + i then coeff p (degree p + i - j) else 0)" lemma vec_of_poly_rev_shifted_dim[simp]: "dim_vec (vec_of_poly_rev_shifted p n j) = n" unfolding vec_of_poly_rev_shifted_def by auto lemma col_sylvester: fixes p q defines "m \ degree p" and "n \ degree q" assumes j: "j < m+n" shows "col (sylvester_mat p q) j = vec_of_poly_rev_shifted p n j @\<^sub>v vec_of_poly_rev_shifted q m j" (is "?l = ?r") proof note [simp] = m_def[symmetric] n_def[symmetric] show "dim_vec ?l = dim_vec ?r" by simp fix i assume "i < dim_vec ?r" hence 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_index_mat) using i apply simp using j apply simp apply (cases "i < n") apply force using i by simp qed lemma inj_on_diff_nat2: "inj_on (\i. (n::nat) - i) {..n}" by (rule inj_onI, auto) lemma image_diff_atMost: "(\i. (n::nat) - i) ` {..n} = {..n}" (is "?l = ?r") unfolding set_eq_iff proof (intro allI iffI) fix x assume x: "x \ ?r" thus "x \ ?l" unfolding image_def mem_Collect_eq by(intro bexI[of _ "n-x"],auto) qed auto lemma sylvester_sum_mat_upper: fixes p q :: "'a :: comm_semiring_1 poly" defines "m \ degree p" and "n \ degree q" assumes i: "i < n" shows "(\j 1" using i by auto define ni1 where "ni1 = n-Suc i" hence ni1: "n-i = Suc ni1" using i by auto define l where "l = m+n-1" hence l: "Suc l = m+n" using n1 by auto let ?g = "\j. monom (coeff (monom 1 (n-Suc i) * p) j) j" let ?p = "\j. l-j" have "sum ?f {..l" have "?f j = ((\j. monom (coeff (monom 1 (n-i) * p) (Suc j)) j) \ ?p) j" apply(subst sylvester_index_mat2) using i j unfolding l_def m_def[symmetric] n_def[symmetric] by (auto simp add: Suc_diff_Suc) also have "... = (?g \ ?p) j" unfolding ni1 unfolding coeff_monom_Suc unfolding ni1_def using i by auto finally have "?f j = (?g \ ?p) j". } hence "(\j\l. ?f j) = (\j\l. (?g\?p) j)" using l by auto also have "... = (\j\l. ?g j)" unfolding l_def using sum.reindex[OF inj_on_diff_nat2,symmetric,unfolded image_diff_atMost]. also have "degree ?r \ l" using degree_mult_le[of "monom 1 (n-Suc i)" p] unfolding l_def m_def unfolding degree_monom_eq[OF one_neq_zero] using i by auto from poly_as_sum_of_monoms'[OF this] have "(\j\l. ?g j) = ?r". finally show ?thesis. qed lemma sylvester_sum_mat_lower: fixes p q :: "'a :: comm_semiring_1 poly" defines "m \ degree p" and "n \ degree q" assumes ni: "n \ i" and imn: "i < m+n" shows "(\jl" have "?f j = ((\j. monom (coeff (monom 1 (m+n-i) * q) (Suc j)) j) \ ?p) j" apply(subst sylvester_index_mat2) using ni imn j unfolding l_def m_def[symmetric] n_def[symmetric] by (auto simp add: Suc_diff_Suc) also have "... = (?g \ ?p) j" unfolding mni1 unfolding coeff_monom_Suc unfolding mni1_def.. finally have "?f j = ...". } hence "(\j\l. ?f j) = (\j\l. (?g\?p) j)" by auto also have "... = (\j\l. ?g j)" using sum.reindex[OF inj_on_diff_nat2,symmetric,unfolded image_diff_atMost]. also have "degree ?r \ l" using degree_mult_le[of "monom 1 (m+n-1-i)" q] unfolding l_def n_def[symmetric] unfolding degree_monom_eq[OF one_neq_zero] using ni imn by auto from poly_as_sum_of_monoms'[OF this] have "(\j\l. ?g j) = ?r". finally show ?thesis. qed definition "vec_of_poly p \ let m = degree p in vec (Suc m) (\i. coeff p (m-i))" definition "poly_of_vec v \ let d = dim_vec v in \iv n) = 0" unfolding poly_of_vec_def Let_def by auto lemma poly_of_vec_0_iff[simp]: fixes v :: "'a :: comm_monoid_add vec" shows "poly_of_vec v = 0 \ v = 0\<^sub>v (dim_vec v)" (is "?v = _ \ _ = ?z") proof assume "?v = 0" hence "\i\{..i. i < dim_vec v \ v $ (dim_vec v - Suc i) = 0" by auto { fix i assume "i < dim_vec v" hence "v $ i = 0" using a[of "dim_vec v - Suc i"] by auto } thus "v = ?z" by auto next assume r: "v = ?z" show "?v = 0" apply (subst r) by auto qed (* TODO: move, copied from no longer existing Cayley-Hamilton/Polynomial_extension *) lemma degree_sum_smaller: assumes "n > 0" "finite A" shows "(\ x. x \A \ degree (f x) < n) \ degree (\x\A. f x) < n" using \finite A\ by(induct rule: finite_induct) (simp_all add: degree_add_less assms) lemma degree_poly_of_vec_less: fixes v :: "'a :: comm_monoid_add vec" assumes dim: "dim_vec v > 0" shows "degree (poly_of_vec v) < dim_vec v" unfolding poly_of_vec_def Let_def apply(rule degree_sum_smaller) using dim apply force apply force unfolding lessThan_iff by (metis degree_0 degree_monom_eq dim monom_eq_0_iff) lemma coeff_poly_of_vec: "coeff (poly_of_vec v) i = (if i < dim_vec v then v $ (dim_vec v - Suc i) else 0)" (is "?l = ?r") proof - have "?l = (\x = ?r" proof (cases "i < dim_vec v") case False show ?thesis by (subst sum.neutral, insert False, auto) next case True show ?thesis by (subst sum.remove[of _ i], force, force simp: True, subst sum.neutral, insert True, auto) qed finally show ?thesis . qed lemma vec_of_poly_rev_shifted_scalar_prod: fixes p v defines "q \ poly_of_vec v" assumes m[simp]: "degree p = m" and n: "dim_vec v = n" assumes j: "j < m+n" shows "vec_of_poly_rev_shifted p n (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" from coeff_eq_0[OF this[folded m]] have "?F x = 0" by simp } 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" from coeff_eq_0[OF this[folded m]] have "?G x = 0" by simp } 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_vec_poly: fixes p q :: "'a :: comm_semiring_0 poly" defines "m \ degree p" and "n \ degree q" assumes v: "v \ carrier_vec (m+n)" shows "poly_of_vec (transpose_mat (sylvester_mat p q) *\<^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 note mn[simp] = m_def[symmetric] n_def[symmetric] let ?Tv = "transpose_mat (sylvester_mat p q) *\<^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 unfolding m_def by auto hence "?f x = 0" by auto } hence "sum ?f {..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) unfolding mn 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,simp,simp) using True apply simp apply (subst add.commute[of n m]) apply (subst vec_of_poly_rev_shifted_scalar_prod,simp,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 subsubsection \Homomorphism and Resultant\ text \Here we prove Lemma~7.3.1 of the textbook.\ lemma(in comm_ring_hom) resultant_sub_map_poly: fixes p q :: "'a poly" shows "hom (resultant_sub m n p q) = resultant_sub m n (map_poly hom p) (map_poly hom q)" (is "?l = ?r'") proof - let ?mh = "map_poly hom" have "?l = det (sylvester_mat_sub m n (?mh p) (?mh q))" unfolding resultant_sub_def apply(subst sylvester_mat_sub_map[symmetric]) by auto thus ?thesis unfolding resultant_sub_def. qed (* lemma (in comm_ring_hom) resultant_map_poly: fixes p q :: "'a poly" defines "p' \ map_poly hom p" defines "q' \ map_poly hom q" defines "m \ degree p" defines "n \ degree q" defines "m' \ degree p'" defines "n' \ degree q'" defines "r \ resultant p q" defines "r' \ resultant p' q'" shows "m' = m \ n' = n \ hom r = r'" and "m' = m \ hom r = hom (coeff p m')^(n-n') * r'" and "m' \ m \ n' = n \ hom r = (if even n then 1 else (-1)^(m-m')) * hom (coeff q n)^(m-m') * r'" (is "_ \ _ \ ?goal") and "m' \ m \ n' \ n \ hom r = 0" proof - have m'm: "m' \ m" unfolding m_def m'_def p'_def using degree_map_poly_le by auto have n'n: "n' \ n" unfolding n_def n'_def q'_def using degree_map_poly_le by auto have coeffp'[simp]: "\i. coeff p' i = hom (coeff p i)" unfolding p'_def by auto have coeffq'[simp]: "\i. coeff q' i = hom (coeff q i)" unfolding q'_def by auto let ?f = "\i. (if even n then 1 else (-1)^i) * hom (coeff q n)^i" have "hom r = resultant_sub m n p' q'" unfolding r_def resultant_sub unfolding m_def n_def p'_def q'_def by(rule resultant_sub_map_poly) also have "... = ?f (m-m') * resultant_sub m' n p' q'" using resultant_sub_trim_upper[of p' "m-m'" n q',folded m'_def] m'm by (auto simp: power_minus[symmetric]) also have "... = ?f (m-m') * hom (coeff p m')^(n-n') * r'" using resultant_sub_trim_lower[of m' q' "n-n'" p'] n'n unfolding r'_def resultant_sub m'_def n'_def by auto finally have main: "hom r = ?f (m-m') * hom (coeff p m')^(n-n') * r'" by auto { assume "m' = m" thus "hom r = hom (coeff p m')^(n-n') * r'" using main by auto thus "n' = n \ hom r = r'" by auto } assume "m' \ m" hence m'm: "m' < m" using m'm by auto thus "n' = n \ ?goal" using main by simp assume "n' \ n" hence "n' < n" using n'n by auto hence "hom (coeff q n) = 0" unfolding coeffq'[symmetric] unfolding n'_def by(rule coeff_eq_0) hence "hom (coeff q n) ^ (m-m') = 0" using m'm by (simp add: power_0_left) from main[unfolded this] show "hom r = 0" using power_0_Suc by auto qed *) subsubsection\Resultant as Polynomial Expression\ context begin text \This context provides notions for proving Lemma 7.2.1 of the textbook.\ private fun mk_poly_sub where "mk_poly_sub A l 0 = A" | "mk_poly_sub A l (Suc j) = mat_addcol (monom 1 (Suc j)) l (l-Suc j) (mk_poly_sub A l j)" definition "mk_poly A = mk_poly_sub (map_mat coeff_lift A) (dim_col A - 1) (dim_col A - 1)" private lemma mk_poly_sub_dim[simp]: "dim_row (mk_poly_sub A l j) = dim_row A" "dim_col (mk_poly_sub A l j) = dim_col A" by (induct j,auto) private lemma mk_poly_sub_carrier: assumes "A \ carrier_mat nr nc" shows "mk_poly_sub A l j \ carrier_mat nr nc" apply (rule carrier_matI) using assms by auto private lemma mk_poly_dim[simp]: "dim_col (mk_poly A) = dim_col A" "dim_row (mk_poly A) = dim_row A" unfolding mk_poly_def by auto private lemma mk_poly_sub_others[simp]: assumes "l \ j'" and "i < dim_row A" and "j' < dim_col A" shows "mk_poly_sub A l j $$ (i,j') = A $$ (i,j')" using assms by (induct j; simp) private lemma mk_poly_others[simp]: assumes i: "i < dim_row A" and j: "j < dim_col A - 1" shows "mk_poly A $$ (i,j) = [: A $$ (i,j) :]" unfolding mk_poly_def apply(subst mk_poly_sub_others) using i j by auto private lemma mk_poly_delete[simp]: assumes i: "i < dim_row A" shows "mat_delete (mk_poly A) i (dim_col A - 1) = map_mat coeff_lift (mat_delete A i (dim_col A - 1))" apply(rule eq_matI) unfolding mat_delete_def by auto private lemma col_mk_poly_sub[simp]: assumes "l \ j'" and "j' < dim_col A" shows "col (mk_poly_sub A l j) j' = col A j'" by(rule eq_vecI; insert assms; simp) private lemma det_mk_poly_sub: assumes A: "(A :: 'a :: comm_ring_1 poly mat) \ carrier_mat n n" and i: "i < n" shows "det (mk_poly_sub A (n-1) i) = det A" using i proof (induct i) case (Suc i) show ?case unfolding mk_poly_sub.simps apply(subst det_addcol[of _ n]) using Suc apply simp using Suc apply simp apply (rule mk_poly_sub_carrier[OF A]) using Suc by auto qed simp private lemma det_mk_poly: fixes A :: "'a :: comm_ring_1 mat" shows "det (mk_poly A) = [: det A :]" proof (cases "dim_row A = dim_col A") case True define n where "n = dim_col A" have "map_mat coeff_lift A \ carrier_mat (dim_row A) (dim_col A)" by simp hence sq: "map_mat coeff_lift A \ carrier_mat (dim_col A) (dim_col A)" unfolding True. show ?thesis proof(cases "dim_col A = 0") case True thus ?thesis unfolding det_def by simp next case False thus ?thesis unfolding mk_poly_def by (subst det_mk_poly_sub[OF sq]; simp) qed next case False hence f2: "dim_row A = dim_col A \ False" by simp hence f3: "dim_row (mk_poly A) = dim_col (mk_poly A) \ False" unfolding mk_poly_dim by auto show ?thesis unfolding det_def unfolding f2 f3 if_False by simp qed private fun mk_poly2_row where "mk_poly2_row A d j pv 0 = pv" | "mk_poly2_row A d j pv (Suc n) = mk_poly2_row A d j pv n |\<^sub>v n \ pv $ n + monom (A$$(n,j)) d" private fun mk_poly2_col where "mk_poly2_col A pv 0 = pv" | "mk_poly2_col A pv (Suc m) = mk_poly2_row A m (dim_col A - Suc m) (mk_poly2_col A pv m) (dim_row A)" private definition "mk_poly2 A \ mk_poly2_col A (0\<^sub>v (dim_row A)) (dim_col A)" private lemma mk_poly2_row_dim[simp]: "dim_vec (mk_poly2_row A d j pv i) = dim_vec pv" by(induct i arbitrary: pv, auto) private lemma mk_poly2_col_dim[simp]: "dim_vec (mk_poly2_col A pv j) = dim_vec pv" by (induct j arbitrary: pv, auto) private lemma mk_poly2_row: assumes n: "n \ dim_vec pv" shows "mk_poly2_row A d j pv n $ i = (if i < n then pv $ i + monom (A $$ (i,j)) d else pv $ i)" using n proof (induct n arbitrary: pv) case (Suc n) thus ?case unfolding mk_poly2_row.simps by (cases rule: linorder_cases[of "i" "n"],auto) qed simp private lemma mk_poly2_row_col: assumes dim[simp]: "dim_vec pv = n" "dim_row A = n" and j: "j < dim_col A" shows "mk_poly2_row A d j pv n = pv + map_vec (\a. monom a d) (col A j)" apply rule using mk_poly2_row[of _ pv] j by auto private lemma mk_poly2_col: fixes pv :: "'a :: comm_semiring_1 poly vec" and A :: "'a mat" assumes i: "i < dim_row A" and dim: "dim_row A = dim_vec pv" shows "mk_poly2_col A pv j $ i = pv $ i + (\j'j' 0" shows "mk_poly2 A $ i = (\j' ?g) ?S" unfolding l_def mk_poly2_pre[OF i] by auto also have "... = sum ?f ?S" unfolding dim unfolding lessThan_Suc_atMost using sum.reindex[OF inj_on_diff_nat2,symmetric,unfolded image_diff_atMost]. finally show ?thesis. qed private lemma mk_poly2_sylvester_upper: fixes p q :: "'a :: comm_semiring_1 poly" assumes i: "i < degree q" shows "mk_poly2 (sylvester_mat p q) $ i = monom 1 (degree q - Suc i) * p" apply (subst mk_poly2) using i apply simp using i apply simp apply (subst sylvester_sum_mat_upper[OF i,symmetric]) apply (rule sum.cong) unfolding sylvester_mat_dim lessThan_Suc_atMost apply simp by auto private lemma mk_poly2_sylvester_lower: fixes p q :: "'a :: comm_semiring_1 poly" assumes mi: "i \ degree q" and imn: "i < degree p + degree q" shows "mk_poly2 (sylvester_mat p q) $ i = monom 1 (degree p + degree q - Suc i) * q" apply (subst mk_poly2) using imn apply simp using mi imn apply simp unfolding sylvester_mat_dim using sylvester_sum_mat_lower[OF mi imn] apply (subst sylvester_sum_mat_lower) using mi imn by auto private lemma foo: fixes v :: "'a :: comm_semiring_1 vec" shows "monom 1 d \\<^sub>v map_vec coeff_lift v = map_vec (\a. monom a d) v" apply (rule eq_vecI) unfolding index_map_vec index_col by (auto simp add: Polynomial.smult_monom) private lemma mk_poly_sub_corresp: assumes dimA[simp]: "dim_col A = Suc l" and dimpv[simp]: "dim_vec pv = dim_row A" and j: "j < dim_col A" shows "pv + col (mk_poly_sub (map_mat coeff_lift A) l j) l = mk_poly2_col A pv (Suc j)" proof(insert j, induct j) have le: "dim_row A \ dim_vec pv" using dimpv by simp have l: "l < dim_col A" using dimA by simp { case 0 show ?case apply (rule eq_vecI) using mk_poly2_row[OF le] by (auto simp add: monom_0) } { case (Suc j) hence j: "j < dim_col A" by simp show ?case unfolding mk_poly_sub.simps apply(subst col_addcol) apply simp - apply simp + apply simp apply(subst(2) comm_add_vec) apply(rule carrier_vecI, simp) - apply(rule carrier_vecI, simp) + apply(rule carrier_vecI, simp) apply(subst assoc_add_vec[symmetric]) - apply(rule carrier_vecI, rule refl) + apply(rule carrier_vecI, rule refl) apply(rule carrier_vecI, simp) - apply(rule carrier_vecI, simp) + apply(rule carrier_vecI, simp) unfolding Suc(1)[OF j] apply(subst(2) mk_poly2_col.simps) apply(subst mk_poly2_row_col) + apply simp apply simp - apply simp - using Suc apply simp + using Suc apply simp apply(subst col_mk_poly_sub) - using Suc apply simp - using Suc apply simp + using Suc apply simp + using Suc apply simp apply(subst col_map_mat) - using dimA apply simp + using dimA apply simp unfolding foo dimA by simp } qed private lemma col_mk_poly_mk_poly2: fixes A :: "'a :: comm_semiring_1 mat" assumes dim: "dim_col A > 0" shows "col (mk_poly A) (dim_col A - 1) = mk_poly2 A" proof - define l where "l = dim_col A - 1" have dim: "dim_col A = Suc l" unfolding l_def using dim by auto show ?thesis unfolding mk_poly_def mk_poly2_def dim apply(subst mk_poly_sub_corresp[symmetric]) apply(rule dim) apply simp using dim apply simp apply(subst left_zero_vec) apply(rule carrier_vecI) using dim apply simp apply simp done qed private lemma mk_poly_mk_poly2: fixes A :: "'a :: comm_semiring_1 mat" assumes dim: "dim_col A > 0" and i: "i < dim_row A" shows "mk_poly A $$ (i,dim_col A - 1) = mk_poly2 A $ i" proof - have "mk_poly A $$ (i,dim_col A - 1) = col (mk_poly A) (dim_col A - 1) $ i" apply (subst index_col(1)) using dim i by auto also note col_mk_poly_mk_poly2[OF dim] finally show ?thesis. qed lemma mk_poly_sylvester_upper: fixes p q :: "'a :: comm_ring_1 poly" defines "m \ degree p" and "n \ degree q" assumes i: "i < n" shows "mk_poly (sylvester_mat p q) $$ (i, m + n - 1) = monom 1 (n - Suc i) * p" (is "?l = ?r") proof - let ?S = "sylvester_mat p q" have c: "m+n = dim_col ?S" and r: "m+n = dim_row ?S" unfolding m_def n_def by auto hence "dim_col ?S > 0" "i < dim_row ?S" using i by auto from mk_poly_mk_poly2[OF this] have "?l = mk_poly2 (sylvester_mat p q) $ i" unfolding m_def n_def by auto also have "... = ?r" apply(subst mk_poly2_sylvester_upper) using i unfolding n_def m_def by auto finally show ?thesis. qed lemma mk_poly_sylvester_lower: fixes p q :: "'a :: comm_ring_1 poly" defines "m \ degree p" and "n \ degree q" assumes ni: "n \ i" and imn: "i < m+n" shows "mk_poly (sylvester_mat p q) $$ (i, m + n - 1) = monom 1 (m + n - Suc i) * q" (is "?l = ?r") proof - let ?S = "sylvester_mat p q" have c: "m+n = dim_col ?S" and r: "m+n = dim_row ?S" unfolding m_def n_def by auto hence "dim_col ?S > 0" "i < dim_row ?S" using imn by auto from mk_poly_mk_poly2[OF this] have "?l = mk_poly2 (sylvester_mat p q) $ i" unfolding m_def n_def by auto also have "... = ?r" apply(subst mk_poly2_sylvester_lower) using ni imn unfolding n_def m_def by auto finally show ?thesis. qed text \The next lemma corresponds to Lemma 7.2.1.\ lemma resultant_as_poly: fixes p q :: "'a :: comm_ring_1 poly" assumes degp: "degree p > 0" and degq: "degree q > 0" shows "\p' q'. degree p' < degree q \ degree q' < degree p \ [: resultant p q :] = p' * p + q' * q" proof (intro exI conjI) define m where "m = degree p" define n where "n = degree q" define d where "d = dim_row (mk_poly (sylvester_mat p q))" define c where "c = (\i. coeff_lift (cofactor (sylvester_mat p q) i (m+n-1)))" define p' where "p' = (\iii. degree (c i) = 0" unfolding c_def by auto have dmn: "d = m+n" and mnd: "m + n = d" unfolding d_def m_def n_def by auto have "[: resultant p q :] = (\ii {n .. {n..i=n..i. i+n) ` {0.. (\i. i+n)) {0.. {.. {..Resultant as Nonzero Polynomial Expression\ lemma resultant_zero: fixes p q :: "'a :: comm_ring_1 poly" assumes deg: "degree p > 0 \ degree q > 0" and xp: "poly p x = 0" and xq: "poly q x = 0" shows "resultant p q = 0" proof - { assume degp: "degree p > 0" and degq: "degree q > 0" obtain p' q' where "[: resultant p q :] = p' * p + q' * q" using resultant_as_poly[OF degp degq] by force hence "resultant p q = poly (p' * p + q' * q) x" using mpoly_base_conv(2)[of "resultant p q"] by auto also have "... = poly p x * poly p' x + poly q x * poly q' x" unfolding poly2_def by simp finally have ?thesis using xp xq by simp } moreover { assume degp: "degree p = 0" have p: "p = [:0:]" using xp degree_0_id[OF degp,symmetric] by (metis mpoly_base_conv(2)) have ?thesis unfolding p using degp deg by simp } moreover { assume degq: "degree q = 0" have q: "q = [:0:]" using xq degree_0_id[OF degq,symmetric] by (metis mpoly_base_conv(2)) have ?thesis unfolding q using degq deg by simp } ultimately show ?thesis by auto qed lemma poly_resultant_zero: fixes p q :: "'a :: comm_ring_1 poly poly" assumes deg: "degree p > 0 \ degree q > 0" assumes p0: "poly2 p x y = 0" and q0: "poly2 q x y = 0" shows "poly (resultant p q) x = 0" proof - { assume "degree p > 0" "degree q > 0" from resultant_as_poly[OF this] obtain p' q' where "[: resultant p q :] = p' * p + q' * q" by force hence "resultant p q = poly (p' * p + q' * q) [:y:]" using mpoly_base_conv(2)[of "resultant p q"] by auto also have "poly ... x = poly2 p x y * poly2 p' x y + poly2 q x y * poly2 q' x y" unfolding poly2_def by simp finally have ?thesis unfolding p0 q0 by simp } moreover { assume degp: "degree p = 0" hence p: "p = [: coeff p 0 :]" by(subst degree_0_id[OF degp,symmetric],simp) hence "resultant p q = coeff p 0 ^ degree q" using resultant_const(1) by metis also have "poly ... x = poly (coeff p 0) x ^ degree q" by auto also have "... = poly2 p x y ^ degree q" unfolding poly2_def by(subst p, auto) finally have ?thesis unfolding p0 using deg degp zero_power by auto } moreover { assume degq: "degree q = 0" hence q: "q = [: coeff q 0 :]" by(subst degree_0_id[OF degq,symmetric],simp) hence "resultant p q = coeff q 0 ^ degree p" using resultant_const(2) by metis also have "poly ... x = poly (coeff q 0) x ^ degree p" by auto also have "... = poly2 q x y ^ degree p" unfolding poly2_def by(subst q, auto) finally have ?thesis unfolding q0 using deg degq zero_power by auto } ultimately show ?thesis by auto qed lemma resultant_as_nonzero_poly_weak: fixes p q :: "'a :: idom poly" assumes degp: "degree p > 0" and degq: "degree q > 0" and r0: "resultant p q \ 0" shows "\p' q'. degree p' < degree q \ degree q' < degree p \ [: resultant p q :] = p' * p + q' * q \ p' \ 0 \ q' \ 0" proof - obtain p' q' where deg: "degree p' < degree q" "degree q' < degree p" and main: "[: resultant p q :] = p' * p + q' * q" using resultant_as_poly[OF degp degq] by auto have p0: "p \ 0" using degp by auto have q0: "q \ 0" using degq by auto show ?thesis proof (intro exI conjI notI) assume "p' = 0" hence "[: resultant p q :] = q' * q" using main by auto also hence d0: "0 = degree (q' * q)" by (metis degree_pCons_0) { assume "q' \ 0" hence "degree (q' * q) = degree q' + degree q" apply(rule degree_mult_eq) using q0 by auto hence False using d0 degq by auto } hence "q' = 0" by auto finally show False using r0 by auto next assume "q' = 0" hence "[: resultant p q :] = p' * p" using main by auto also hence d0: "0 = degree (p' * p)" by (metis degree_pCons_0) { assume "p' \ 0" hence "degree (p' * p) = degree p' + degree p" apply(rule degree_mult_eq) using p0 by auto hence False using d0 degp by auto } hence "p' = 0" by auto finally show False using r0 by auto qed fact+ qed text \ Next lemma corresponds to Lemma 7.2.2 of the textbook \ lemma resultant_as_nonzero_poly: fixes p q :: "'a :: idom poly" defines "m \ degree p" and "n \ degree q" assumes degp: "m > 0" and degq: "n > 0" shows "\p' q'. degree p' < n \ degree q' < m \ [: resultant p q :] = p' * p + q' * q \ p' \ 0 \ q' \ 0" proof (cases "resultant p q = 0") case False thus ?thesis using resultant_as_nonzero_poly_weak degp degq unfolding m_def n_def by auto next case True define S where "S = transpose_mat (sylvester_mat p q)" have S: "S \ carrier_mat (m+n) (m+n)" unfolding S_def m_def n_def by auto have "det S = 0" using True unfolding resultant_def S_def apply (subst det_transpose) by auto then obtain v where v: "v \ carrier_vec (m+n)" and v0: "v \ 0\<^sub>v (m+n)" and "S *\<^sub>v v = 0\<^sub>v (m+n)" using det_0_iff_vec_prod_zero[OF S] by auto hence "poly_of_vec (S *\<^sub>v v) = 0" by auto hence main: "poly_of_vec (vec_first v n) * p + poly_of_vec (vec_last v m) * q = 0" (is "?p * _ + ?q * _ = _") using sylvester_vec_poly[OF v[unfolded m_def n_def], folded m_def n_def S_def] by auto have split: "vec_first v n @\<^sub>v vec_last v m = v" using vec_first_last_append[simplified add.commute] v by auto show ?thesis proof(intro exI conjI) show "[: resultant p q :] = ?p * p + ?q * q" unfolding True using main by auto show "?p \ 0" proof assume p'0: "?p = 0" hence "?q * q = 0" using main by auto hence "?q = 0" using degq n_def by auto hence "vec_last v m = 0\<^sub>v m" unfolding poly_of_vec_0_iff by auto also have "vec_first v n @\<^sub>v ... = 0\<^sub>v (m+n)" using p'0 unfolding poly_of_vec_0_iff by auto finally have "v = 0\<^sub>v (m+n)" using split by auto thus False using v0 by auto qed show "?q \ 0" proof assume q'0: "?q = 0" hence "?p * p = 0" using main by auto hence "?p = 0" using degp m_def by auto hence "vec_first v n = 0\<^sub>v n" unfolding poly_of_vec_0_iff by auto also have "... @\<^sub>v vec_last v m = 0\<^sub>v (m+n)" using q'0 unfolding poly_of_vec_0_iff by auto finally have "v = 0\<^sub>v (m+n)" using split by auto thus False using v0 by auto qed show "degree ?p < n" using degree_poly_of_vec_less[of "vec_first v n"] using degq by auto show "degree ?q < m" using degree_poly_of_vec_less[of "vec_last v m"] using degp by auto qed qed text\Corresponds to Lemma 7.2.3 of the textbook\ lemma resultant_zero_imp_common_factor: fixes p q :: "'a :: ufd poly" assumes deg: "degree p > 0 \ degree q > 0" and r0: "resultant p q = 0" shows "\ coprime p q" unfolding neq0_conv[symmetric] proof - { assume degp: "degree p > 0" and degq: "degree q > 0" assume cop: "coprime p q" obtain p' q' where "p' * p + q' * q = 0" and p': "degree p' < degree q" and q': "degree q' < degree p" and p'0: "p' \ 0" and q'0: "q' \ 0" using resultant_as_nonzero_poly[OF degp degq] r0 by auto hence "p' * p = - q' * q" by (simp add: eq_neg_iff_add_eq_0) from some_gcd.coprime_mult_cross_dvd[OF cop this] have "p dvd q'" by auto from dvd_imp_degree_le[OF this q'0] have "degree p \ degree q'" by auto hence False using q' by auto } moreover { assume degp: "degree p = 0" then obtain x where "p = [:x:]" by (elim degree_eq_zeroE) moreover hence "resultant p q = x ^ degree q" using resultant_const by auto hence "x = 0" using r0 by auto ultimately have "p = 0" by auto hence ?thesis unfolding not_coprime_iff_common_factor by (metis deg degp dvd_0_right dvd_refl less_numeral_extra(3) poly_dvd_1) } moreover { assume degq: "degree q = 0" then obtain x where "q = [:x:]" by (elim degree_eq_zeroE) moreover hence "resultant p q = x ^ degree p" using resultant_const by auto hence "x = 0" using r0 by auto ultimately have "q = 0" by auto hence ?thesis unfolding not_coprime_iff_common_factor by (metis deg degq dvd_0_right dvd_refl less_numeral_extra(3) poly_dvd_1) } ultimately show ?thesis by auto qed lemma resultant_non_zero_imp_coprime: assumes nz: "resultant (f :: 'a :: field poly) g \ 0" and nz': "f \ 0 \ g \ 0" shows "coprime f g" proof (cases "degree f = 0 \ degree g = 0") case False define r where "r = [:resultant f g:]" from nz have r: "r \ 0" unfolding r_def by auto from False have "degree f > 0" "degree g > 0" by auto from resultant_as_nonzero_poly_weak[OF this nz] obtain p q where "degree p < degree g" "degree q < degree f" and id: "r = p * f + q * g" and "p \ 0" "q \ 0" unfolding r_def by auto define h where "h = some_gcd f g" have "h dvd f" "h dvd g" unfolding h_def by auto then obtain j k where f: "f = h * j" and g: "g = h * k" unfolding dvd_def by auto from id[unfolded f g] have id: "h * (p * j + q * k) = r" by (auto simp: field_simps) from arg_cong[OF id, of degree] have "degree (h * (p * j + q * k)) = 0" unfolding r_def by auto also have "degree (h * (p * j + q * k)) = degree h + degree (p * j + q * k)" by (subst degree_mult_eq, insert id r, auto) finally have h: "degree h = 0" "h \ 0" using r id by auto thus ?thesis unfolding h_def using is_unit_iff_degree some_gcd.gcd_dvd_1 by blast next case True - { - fix f g :: "'a poly" - assume deg_g: "degree g = 0" and res: "resultant f g \ 0" and nz: "f \ 0 \ g \ 0" - have "coprime f g" + thus ?thesis + proof + assume deg_g: "degree g = 0" + show ?thesis proof (cases "g = 0") case False then show ?thesis using divides_degree[of _ g, unfolded deg_g] by (simp add: is_unit_right_imp_coprime) next case g: True then have "g = [:0:]" by auto - from res[unfolded this resultant_const] have "degree f = 0" by auto - with nz show ?thesis unfolding g by auto + from nz[unfolded this resultant_const] have "degree f = 0" by auto + with nz' show ?thesis unfolding g by auto qed - } note main = this - from True - show ?thesis - proof - assume f: "degree f = 0" - from nz[unfolded resultant_swap[of f g]] have "resultant g f \ 0" by (auto split: if_splits) - from main[OF f this] nz' show ?thesis by (auto simp: ac_simps) next - assume "degree g = 0" - from main[OF this nz nz'] show ?thesis . + assume deg_f: "degree f = 0" + show ?thesis + proof (cases "f = 0") + case False + then show ?thesis using divides_degree[of _ f, unfolded deg_f] + by (simp add: is_unit_left_imp_coprime) + next + case f: True + then have "f = [:0:]" by auto + from nz[unfolded this resultant_const] have "degree g = 0" by auto + with nz' show ?thesis unfolding f by auto + qed qed 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,322 @@ 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 + Berlekamp_Zassenhaus.Square_Free_Int_To_Square_Free_GFp begin +hide_const (open) UnivPoly.coeff + (* 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 :: {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