| author | wenzelm | 
| Wed, 28 Feb 2024 22:11:11 +0100 | |
| changeset 79740 | ea1913c953ef | 
| parent 76121 | f58ad163bb75 | 
| child 80084 | 173548e4d5d0 | 
| permissions | -rw-r--r-- | 
| 65435 | 1  | 
(* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy  | 
| 63764 | 2  | 
Author: Manuel Eberl  | 
3  | 
*)  | 
|
4  | 
||
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
5  | 
section \<open>Polynomials, fractions and rings\<close>  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
6  | 
|
| 63498 | 7  | 
theory Polynomial_Factorial  | 
8  | 
imports  | 
|
9  | 
Complex_Main  | 
|
| 65366 | 10  | 
Polynomial  | 
11  | 
Normalized_Fraction  | 
|
| 63498 | 12  | 
begin  | 
13  | 
||
14  | 
subsection \<open>Lifting elements into the field of fractions\<close>  | 
|
15  | 
||
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
16  | 
definition to_fract :: "'a :: idom \<Rightarrow> 'a fract"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
17  | 
where "to_fract x = Fract x 1"  | 
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66806 
diff
changeset
 | 
18  | 
\<comment> \<open>FIXME: more idiomatic name, abbreviation\<close>  | 
| 63498 | 19  | 
|
20  | 
lemma to_fract_0 [simp]: "to_fract 0 = 0"  | 
|
21  | 
by (simp add: to_fract_def eq_fract Zero_fract_def)  | 
|
22  | 
||
23  | 
lemma to_fract_1 [simp]: "to_fract 1 = 1"  | 
|
24  | 
by (simp add: to_fract_def eq_fract One_fract_def)  | 
|
25  | 
||
26  | 
lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y"  | 
|
27  | 
by (simp add: to_fract_def)  | 
|
28  | 
||
29  | 
lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y"  | 
|
30  | 
by (simp add: to_fract_def)  | 
|
31  | 
||
32  | 
lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x"  | 
|
33  | 
by (simp add: to_fract_def)  | 
|
34  | 
||
35  | 
lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y"  | 
|
36  | 
by (simp add: to_fract_def)  | 
|
37  | 
||
38  | 
lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \<longleftrightarrow> x = y"  | 
|
39  | 
by (simp add: to_fract_def eq_fract)  | 
|
40  | 
||
41  | 
lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \<longleftrightarrow> x = 0"  | 
|
42  | 
by (simp add: to_fract_def Zero_fract_def eq_fract)  | 
|
43  | 
||
44  | 
lemma to_fract_quot_of_fract:  | 
|
45  | 
assumes "snd (quot_of_fract x) = 1"  | 
|
46  | 
shows "to_fract (fst (quot_of_fract x)) = x"  | 
|
47  | 
proof -  | 
|
48  | 
have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp  | 
|
49  | 
also note assms  | 
|
50  | 
finally show ?thesis by (simp add: to_fract_def)  | 
|
51  | 
qed  | 
|
52  | 
||
53  | 
lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b"  | 
|
54  | 
by (simp add: to_fract_def)  | 
|
55  | 
||
56  | 
lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)"  | 
|
57  | 
unfolding to_fract_def by transfer (simp add: normalize_quot_def)  | 
|
58  | 
||
59  | 
lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1"  | 
|
60  | 
unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all  | 
|
61  | 
||
62  | 
||
63  | 
subsection \<open>Lifting polynomial coefficients to the field of fractions\<close>  | 
|
64  | 
||
| 76121 | 65  | 
abbreviation (input) fract_poly :: \<open>'a::idom poly \<Rightarrow> 'a fract poly\<close>  | 
| 63498 | 66  | 
where "fract_poly \<equiv> map_poly to_fract"  | 
67  | 
||
| 76121 | 68  | 
abbreviation (input) unfract_poly :: \<open>'a::{ring_gcd,semiring_gcd_mult_normalize,idom_divide} fract poly \<Rightarrow> 'a poly\<close>
 | 
| 63498 | 69  | 
where "unfract_poly \<equiv> map_poly (fst \<circ> quot_of_fract)"  | 
| 76121 | 70  | 
|
| 63498 | 71  | 
lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)"  | 
72  | 
by (simp add: smult_conv_map_poly map_poly_map_poly o_def)  | 
|
73  | 
||
74  | 
lemma fract_poly_0 [simp]: "fract_poly 0 = 0"  | 
|
75  | 
by (simp add: poly_eqI coeff_map_poly)  | 
|
76  | 
||
77  | 
lemma fract_poly_1 [simp]: "fract_poly 1 = 1"  | 
|
| 65486 | 78  | 
by (simp add: map_poly_pCons)  | 
| 63498 | 79  | 
|
80  | 
lemma fract_poly_add [simp]:  | 
|
81  | 
"fract_poly (p + q) = fract_poly p + fract_poly q"  | 
|
82  | 
by (intro poly_eqI) (simp_all add: coeff_map_poly)  | 
|
83  | 
||
84  | 
lemma fract_poly_diff [simp]:  | 
|
85  | 
"fract_poly (p - q) = fract_poly p - fract_poly q"  | 
|
86  | 
by (intro poly_eqI) (simp_all add: coeff_map_poly)  | 
|
87  | 
||
| 64267 | 88  | 
lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\<lambda>x. to_fract (f x)) A"  | 
| 63498 | 89  | 
by (cases "finite A", induction A rule: finite_induct) simp_all  | 
90  | 
||
91  | 
lemma fract_poly_mult [simp]:  | 
|
92  | 
"fract_poly (p * q) = fract_poly p * fract_poly q"  | 
|
93  | 
by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult)  | 
|
94  | 
||
95  | 
lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \<longleftrightarrow> p = q"  | 
|
96  | 
by (auto simp: poly_eq_iff coeff_map_poly)  | 
|
97  | 
||
98  | 
lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \<longleftrightarrow> p = 0"  | 
|
99  | 
using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff)  | 
|
100  | 
||
101  | 
lemma fract_poly_dvd: "p dvd q \<Longrightarrow> fract_poly p dvd fract_poly q"  | 
|
| 76121 | 102  | 
by auto  | 
| 63498 | 103  | 
|
| 63830 | 104  | 
lemma prod_mset_fract_poly:  | 
| 65390 | 105  | 
"(\<Prod>x\<in>#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))"  | 
106  | 
by (induct A) (simp_all add: ac_simps)  | 
|
| 63498 | 107  | 
|
108  | 
lemma is_unit_fract_poly_iff:  | 
|
109  | 
"p dvd 1 \<longleftrightarrow> fract_poly p dvd 1 \<and> content p = 1"  | 
|
110  | 
proof safe  | 
|
111  | 
assume A: "p dvd 1"  | 
|
| 65389 | 112  | 
with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)"  | 
113  | 
by simp  | 
|
| 63498 | 114  | 
from A show "content p = 1"  | 
115  | 
by (auto simp: is_unit_poly_iff normalize_1_iff)  | 
|
116  | 
next  | 
|
117  | 
assume A: "fract_poly p dvd 1" and B: "content p = 1"  | 
|
118  | 
from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff)  | 
|
119  | 
  {
 | 
|
120  | 
fix n :: nat assume "n > 0"  | 
|
121  | 
have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly)  | 
|
122  | 
also note c  | 
|
123  | 
also from \<open>n > 0\<close> have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits)  | 
|
124  | 
finally have "coeff p n = 0" by simp  | 
|
125  | 
}  | 
|
126  | 
hence "degree p \<le> 0" by (intro degree_le) simp_all  | 
|
127  | 
with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE)  | 
|
128  | 
qed  | 
|
129  | 
||
130  | 
lemma fract_poly_is_unit: "p dvd 1 \<Longrightarrow> fract_poly p dvd 1"  | 
|
131  | 
using fract_poly_dvd[of p 1] by simp  | 
|
132  | 
||
133  | 
lemma fract_poly_smult_eqE:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
134  | 
  fixes c :: "'a :: {idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract"
 | 
| 63498 | 135  | 
assumes "fract_poly p = smult c (fract_poly q)"  | 
136  | 
obtains a b  | 
|
137  | 
where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a"  | 
|
138  | 
proof -  | 
|
139  | 
define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)"  | 
|
140  | 
have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)"  | 
|
141  | 
by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms)  | 
|
142  | 
hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff)  | 
|
143  | 
hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff)  | 
|
144  | 
moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b"  | 
|
| 67051 | 145  | 
by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps  | 
| 63498 | 146  | 
normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric])  | 
147  | 
ultimately show ?thesis by (intro that[of a b])  | 
|
148  | 
qed  | 
|
149  | 
||
150  | 
||
151  | 
subsection \<open>Fractional content\<close>  | 
|
152  | 
||
153  | 
abbreviation (input) Lcm_coeff_denoms  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
154  | 
    :: "'a :: {semiring_Gcd,idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract poly \<Rightarrow> 'a"
 | 
| 63498 | 155  | 
where "Lcm_coeff_denoms p \<equiv> Lcm (snd ` quot_of_fract ` set (coeffs p))"  | 
156  | 
||
157  | 
definition fract_content ::  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
158  | 
      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \<Rightarrow> 'a fract" where
 | 
| 63498 | 159  | 
"fract_content p =  | 
160  | 
(let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)"  | 
|
161  | 
||
162  | 
definition primitive_part_fract ::  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
163  | 
      "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \<Rightarrow> 'a poly" where
 | 
| 63498 | 164  | 
"primitive_part_fract p =  | 
165  | 
primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))"  | 
|
166  | 
||
167  | 
lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0"  | 
|
168  | 
by (simp add: primitive_part_fract_def)  | 
|
169  | 
||
170  | 
lemma fract_content_eq_0_iff [simp]:  | 
|
171  | 
"fract_content p = 0 \<longleftrightarrow> p = 0"  | 
|
172  | 
unfolding fract_content_def Let_def Zero_fract_def  | 
|
173  | 
by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff)  | 
|
174  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
175  | 
lemma content_primitive_part_fract [simp]:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
176  | 
  fixes p :: "'a :: {semiring_gcd_mult_normalize,
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
177  | 
factorial_semiring, ring_gcd, semiring_Gcd,idom_divide} fract poly"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
178  | 
shows "p \<noteq> 0 \<Longrightarrow> content (primitive_part_fract p) = 1"  | 
| 63498 | 179  | 
unfolding primitive_part_fract_def  | 
180  | 
by (rule content_primitive_part)  | 
|
181  | 
(auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff)  | 
|
182  | 
||
183  | 
lemma content_times_primitive_part_fract:  | 
|
184  | 
"smult (fract_content p) (fract_poly (primitive_part_fract p)) = p"  | 
|
185  | 
proof -  | 
|
186  | 
define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)"  | 
|
187  | 
have "fract_poly p' =  | 
|
188  | 
map_poly (to_fract \<circ> fst \<circ> quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)"  | 
|
189  | 
unfolding primitive_part_fract_def p'_def  | 
|
190  | 
by (subst map_poly_map_poly) (simp_all add: o_assoc)  | 
|
191  | 
also have "\<dots> = smult (to_fract (Lcm_coeff_denoms p)) p"  | 
|
192  | 
proof (intro map_poly_idI, unfold o_apply)  | 
|
193  | 
fix c assume "c \<in> set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))"  | 
|
194  | 
then obtain c' where c: "c' \<in> set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'"  | 
|
195  | 
by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits)  | 
|
196  | 
note c(2)  | 
|
197  | 
also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))"  | 
|
198  | 
by simp  | 
|
199  | 
also have "to_fract (Lcm_coeff_denoms p) * \<dots> =  | 
|
200  | 
Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))"  | 
|
201  | 
unfolding to_fract_def by (subst mult_fract) simp_all  | 
|
202  | 
also have "snd (quot_of_fract \<dots>) = 1"  | 
|
203  | 
by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto)  | 
|
204  | 
finally show "to_fract (fst (quot_of_fract c)) = c"  | 
|
205  | 
by (rule to_fract_quot_of_fract)  | 
|
206  | 
qed  | 
|
207  | 
also have "p' = smult (content p') (primitive_part p')"  | 
|
208  | 
by (rule content_times_primitive_part [symmetric])  | 
|
209  | 
also have "primitive_part p' = primitive_part_fract p"  | 
|
210  | 
by (simp add: primitive_part_fract_def p'_def)  | 
|
211  | 
also have "fract_poly (smult (content p') (primitive_part_fract p)) =  | 
|
212  | 
smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp  | 
|
213  | 
finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) =  | 
|
214  | 
smult (to_fract (Lcm_coeff_denoms p)) p" .  | 
|
215  | 
thus ?thesis  | 
|
216  | 
by (subst (asm) smult_eq_iff)  | 
|
217  | 
(auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def)  | 
|
218  | 
qed  | 
|
219  | 
||
220  | 
lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)"  | 
|
221  | 
proof -  | 
|
222  | 
have "Lcm_coeff_denoms (fract_poly p) = 1"  | 
|
| 63905 | 223  | 
by (auto simp: set_coeffs_map_poly)  | 
| 63498 | 224  | 
hence "fract_content (fract_poly p) =  | 
225  | 
to_fract (content (map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p))"  | 
|
226  | 
by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff)  | 
|
227  | 
also have "map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p = p"  | 
|
228  | 
by (intro map_poly_idI) simp_all  | 
|
229  | 
finally show ?thesis .  | 
|
230  | 
qed  | 
|
231  | 
||
232  | 
lemma content_decompose_fract:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
233  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
234  | 
semiring_gcd_mult_normalize} fract poly"  | 
| 63498 | 235  | 
obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1"  | 
236  | 
proof (cases "p = 0")  | 
|
237  | 
case True  | 
|
238  | 
hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all  | 
|
239  | 
thus ?thesis ..  | 
|
240  | 
next  | 
|
241  | 
case False  | 
|
242  | 
thus ?thesis  | 
|
243  | 
by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract])  | 
|
244  | 
qed  | 
|
245  | 
||
246  | 
lemma fract_poly_dvdD:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
247  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
248  | 
semiring_gcd_mult_normalize} poly"  | 
| 63498 | 249  | 
assumes "fract_poly p dvd fract_poly q" "content p = 1"  | 
250  | 
shows "p dvd q"  | 
|
251  | 
proof -  | 
|
252  | 
from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE)  | 
|
| 74362 | 253  | 
from content_decompose_fract[of r]  | 
254  | 
obtain c r' where r': "r = smult c (map_poly to_fract r')" "content r' = 1" .  | 
|
| 63498 | 255  | 
from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp  | 
| 74362 | 256  | 
from fract_poly_smult_eqE[OF this] obtain a b  | 
257  | 
where ab:  | 
|
258  | 
"c = to_fract b / to_fract a"  | 
|
259  | 
"smult a q = smult b (p * r')"  | 
|
260  | 
"coprime a b"  | 
|
261  | 
"normalize a = a" .  | 
|
| 63498 | 262  | 
have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2))  | 
263  | 
hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4))  | 
|
264  | 
have "1 = gcd a (normalize b)" by (simp add: ab)  | 
|
265  | 
also note eq'  | 
|
266  | 
also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4))  | 
|
267  | 
finally have [simp]: "a = 1" by simp  | 
|
268  | 
from eq ab have "q = p * ([:b:] * r')" by simp  | 
|
269  | 
thus ?thesis by (rule dvdI)  | 
|
270  | 
qed  | 
|
271  | 
||
272  | 
||
273  | 
subsection \<open>Polynomials over a field are a Euclidean ring\<close>  | 
|
274  | 
||
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
275  | 
context  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
276  | 
begin  | 
| 63498 | 277  | 
|
278  | 
interpretation field_poly:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
279  | 
normalization_euclidean_semiring_multiplicative where zero = "0 :: 'a :: field poly"  | 
| 66817 | 280  | 
and one = 1 and plus = plus and minus = minus  | 
| 
64164
 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 
haftmann 
parents: 
63954 
diff
changeset
 | 
281  | 
and times = times  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
282  | 
and normalize = "\<lambda>p. smult (inverse (lead_coeff p)) p"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
283  | 
and unit_factor = "\<lambda>p. [:lead_coeff p:]"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
284  | 
and euclidean_size = "\<lambda>p. if p = 0 then 0 else 2 ^ degree p"  | 
| 
64164
 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 
haftmann 
parents: 
63954 
diff
changeset
 | 
285  | 
and divide = divide and modulo = modulo  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
286  | 
rewrites "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
287  | 
and "comm_monoid_mult.prod_mset times 1 = prod_mset"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
288  | 
and "comm_semiring_1.irreducible times 1 0 = irreducible"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
289  | 
and "comm_semiring_1.prime_elem times 1 0 = prime_elem"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
290  | 
proof -  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
291  | 
show "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
292  | 
by (simp add: dvd_dict)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
293  | 
show "comm_monoid_mult.prod_mset times 1 = prod_mset"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
294  | 
by (simp add: prod_mset_dict)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
295  | 
show "comm_semiring_1.irreducible times 1 0 = irreducible"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
296  | 
by (simp add: irreducible_dict)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
297  | 
show "comm_semiring_1.prime_elem times 1 0 = prime_elem"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
298  | 
by (simp add: prime_elem_dict)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
299  | 
show "class.normalization_euclidean_semiring_multiplicative divide plus minus (0 :: 'a poly) times 1  | 
| 66817 | 300  | 
modulo (\<lambda>p. if p = 0 then 0 else 2 ^ degree p)  | 
301  | 
(\<lambda>p. [:lead_coeff p:]) (\<lambda>p. smult (inverse (lead_coeff p)) p)"  | 
|
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
302  | 
proof (standard, fold dvd_dict)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
303  | 
fix p :: "'a poly"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
304  | 
show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
305  | 
by (cases "p = 0") simp_all  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
306  | 
next  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
307  | 
fix p :: "'a poly" assume "is_unit p"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
308  | 
then show "[:lead_coeff p:] = p"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
309  | 
by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
310  | 
next  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
311  | 
fix p :: "'a poly" assume "p \<noteq> 0"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
312  | 
then show "is_unit [:lead_coeff p:]"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
313  | 
by (simp add: is_unit_pCons_iff)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
314  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
315  | 
fix a b :: "'a poly" assume "is_unit a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
316  | 
thus "[:lead_coeff (a * b):] = a * [:lead_coeff b:]"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
317  | 
by (auto elim!: is_unit_polyE)  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
318  | 
qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
319  | 
qed  | 
| 63498 | 320  | 
|
| 
63722
 
b9c8da46443b
Deprivatisation of lemmas in Polynomial_Factorial
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63705 
diff
changeset
 | 
321  | 
lemma field_poly_irreducible_imp_prime:  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
322  | 
"prime_elem p" if "irreducible p" for p :: "'a :: field poly"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
323  | 
using that by (fact field_poly.irreducible_imp_prime_elem)  | 
| 72265 | 324  | 
|
| 63830 | 325  | 
lemma field_poly_prod_mset_prime_factorization:  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
326  | 
"prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
327  | 
if "p \<noteq> 0" for p :: "'a :: field poly"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
328  | 
using that by (fact field_poly.prod_mset_prime_factorization)  | 
| 63498 | 329  | 
|
| 
63722
 
b9c8da46443b
Deprivatisation of lemmas in Polynomial_Factorial
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63705 
diff
changeset
 | 
330  | 
lemma field_poly_in_prime_factorization_imp_prime:  | 
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
331  | 
"prime_elem p" if "p \<in># field_poly.prime_factorization x"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
332  | 
for p :: "'a :: field poly"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
333  | 
by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
334  | 
(fact that)  | 
| 63498 | 335  | 
|
336  | 
||
337  | 
subsection \<open>Primality and irreducibility in polynomial rings\<close>  | 
|
338  | 
||
339  | 
lemma nonconst_poly_irreducible_iff:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
340  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 341  | 
assumes "degree p \<noteq> 0"  | 
342  | 
shows "irreducible p \<longleftrightarrow> irreducible (fract_poly p) \<and> content p = 1"  | 
|
343  | 
proof safe  | 
|
344  | 
assume p: "irreducible p"  | 
|
345  | 
||
| 74362 | 346  | 
from content_decompose[of p] obtain p' where p': "p = smult (content p) p'" "content p' = 1" .  | 
| 63498 | 347  | 
hence "p = [:content p:] * p'" by simp  | 
348  | 
from p this have "[:content p:] dvd 1 \<or> p' dvd 1" by (rule irreducibleD)  | 
|
349  | 
moreover have "\<not>p' dvd 1"  | 
|
350  | 
proof  | 
|
351  | 
assume "p' dvd 1"  | 
|
352  | 
hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff)  | 
|
353  | 
with assms show False by contradiction  | 
|
354  | 
qed  | 
|
355  | 
ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff)  | 
|
356  | 
||
357  | 
show "irreducible (map_poly to_fract p)"  | 
|
358  | 
proof (rule irreducibleI)  | 
|
359  | 
have "fract_poly p = 0 \<longleftrightarrow> p = 0" by (intro map_poly_eq_0_iff) auto  | 
|
360  | 
with assms show "map_poly to_fract p \<noteq> 0" by auto  | 
|
361  | 
next  | 
|
362  | 
show "\<not>is_unit (fract_poly p)"  | 
|
363  | 
proof  | 
|
364  | 
assume "is_unit (map_poly to_fract p)"  | 
|
365  | 
hence "degree (map_poly to_fract p) = 0"  | 
|
366  | 
by (auto simp: is_unit_poly_iff)  | 
|
367  | 
hence "degree p = 0" by (simp add: degree_map_poly)  | 
|
368  | 
with assms show False by contradiction  | 
|
369  | 
qed  | 
|
370  | 
next  | 
|
371  | 
fix q r assume qr: "fract_poly p = q * r"  | 
|
| 74362 | 372  | 
from content_decompose_fract[of q]  | 
373  | 
obtain cg q' where q: "q = smult cg (map_poly to_fract q')" "content q' = 1" .  | 
|
374  | 
from content_decompose_fract[of r]  | 
|
375  | 
obtain cr r' where r: "r = smult cr (map_poly to_fract r')" "content r' = 1" .  | 
|
| 63498 | 376  | 
from qr q r p have nz: "cg \<noteq> 0" "cr \<noteq> 0" by auto  | 
377  | 
from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))"  | 
|
378  | 
by (simp add: q r)  | 
|
| 74362 | 379  | 
from fract_poly_smult_eqE[OF this] obtain a b  | 
380  | 
where ab: "cr * cg = to_fract b / to_fract a"  | 
|
381  | 
"smult a p = smult b (q' * r')" "coprime a b" "normalize a = a" .  | 
|
| 63498 | 382  | 
hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:)  | 
383  | 
with ab(4) have a: "a = normalize b" by (simp add: content_mult q r)  | 
|
| 67051 | 384  | 
then have "normalize b = gcd a b"  | 
385  | 
by simp  | 
|
386  | 
with \<open>coprime a b\<close> have "normalize b = 1"  | 
|
387  | 
by simp  | 
|
388  | 
then have "a = 1" "is_unit b"  | 
|
389  | 
by (simp_all add: a normalize_1_iff)  | 
|
| 63498 | 390  | 
|
391  | 
note eq  | 
|
392  | 
also from ab(1) \<open>a = 1\<close> have "cr * cg = to_fract b" by simp  | 
|
393  | 
also have "smult \<dots> (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp  | 
|
394  | 
finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult)  | 
|
395  | 
from p and this have "([:b:] * q') dvd 1 \<or> r' dvd 1" by (rule irreducibleD)  | 
|
396  | 
hence "q' dvd 1 \<or> r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left)  | 
|
397  | 
hence "fract_poly q' dvd 1 \<or> fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit)  | 
|
398  | 
with q r show "is_unit q \<or> is_unit r"  | 
|
399  | 
by (auto simp add: is_unit_smult_iff dvd_field_iff nz)  | 
|
400  | 
qed  | 
|
401  | 
||
402  | 
next  | 
|
403  | 
||
404  | 
assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1"  | 
|
405  | 
show "irreducible p"  | 
|
406  | 
proof (rule irreducibleI)  | 
|
407  | 
from irred show "p \<noteq> 0" by auto  | 
|
408  | 
next  | 
|
409  | 
from irred show "\<not>p dvd 1"  | 
|
410  | 
by (auto simp: irreducible_def dest: fract_poly_is_unit)  | 
|
411  | 
next  | 
|
412  | 
fix q r assume qr: "p = q * r"  | 
|
413  | 
hence "fract_poly p = fract_poly q * fract_poly r" by simp  | 
|
414  | 
from irred and this have "fract_poly q dvd 1 \<or> fract_poly r dvd 1"  | 
|
415  | 
by (rule irreducibleD)  | 
|
416  | 
with primitive qr show "q dvd 1 \<or> r dvd 1"  | 
|
417  | 
by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff)  | 
|
418  | 
qed  | 
|
419  | 
qed  | 
|
420  | 
||
| 
74542
 
d592354c4a26
removed some 'private' modifiers from HOL-Computational_Algebra
 
Manuel Eberl <manuel@pruvisto.org> 
parents: 
74362 
diff
changeset
 | 
421  | 
lemma irreducible_imp_prime_poly:  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
422  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 423  | 
assumes "irreducible p"  | 
| 63633 | 424  | 
shows "prime_elem p"  | 
| 63498 | 425  | 
proof (cases "degree p = 0")  | 
426  | 
case True  | 
|
427  | 
with assms show ?thesis  | 
|
428  | 
by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff  | 
|
| 63633 | 429  | 
intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE)  | 
| 63498 | 430  | 
next  | 
431  | 
case False  | 
|
432  | 
from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1"  | 
|
433  | 
by (simp_all add: nonconst_poly_irreducible_iff)  | 
|
| 63633 | 434  | 
from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime)  | 
| 63498 | 435  | 
show ?thesis  | 
| 63633 | 436  | 
proof (rule prime_elemI)  | 
| 63498 | 437  | 
fix q r assume "p dvd q * r"  | 
438  | 
hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd)  | 
|
439  | 
hence "fract_poly p dvd fract_poly q * fract_poly r" by simp  | 
|
440  | 
from prime and this have "fract_poly p dvd fract_poly q \<or> fract_poly p dvd fract_poly r"  | 
|
| 63633 | 441  | 
by (rule prime_elem_dvd_multD)  | 
| 63498 | 442  | 
with primitive show "p dvd q \<or> p dvd r" by (auto dest: fract_poly_dvdD)  | 
443  | 
qed (insert assms, auto simp: irreducible_def)  | 
|
444  | 
qed  | 
|
445  | 
||
446  | 
lemma degree_primitive_part_fract [simp]:  | 
|
447  | 
"degree (primitive_part_fract p) = degree p"  | 
|
448  | 
proof -  | 
|
449  | 
have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))"  | 
|
450  | 
by (simp add: content_times_primitive_part_fract)  | 
|
451  | 
also have "degree \<dots> = degree (primitive_part_fract p)"  | 
|
452  | 
by (auto simp: degree_map_poly)  | 
|
453  | 
finally show ?thesis ..  | 
|
454  | 
qed  | 
|
455  | 
||
456  | 
lemma irreducible_primitive_part_fract:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
457  | 
  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly"
 | 
| 63498 | 458  | 
assumes "irreducible p"  | 
459  | 
shows "irreducible (primitive_part_fract p)"  | 
|
460  | 
proof -  | 
|
461  | 
from assms have deg: "degree (primitive_part_fract p) \<noteq> 0"  | 
|
462  | 
by (intro notI)  | 
|
463  | 
(auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff)  | 
|
464  | 
hence [simp]: "p \<noteq> 0" by auto  | 
|
465  | 
||
466  | 
note \<open>irreducible p\<close>  | 
|
467  | 
also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)"  | 
|
468  | 
by (simp add: content_times_primitive_part_fract)  | 
|
469  | 
also have "irreducible \<dots> \<longleftrightarrow> irreducible (fract_poly (primitive_part_fract p))"  | 
|
470  | 
by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff)  | 
|
471  | 
finally show ?thesis using deg  | 
|
472  | 
by (simp add: nonconst_poly_irreducible_iff)  | 
|
473  | 
qed  | 
|
474  | 
||
| 63633 | 475  | 
lemma prime_elem_primitive_part_fract:  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
476  | 
  fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly"
 | 
| 63633 | 477  | 
shows "irreducible p \<Longrightarrow> prime_elem (primitive_part_fract p)"  | 
| 63498 | 478  | 
by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract)  | 
479  | 
||
480  | 
lemma irreducible_linear_field_poly:  | 
|
481  | 
fixes a b :: "'a::field"  | 
|
482  | 
assumes "b \<noteq> 0"  | 
|
483  | 
shows "irreducible [:a,b:]"  | 
|
484  | 
proof (rule irreducibleI)  | 
|
485  | 
fix p q assume pq: "[:a,b:] = p * q"  | 
|
| 63539 | 486  | 
also from pq assms have "degree \<dots> = degree p + degree q"  | 
| 63498 | 487  | 
by (intro degree_mult_eq) auto  | 
488  | 
finally have "degree p = 0 \<or> degree q = 0" using assms by auto  | 
|
489  | 
with assms pq show "is_unit p \<or> is_unit q"  | 
|
490  | 
by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE)  | 
|
491  | 
qed (insert assms, auto simp: is_unit_poly_iff)  | 
|
492  | 
||
| 63633 | 493  | 
lemma prime_elem_linear_field_poly:  | 
494  | 
"(b :: 'a :: field) \<noteq> 0 \<Longrightarrow> prime_elem [:a,b:]"  | 
|
| 63498 | 495  | 
by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly)  | 
496  | 
||
497  | 
lemma irreducible_linear_poly:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
498  | 
  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}"
 | 
| 63498 | 499  | 
shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> irreducible [:a,b:]"  | 
500  | 
by (auto intro!: irreducible_linear_field_poly  | 
|
501  | 
simp: nonconst_poly_irreducible_iff content_def map_poly_pCons)  | 
|
502  | 
||
| 63633 | 503  | 
lemma prime_elem_linear_poly:  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
504  | 
  fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}"
 | 
| 63633 | 505  | 
shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> prime_elem [:a,b:]"  | 
| 63498 | 506  | 
by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly)  | 
507  | 
||
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
508  | 
|
| 63498 | 509  | 
subsection \<open>Prime factorisation of polynomials\<close>  | 
510  | 
||
| 
74542
 
d592354c4a26
removed some 'private' modifiers from HOL-Computational_Algebra
 
Manuel Eberl <manuel@pruvisto.org> 
parents: 
74362 
diff
changeset
 | 
511  | 
lemma poly_prime_factorization_exists_content_1:  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
512  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 513  | 
assumes "p \<noteq> 0" "content p = 1"  | 
| 63830 | 514  | 
shows "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p"  | 
| 63498 | 515  | 
proof -  | 
516  | 
let ?P = "field_poly.prime_factorization (fract_poly p)"  | 
|
| 63830 | 517  | 
define c where "c = prod_mset (image_mset fract_content ?P)"  | 
| 63498 | 518  | 
define c' where "c' = c * to_fract (lead_coeff p)"  | 
| 63830 | 519  | 
define e where "e = prod_mset (image_mset primitive_part_fract ?P)"  | 
| 63498 | 520  | 
define A where "A = image_mset (normalize \<circ> primitive_part_fract) ?P"  | 
521  | 
have "content e = (\<Prod>x\<in>#field_poly.prime_factorization (map_poly to_fract p).  | 
|
522  | 
content (primitive_part_fract x))"  | 
|
| 63830 | 523  | 
by (simp add: e_def content_prod_mset multiset.map_comp o_def)  | 
| 63498 | 524  | 
also have "image_mset (\<lambda>x. content (primitive_part_fract x)) ?P = image_mset (\<lambda>_. 1) ?P"  | 
525  | 
by (intro image_mset_cong content_primitive_part_fract) auto  | 
|
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
526  | 
finally have content_e: "content e = 1"  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
527  | 
by simp  | 
| 63498 | 528  | 
|
| 
66805
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
529  | 
from \<open>p \<noteq> 0\<close> have "fract_poly p = [:lead_coeff (fract_poly p):] *  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
530  | 
smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
531  | 
by simp  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
532  | 
also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]"  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
533  | 
by (simp add: monom_0 degree_map_poly coeff_map_poly)  | 
| 
 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 
haftmann 
parents: 
65965 
diff
changeset
 | 
534  | 
also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P"  | 
| 63830 | 535  | 
by (subst field_poly_prod_mset_prime_factorization) simp_all  | 
536  | 
also have "\<dots> = prod_mset (image_mset id ?P)" by simp  | 
|
| 63498 | 537  | 
also have "image_mset id ?P =  | 
538  | 
image_mset (\<lambda>x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P"  | 
|
539  | 
by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract)  | 
|
| 63830 | 540  | 
also have "prod_mset \<dots> = smult c (fract_poly e)"  | 
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
541  | 
by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def)  | 
| 63498 | 542  | 
also have "[:to_fract (lead_coeff p):] * \<dots> = smult c' (fract_poly e)"  | 
543  | 
by (simp add: c'_def)  | 
|
544  | 
finally have eq: "fract_poly p = smult c' (fract_poly e)" .  | 
|
545  | 
also obtain b where b: "c' = to_fract b" "is_unit b"  | 
|
546  | 
proof -  | 
|
| 74362 | 547  | 
from fract_poly_smult_eqE[OF eq]  | 
548  | 
obtain a b where ab:  | 
|
549  | 
"c' = to_fract b / to_fract a"  | 
|
550  | 
"smult a p = smult b e"  | 
|
551  | 
"coprime a b"  | 
|
552  | 
"normalize a = a" .  | 
|
| 63498 | 553  | 
from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: )  | 
554  | 
with assms content_e have "a = normalize b" by (simp add: ab(4))  | 
|
| 67051 | 555  | 
with ab have ab': "a = 1" "is_unit b"  | 
556  | 
by (simp_all add: normalize_1_iff)  | 
|
| 63498 | 557  | 
with ab ab' have "c' = to_fract b" by auto  | 
558  | 
from this and \<open>is_unit b\<close> show ?thesis by (rule that)  | 
|
559  | 
qed  | 
|
560  | 
hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp  | 
|
561  | 
finally have "p = smult b e" by (simp only: fract_poly_eq_iff)  | 
|
562  | 
hence "p = [:b:] * e" by simp  | 
|
563  | 
with b have "normalize p = normalize e"  | 
|
564  | 
by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff)  | 
|
| 63830 | 565  | 
also have "normalize e = prod_mset A"  | 
566  | 
by (simp add: multiset.map_comp e_def A_def normalize_prod_mset)  | 
|
567  | 
finally have "prod_mset A = normalize p" ..  | 
|
| 63498 | 568  | 
|
| 63633 | 569  | 
have "prime_elem p" if "p \<in># A" for p  | 
570  | 
using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible  | 
|
| 63498 | 571  | 
dest!: field_poly_in_prime_factorization_imp_prime )  | 
| 63830 | 572  | 
from this and \<open>prod_mset A = normalize p\<close> show ?thesis  | 
| 63498 | 573  | 
by (intro exI[of _ A]) blast  | 
574  | 
qed  | 
|
575  | 
||
576  | 
lemma poly_prime_factorization_exists:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
577  | 
  fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 578  | 
assumes "p \<noteq> 0"  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
579  | 
shows "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> normalize (prod_mset A) = normalize p"  | 
| 63498 | 580  | 
proof -  | 
581  | 
define B where "B = image_mset (\<lambda>x. [:x:]) (prime_factorization (content p))"  | 
|
| 63830 | 582  | 
have "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize (primitive_part p)"  | 
| 63498 | 583  | 
by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all)  | 
| 74362 | 584  | 
then obtain A where A: "\<forall>p. p \<in># A \<longrightarrow> prime_elem p" "\<Prod>\<^sub># A = normalize (primitive_part p)"  | 
585  | 
by blast  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
586  | 
have "normalize (prod_mset (A + B)) = normalize (prod_mset A * normalize (prod_mset B))"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
587  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
588  | 
also from assms have "normalize (prod_mset B) = normalize [:content p:]"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
589  | 
by (simp add: prod_mset_const_poly normalize_const_poly prod_mset_prime_factorization_weak B_def)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
590  | 
also have "prod_mset A = normalize (primitive_part p)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
591  | 
using A by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
592  | 
finally have "normalize (prod_mset (A + B)) = normalize (primitive_part p * [:content p:])"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
593  | 
by simp  | 
| 63633 | 594  | 
moreover have "\<forall>p. p \<in># B \<longrightarrow> prime_elem p"  | 
| 63905 | 595  | 
by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
596  | 
ultimately show ?thesis using A by (intro exI[of _ "A + B"]) (auto)  | 
| 63498 | 597  | 
qed  | 
598  | 
||
599  | 
end  | 
|
600  | 
||
601  | 
||
602  | 
subsection \<open>Typeclass instances\<close>  | 
|
603  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
604  | 
instance poly :: ("{factorial_ring_gcd,semiring_gcd_mult_normalize}") factorial_semiring
 | 
| 63498 | 605  | 
by standard (rule poly_prime_factorization_exists)  | 
606  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
607  | 
instantiation poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") factorial_ring_gcd
 | 
| 63498 | 608  | 
begin  | 
609  | 
||
610  | 
definition gcd_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where  | 
|
611  | 
[code del]: "gcd_poly = gcd_factorial"  | 
|
612  | 
||
613  | 
definition lcm_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where  | 
|
614  | 
[code del]: "lcm_poly = lcm_factorial"  | 
|
615  | 
||
616  | 
definition Gcd_poly :: "'a poly set \<Rightarrow> 'a poly" where  | 
|
617  | 
[code del]: "Gcd_poly = Gcd_factorial"  | 
|
618  | 
||
619  | 
definition Lcm_poly :: "'a poly set \<Rightarrow> 'a poly" where  | 
|
620  | 
[code del]: "Lcm_poly = Lcm_factorial"  | 
|
621  | 
||
622  | 
instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def)  | 
|
623  | 
||
624  | 
end  | 
|
625  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
626  | 
instance poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") semiring_gcd_mult_normalize ..
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
627  | 
|
| 76121 | 628  | 
instance poly :: ("{field,factorial_ring_gcd,semiring_gcd_mult_normalize}")
 | 
629  | 
"normalization_euclidean_semiring" ..  | 
|
| 63498 | 630  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
631  | 
instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd,
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
632  | 
semiring_gcd_mult_normalize}") euclidean_ring_gcd  | 
| 66817 | 633  | 
by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard  | 
| 63498 | 634  | 
|
635  | 
||
636  | 
subsection \<open>Polynomial GCD\<close>  | 
|
637  | 
||
638  | 
lemma gcd_poly_decompose:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
639  | 
  fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 640  | 
shows "gcd p q =  | 
641  | 
smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"  | 
|
642  | 
proof (rule sym, rule gcdI)  | 
|
643  | 
have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd  | 
|
644  | 
[:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all  | 
|
645  | 
thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p"  | 
|
646  | 
by simp  | 
|
647  | 
next  | 
|
648  | 
have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd  | 
|
649  | 
[:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all  | 
|
650  | 
thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q"  | 
|
651  | 
by simp  | 
|
652  | 
next  | 
|
653  | 
fix d assume "d dvd p" "d dvd q"  | 
|
654  | 
hence "[:content d:] * primitive_part d dvd  | 
|
655  | 
[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)"  | 
|
656  | 
by (intro mult_dvd_mono) auto  | 
|
657  | 
thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))"  | 
|
658  | 
by simp  | 
|
659  | 
qed (auto simp: normalize_smult)  | 
|
660  | 
||
661  | 
||
662  | 
lemma gcd_poly_pseudo_mod:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
663  | 
  fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"
 | 
| 63498 | 664  | 
assumes nz: "q \<noteq> 0" and prim: "content p = 1" "content q = 1"  | 
665  | 
shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))"  | 
|
666  | 
proof -  | 
|
667  | 
define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)"  | 
|
668  | 
define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]"  | 
|
669  | 
have [simp]: "primitive_part a = unit_factor a"  | 
|
670  | 
by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0)  | 
|
671  | 
from nz have [simp]: "a \<noteq> 0" by (auto simp: a_def)  | 
|
672  | 
||
673  | 
have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def)  | 
|
674  | 
have "gcd (q * r + s) q = gcd q s"  | 
|
675  | 
using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac)  | 
|
676  | 
with pseudo_divmod(1)[OF nz rs]  | 
|
677  | 
have "gcd (p * a) q = gcd q s" by (simp add: a_def)  | 
|
678  | 
also from prim have "gcd (p * a) q = gcd p q"  | 
|
679  | 
by (subst gcd_poly_decompose)  | 
|
680  | 
(auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim  | 
|
681  | 
simp del: mult_pCons_right )  | 
|
682  | 
also from prim have "gcd q s = gcd q (primitive_part s)"  | 
|
683  | 
by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim)  | 
|
684  | 
also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def)  | 
|
685  | 
finally show ?thesis .  | 
|
686  | 
qed  | 
|
687  | 
||
688  | 
lemma degree_pseudo_mod_less:  | 
|
689  | 
assumes "q \<noteq> 0" "pseudo_mod p q \<noteq> 0"  | 
|
690  | 
shows "degree (pseudo_mod p q) < degree q"  | 
|
691  | 
using pseudo_mod(2)[of q p] assms by auto  | 
|
692  | 
||
693  | 
function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where  | 
|
694  | 
"gcd_poly_code_aux p q =  | 
|
695  | 
(if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))"  | 
|
696  | 
by auto  | 
|
697  | 
termination  | 
|
698  | 
by (relation "measure ((\<lambda>p. if p = 0 then 0 else Suc (degree p)) \<circ> snd)")  | 
|
| 
64164
 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 
haftmann 
parents: 
63954 
diff
changeset
 | 
699  | 
(auto simp: degree_pseudo_mod_less)  | 
| 63498 | 700  | 
|
701  | 
declare gcd_poly_code_aux.simps [simp del]  | 
|
702  | 
||
703  | 
lemma gcd_poly_code_aux_correct:  | 
|
704  | 
assumes "content p = 1" "q = 0 \<or> content q = 1"  | 
|
705  | 
shows "gcd_poly_code_aux p q = gcd p q"  | 
|
706  | 
using assms  | 
|
707  | 
proof (induction p q rule: gcd_poly_code_aux.induct)  | 
|
708  | 
case (1 p q)  | 
|
709  | 
show ?case  | 
|
710  | 
proof (cases "q = 0")  | 
|
711  | 
case True  | 
|
712  | 
thus ?thesis by (subst gcd_poly_code_aux.simps) auto  | 
|
713  | 
next  | 
|
714  | 
case False  | 
|
715  | 
hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))"  | 
|
716  | 
by (subst gcd_poly_code_aux.simps) simp_all  | 
|
717  | 
also from "1.prems" False  | 
|
718  | 
have "primitive_part (pseudo_mod p q) = 0 \<or>  | 
|
719  | 
content (primitive_part (pseudo_mod p q)) = 1"  | 
|
720  | 
by (cases "pseudo_mod p q = 0") auto  | 
|
721  | 
with "1.prems" False  | 
|
722  | 
have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) =  | 
|
723  | 
gcd q (primitive_part (pseudo_mod p q))"  | 
|
724  | 
by (intro 1) simp_all  | 
|
725  | 
also from "1.prems" False  | 
|
726  | 
have "\<dots> = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto  | 
|
727  | 
finally show ?thesis .  | 
|
728  | 
qed  | 
|
729  | 
qed  | 
|
730  | 
||
731  | 
definition gcd_poly_code  | 
|
732  | 
:: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"  | 
|
733  | 
where "gcd_poly_code p q =  | 
|
734  | 
(if p = 0 then normalize q else if q = 0 then normalize p else  | 
|
735  | 
smult (gcd (content p) (content q))  | 
|
736  | 
(gcd_poly_code_aux (primitive_part p) (primitive_part q)))"  | 
|
737  | 
||
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
738  | 
lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q"  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
739  | 
by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric])  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
740  | 
|
| 63498 | 741  | 
lemma lcm_poly_code [code]:  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
742  | 
  fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
743  | 
shows "lcm p q = normalize (p * q div gcd p q)"  | 
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
744  | 
by (fact lcm_gcd)  | 
| 63498 | 745  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
746  | 
lemmas Gcd_poly_set_eq_fold [code] =  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
747  | 
  Gcd_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"]
 | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
748  | 
lemmas Lcm_poly_set_eq_fold [code] =  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68790 
diff
changeset
 | 
749  | 
  Lcm_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"]
 | 
| 64860 | 750  | 
|
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
751  | 
text \<open>Example:  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
752  | 
  @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval}
 | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64267 
diff
changeset
 | 
753  | 
\<close>  | 
| 63498 | 754  | 
|
| 63764 | 755  | 
end  |