| author | blanchet | 
| Thu, 06 Dec 2012 16:07:09 +0100 | |
| changeset 50402 | 923f1e199f4f | 
| parent 49834 | b27bbb021df1 | 
| child 51489 | f738e6dbd844 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Algebra/poly/UnivPoly2.thy  | 
| 35849 | 2  | 
Author: Clemens Ballarin, started 9 December 1996  | 
3  | 
Copyright: Clemens Ballarin  | 
|
| 13936 | 4  | 
*)  | 
5  | 
||
6  | 
header {* Univariate Polynomials *}
 | 
|
7  | 
||
| 15481 | 8  | 
theory UnivPoly2  | 
9  | 
imports "../abstract/Abstract"  | 
|
10  | 
begin  | 
|
| 13936 | 11  | 
|
| 
16640
 
03bdf544a552
Removed setsubgoaler hack (thanks to strong_setsum_cong).
 
berghofe 
parents: 
16052 
diff
changeset
 | 
12  | 
(* With this variant of setsum_cong, assumptions  | 
| 13936 | 13  | 
   like i:{m..n} get simplified (to m <= i & i <= n). *)
 | 
14  | 
||
| 
16640
 
03bdf544a552
Removed setsubgoaler hack (thanks to strong_setsum_cong).
 
berghofe 
parents: 
16052 
diff
changeset
 | 
15  | 
declare strong_setsum_cong [cong]  | 
| 13936 | 16  | 
|
| 35849 | 17  | 
|
| 13936 | 18  | 
section {* Definition of type up *}
 | 
19  | 
||
| 21423 | 20  | 
definition  | 
21  | 
bound :: "[nat, nat => 'a::zero] => bool" where  | 
|
22  | 
"bound n f = (ALL i. n < i --> f i = 0)"  | 
|
| 13936 | 23  | 
|
24  | 
lemma boundI [intro!]: "[| !! m. n < m ==> f m = 0 |] ==> bound n f"  | 
|
| 21423 | 25  | 
unfolding bound_def by blast  | 
| 13936 | 26  | 
|
27  | 
lemma boundE [elim?]: "[| bound n f; (!! m. n < m ==> f m = 0) ==> P |] ==> P"  | 
|
| 21423 | 28  | 
unfolding bound_def by blast  | 
| 13936 | 29  | 
|
30  | 
lemma boundD [dest]: "[| bound n f; n < m |] ==> f m = 0"  | 
|
| 21423 | 31  | 
unfolding bound_def by blast  | 
| 13936 | 32  | 
|
33  | 
lemma bound_below:  | 
|
34  | 
assumes bound: "bound m f" and nonzero: "f n ~= 0" shows "n <= m"  | 
|
35  | 
proof (rule classical)  | 
|
36  | 
assume "~ ?thesis"  | 
|
37  | 
then have "m < n" by arith  | 
|
38  | 
with bound have "f n = 0" ..  | 
|
39  | 
with nonzero show ?thesis by contradiction  | 
|
40  | 
qed  | 
|
41  | 
||
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
42  | 
|
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
43  | 
definition "UP = {f :: nat => 'a::zero. EX n. bound n f}"
 | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
44  | 
|
| 49834 | 45  | 
typedef 'a up = "UP :: (nat => 'a::zero) set"  | 
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
46  | 
morphisms Rep_UP Abs_UP  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
47  | 
proof -  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
48  | 
have "bound 0 (\<lambda>_. 0::'a)" by (rule boundI) (rule refl)  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
49  | 
then show ?thesis unfolding UP_def by blast  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
50  | 
qed  | 
| 21423 | 51  | 
|
| 13936 | 52  | 
|
53  | 
section {* Constants *}
 | 
|
54  | 
||
| 21423 | 55  | 
definition  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
56  | 
  coeff :: "['a up, nat] => ('a::zero)"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
57  | 
where "coeff p n = Rep_UP p n"  | 
| 13936 | 58  | 
|
| 21423 | 59  | 
definition  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
60  | 
  monom :: "['a::zero, nat] => 'a up"  ("(3_*X^/_)" [71, 71] 70)
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
61  | 
where "monom a n = Abs_UP (%i. if i=n then a else 0)"  | 
| 21423 | 62  | 
|
63  | 
definition  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
64  | 
  smult :: "['a::{zero, times}, 'a up] => 'a up"  (infixl "*s" 70)
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
65  | 
where "a *s p = Abs_UP (%i. a * Rep_UP p i)"  | 
| 13936 | 66  | 
|
67  | 
lemma coeff_bound_ex: "EX n. bound n (coeff p)"  | 
|
68  | 
proof -  | 
|
69  | 
have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)  | 
|
70  | 
then obtain n where "bound n (coeff p)" by (unfold UP_def) fast  | 
|
71  | 
then show ?thesis ..  | 
|
72  | 
qed  | 
|
73  | 
||
74  | 
lemma bound_coeff_obtain:  | 
|
75  | 
assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P"  | 
|
76  | 
proof -  | 
|
77  | 
have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)  | 
|
78  | 
then obtain n where "bound n (coeff p)" by (unfold UP_def) fast  | 
|
79  | 
with prem show P .  | 
|
80  | 
qed  | 
|
81  | 
||
| 21423 | 82  | 
|
| 13936 | 83  | 
text {* Ring operations *}
 | 
84  | 
||
| 25762 | 85  | 
instantiation up :: (zero) zero  | 
86  | 
begin  | 
|
87  | 
||
88  | 
definition  | 
|
89  | 
up_zero_def: "0 = monom 0 0"  | 
|
90  | 
||
91  | 
instance ..  | 
|
92  | 
||
93  | 
end  | 
|
94  | 
||
95  | 
instantiation up :: ("{one, zero}") one
 | 
|
96  | 
begin  | 
|
97  | 
||
98  | 
definition  | 
|
99  | 
up_one_def: "1 = monom 1 0"  | 
|
100  | 
||
101  | 
instance ..  | 
|
102  | 
||
103  | 
end  | 
|
104  | 
||
105  | 
instantiation up :: ("{plus, zero}") plus
 | 
|
106  | 
begin  | 
|
107  | 
||
108  | 
definition  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
109  | 
up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"  | 
| 25762 | 110  | 
|
111  | 
instance ..  | 
|
| 13936 | 112  | 
|
| 25762 | 113  | 
end  | 
114  | 
||
115  | 
instantiation up :: ("{one, times, uminus, zero}") uminus
 | 
|
116  | 
begin  | 
|
117  | 
||
118  | 
definition  | 
|
119  | 
(* note: - 1 is different from -1; latter is of class number *)  | 
|
120  | 
up_uminus_def:"- p = (- 1) *s p"  | 
|
121  | 
(* easier to use than "Abs_UP (%i. - Rep_UP p i)" *)  | 
|
122  | 
||
123  | 
instance ..  | 
|
124  | 
||
125  | 
end  | 
|
126  | 
||
127  | 
instantiation up :: ("{one, plus, times, minus, uminus, zero}") minus
 | 
|
128  | 
begin  | 
|
129  | 
||
130  | 
definition  | 
|
131  | 
up_minus_def: "(a :: 'a up) - b = a + (-b)"  | 
|
132  | 
||
133  | 
instance ..  | 
|
134  | 
||
135  | 
end  | 
|
136  | 
||
137  | 
instantiation up :: ("{times, comm_monoid_add}") times
 | 
|
138  | 
begin  | 
|
139  | 
||
140  | 
definition  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
141  | 
up_mult_def: "p * q = Abs_UP (%n::nat. setsum  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
142  | 
                     (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
 | 
| 25762 | 143  | 
|
144  | 
instance ..  | 
|
145  | 
||
146  | 
end  | 
|
147  | 
||
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
33657 
diff
changeset
 | 
148  | 
instance up :: ("{times, comm_monoid_add}") Rings.dvd ..
 | 
| 26563 | 149  | 
|
150  | 
instantiation up :: ("{times, one, comm_monoid_add, uminus, minus}") inverse
 | 
|
| 25762 | 151  | 
begin  | 
152  | 
||
153  | 
definition  | 
|
154  | 
up_inverse_def: "inverse (a :: 'a up) = (if a dvd 1 then  | 
|
155  | 
THE x. a * x = 1 else 0)"  | 
|
| 13936 | 156  | 
|
| 25762 | 157  | 
definition  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
158  | 
up_divide_def: "(a :: 'a up) / b = a * inverse b"  | 
| 25762 | 159  | 
|
160  | 
instance ..  | 
|
161  | 
||
162  | 
end  | 
|
| 13936 | 163  | 
|
| 21423 | 164  | 
|
| 13936 | 165  | 
subsection {* Effect of operations on coefficients *}
 | 
166  | 
||
167  | 
lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)"  | 
|
168  | 
proof -  | 
|
169  | 
have "(%n. if n = m then a else 0) : UP"  | 
|
170  | 
using UP_def by force  | 
|
171  | 
from this show ?thesis  | 
|
172  | 
by (simp add: coeff_def monom_def Abs_UP_inverse Rep_UP)  | 
|
173  | 
qed  | 
|
174  | 
||
175  | 
lemma coeff_zero [simp]: "coeff 0 n = 0"  | 
|
176  | 
proof (unfold up_zero_def)  | 
|
177  | 
qed simp  | 
|
178  | 
||
179  | 
lemma coeff_one [simp]: "coeff 1 n = (if n=0 then 1 else 0)"  | 
|
180  | 
proof (unfold up_one_def)  | 
|
181  | 
qed simp  | 
|
182  | 
||
183  | 
(* term order  | 
|
184  | 
lemma coeff_smult [simp]: "coeff (a *s p) n = (a::'a::ring) * coeff p n"  | 
|
185  | 
proof -  | 
|
186  | 
have "!!f. f : UP ==> (%n. a * f n) : UP"  | 
|
| 29667 | 187  | 
by (unfold UP_def) (force simp add: algebra_simps)  | 
| 13936 | 188  | 
*) (* this force step is slow *)  | 
189  | 
(* then show ?thesis  | 
|
190  | 
apply (simp add: coeff_def smult_def Abs_UP_inverse Rep_UP)  | 
|
191  | 
qed  | 
|
192  | 
*)  | 
|
193  | 
lemma coeff_smult [simp]: "coeff (a *s p) n = (a::'a::ring) * coeff p n"  | 
|
194  | 
proof -  | 
|
195  | 
have "Rep_UP p : UP ==> (%n. a * Rep_UP p n) : UP"  | 
|
| 29667 | 196  | 
by (unfold UP_def) (force simp add: algebra_simps)  | 
| 13936 | 197  | 
(* this force step is slow *)  | 
198  | 
then show ?thesis  | 
|
199  | 
by (simp add: coeff_def smult_def Abs_UP_inverse Rep_UP)  | 
|
200  | 
qed  | 
|
201  | 
||
202  | 
lemma coeff_add [simp]: "coeff (p+q) n = (coeff p n + coeff q n::'a::ring)"  | 
|
203  | 
proof -  | 
|
204  | 
  {
 | 
|
205  | 
fix f g  | 
|
206  | 
assume fup: "(f::nat=>'a::ring) : UP" and gup: "(g::nat=>'a::ring) : UP"  | 
|
207  | 
have "(%i. f i + g i) : UP"  | 
|
208  | 
proof -  | 
|
209  | 
from fup obtain n where boundn: "bound n f"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
210  | 
by (unfold UP_def) fast  | 
| 13936 | 211  | 
from gup obtain m where boundm: "bound m g"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
212  | 
by (unfold UP_def) fast  | 
| 13936 | 213  | 
have "bound (max n m) (%i. (f i + g i))"  | 
214  | 
proof  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
215  | 
fix i  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
216  | 
assume "max n m < i"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
217  | 
with boundn and boundm show "f i + g i = 0"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42768 
diff
changeset
 | 
218  | 
by (fastforce simp add: algebra_simps)  | 
| 13936 | 219  | 
qed  | 
220  | 
then show "(%i. (f i + g i)) : UP"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
221  | 
by (unfold UP_def) fast  | 
| 13936 | 222  | 
qed  | 
223  | 
}  | 
|
224  | 
then show ?thesis  | 
|
225  | 
by (simp add: coeff_def up_add_def Abs_UP_inverse Rep_UP)  | 
|
226  | 
qed  | 
|
227  | 
||
228  | 
lemma coeff_mult [simp]:  | 
|
229  | 
  "coeff (p * q) n = (setsum (%i. coeff p i * coeff q (n-i)) {..n}::'a::ring)"
 | 
|
230  | 
proof -  | 
|
231  | 
  {
 | 
|
232  | 
fix f g  | 
|
233  | 
assume fup: "(f::nat=>'a::ring) : UP" and gup: "(g::nat=>'a::ring) : UP"  | 
|
234  | 
    have "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
 | 
|
235  | 
proof -  | 
|
236  | 
from fup obtain n where "bound n f"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
237  | 
by (unfold UP_def) fast  | 
| 13936 | 238  | 
from gup obtain m where "bound m g"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
239  | 
by (unfold UP_def) fast  | 
| 13936 | 240  | 
      have "bound (n + m) (%n. setsum (%i. f i * g (n-i)) {..n})"
 | 
241  | 
proof  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
242  | 
fix k  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
243  | 
assume bound: "n + m < k"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
244  | 
        {
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
245  | 
fix i  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
246  | 
have "f i * g (k-i) = 0"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
247  | 
proof cases  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
248  | 
assume "n < i"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
249  | 
with `bound n f` show ?thesis by (auto simp add: algebra_simps)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
250  | 
next  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
251  | 
assume "~ (n < i)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
252  | 
with bound have "m < k-i" by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
253  | 
with `bound m g` show ?thesis by (auto simp add: algebra_simps)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
254  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
255  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
256  | 
        then show "setsum (%i. f i * g (k-i)) {..k} = 0"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
257  | 
by (simp add: algebra_simps)  | 
| 13936 | 258  | 
qed  | 
259  | 
      then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
 | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
260  | 
by (unfold UP_def) fast  | 
| 13936 | 261  | 
qed  | 
262  | 
}  | 
|
263  | 
then show ?thesis  | 
|
264  | 
by (simp add: coeff_def up_mult_def Abs_UP_inverse Rep_UP)  | 
|
265  | 
qed  | 
|
266  | 
||
267  | 
lemma coeff_uminus [simp]: "coeff (-p) n = (-coeff p n::'a::ring)"  | 
|
| 29667 | 268  | 
by (unfold up_uminus_def) (simp add: algebra_simps)  | 
| 13936 | 269  | 
|
270  | 
(* Other lemmas *)  | 
|
271  | 
||
272  | 
lemma up_eqI: assumes prem: "(!! n. coeff p n = coeff q n)" shows "p = q"  | 
|
273  | 
proof -  | 
|
274  | 
have "p = Abs_UP (%u. Rep_UP p u)" by (simp add: Rep_UP_inverse)  | 
|
275  | 
also from prem have "... = Abs_UP (Rep_UP q)" by (simp only: coeff_def)  | 
|
276  | 
also have "... = q" by (simp add: Rep_UP_inverse)  | 
|
277  | 
finally show ?thesis .  | 
|
278  | 
qed  | 
|
279  | 
||
280  | 
instance up :: (ring) ring  | 
|
281  | 
proof  | 
|
282  | 
fix p q r :: "'a::ring up"  | 
|
283  | 
show "(p + q) + r = p + (q + r)"  | 
|
284  | 
by (rule up_eqI) simp  | 
|
285  | 
show "0 + p = p"  | 
|
286  | 
by (rule up_eqI) simp  | 
|
287  | 
show "(-p) + p = 0"  | 
|
288  | 
by (rule up_eqI) simp  | 
|
289  | 
show "p + q = q + p"  | 
|
290  | 
by (rule up_eqI) simp  | 
|
291  | 
show "(p * q) * r = p * (q * r)"  | 
|
292  | 
proof (rule up_eqI)  | 
|
293  | 
fix n  | 
|
294  | 
    {
 | 
|
295  | 
fix k and a b c :: "nat=>'a::ring"  | 
|
296  | 
have "k <= n ==>  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
297  | 
        setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} = 
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
298  | 
        setsum (%j. a j * setsum  (%i. b i * c (n-j-i)) {..k-j}) {..k}"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
299  | 
(is "_ ==> ?eq k")  | 
| 13936 | 300  | 
proof (induct k)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
301  | 
case 0 show ?case by simp  | 
| 13936 | 302  | 
next  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
303  | 
case (Suc k)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
304  | 
then have "k <= n" by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
305  | 
then have "?eq k" by (rule Suc)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
306  | 
then show ?case  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
307  | 
by (simp add: Suc_diff_le natsum_ldistr)  | 
| 13936 | 308  | 
qed  | 
309  | 
}  | 
|
310  | 
then show "coeff ((p * q) * r) n = coeff (p * (q * r)) n"  | 
|
311  | 
by simp  | 
|
312  | 
qed  | 
|
313  | 
show "1 * p = p"  | 
|
314  | 
proof (rule up_eqI)  | 
|
315  | 
fix n  | 
|
316  | 
show "coeff (1 * p) n = coeff p n"  | 
|
317  | 
proof (cases n)  | 
|
318  | 
case 0 then show ?thesis by simp  | 
|
319  | 
next  | 
|
| 16052 | 320  | 
case Suc then show ?thesis by (simp del: setsum_atMost_Suc add: natsum_Suc2)  | 
| 13936 | 321  | 
qed  | 
322  | 
qed  | 
|
323  | 
show "(p + q) * r = p * r + q * r"  | 
|
324  | 
by (rule up_eqI) simp  | 
|
| 
30968
 
10fef94f40fc
adaptions due to rearrangment of power operation
 
haftmann 
parents: 
30012 
diff
changeset
 | 
325  | 
show "\<And>q. p * q = q * p"  | 
| 13936 | 326  | 
proof (rule up_eqI)  | 
| 
30968
 
10fef94f40fc
adaptions due to rearrangment of power operation
 
haftmann 
parents: 
30012 
diff
changeset
 | 
327  | 
fix q  | 
| 13936 | 328  | 
fix n  | 
329  | 
    {
 | 
|
330  | 
fix k  | 
|
331  | 
fix a b :: "nat=>'a::ring"  | 
|
332  | 
have "k <= n ==>  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
333  | 
        setsum (%i. a i * b (n-i)) {..k} =
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
334  | 
        setsum (%i. a (k-i) * b (i+n-k)) {..k}"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
335  | 
(is "_ ==> ?eq k")  | 
| 13936 | 336  | 
proof (induct k)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
337  | 
case 0 show ?case by simp  | 
| 13936 | 338  | 
next  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
339  | 
case (Suc k) then show ?case by (subst natsum_Suc2) simp  | 
| 13936 | 340  | 
qed  | 
341  | 
}  | 
|
342  | 
then show "coeff (p * q) n = coeff (q * p) n"  | 
|
343  | 
by simp  | 
|
344  | 
qed  | 
|
345  | 
||
346  | 
show "p - q = p + (-q)"  | 
|
347  | 
by (simp add: up_minus_def)  | 
|
348  | 
show "inverse p = (if p dvd 1 then THE x. p*x = 1 else 0)"  | 
|
349  | 
by (simp add: up_inverse_def)  | 
|
350  | 
show "p / q = p * inverse q"  | 
|
351  | 
by (simp add: up_divide_def)  | 
|
| 27540 | 352  | 
qed  | 
| 13936 | 353  | 
|
354  | 
(* Further properties of monom *)  | 
|
355  | 
||
356  | 
lemma monom_zero [simp]:  | 
|
357  | 
"monom 0 n = 0"  | 
|
358  | 
by (simp add: monom_def up_zero_def)  | 
|
359  | 
(* term order: application of coeff_mult goes wrong: rule not symmetric  | 
|
360  | 
lemma monom_mult_is_smult:  | 
|
361  | 
"monom (a::'a::ring) 0 * p = a *s p"  | 
|
362  | 
proof (rule up_eqI)  | 
|
363  | 
fix k  | 
|
364  | 
show "coeff (monom a 0 * p) k = coeff (a *s p) k"  | 
|
365  | 
proof (cases k)  | 
|
366  | 
case 0 then show ?thesis by simp  | 
|
367  | 
next  | 
|
368  | 
case Suc then show ?thesis by simp  | 
|
369  | 
qed  | 
|
370  | 
qed  | 
|
371  | 
*)  | 
|
372  | 
||
373  | 
lemma monom_mult_is_smult:  | 
|
374  | 
"monom (a::'a::ring) 0 * p = a *s p"  | 
|
375  | 
proof (rule up_eqI)  | 
|
| 42768 | 376  | 
note [[simproc del: ring]]  | 
| 13936 | 377  | 
fix k  | 
378  | 
have "coeff (p * monom a 0) k = coeff (a *s p) k"  | 
|
379  | 
proof (cases k)  | 
|
380  | 
case 0 then show ?thesis by simp ring  | 
|
381  | 
next  | 
|
| 30012 | 382  | 
case Suc then show ?thesis by simp (ring, simp)  | 
| 13936 | 383  | 
qed  | 
384  | 
then show "coeff (monom a 0 * p) k = coeff (a *s p) k" by ring  | 
|
385  | 
qed  | 
|
386  | 
||
387  | 
lemma monom_add [simp]:  | 
|
388  | 
"monom (a + b) n = monom (a::'a::ring) n + monom b n"  | 
|
389  | 
by (rule up_eqI) simp  | 
|
390  | 
||
391  | 
lemma monom_mult_smult:  | 
|
392  | 
"monom (a * b) n = a *s monom (b::'a::ring) n"  | 
|
393  | 
by (rule up_eqI) simp  | 
|
394  | 
||
395  | 
lemma monom_uminus [simp]:  | 
|
396  | 
"monom (-a) n = - monom (a::'a::ring) n"  | 
|
397  | 
by (rule up_eqI) simp  | 
|
398  | 
||
399  | 
lemma monom_one [simp]:  | 
|
400  | 
"monom 1 0 = 1"  | 
|
401  | 
by (simp add: up_one_def)  | 
|
402  | 
||
403  | 
lemma monom_inj:  | 
|
404  | 
"(monom a n = monom b n) = (a = b)"  | 
|
405  | 
proof  | 
|
406  | 
assume "monom a n = monom b n"  | 
|
407  | 
then have "coeff (monom a n) n = coeff (monom b n) n" by simp  | 
|
408  | 
then show "a = b" by simp  | 
|
409  | 
next  | 
|
410  | 
assume "a = b" then show "monom a n = monom b n" by simp  | 
|
411  | 
qed  | 
|
412  | 
||
413  | 
(* Properties of *s:  | 
|
414  | 
Polynomials form a module *)  | 
|
415  | 
||
416  | 
lemma smult_l_distr:  | 
|
417  | 
"(a + b::'a::ring) *s p = a *s p + b *s p"  | 
|
418  | 
by (rule up_eqI) simp  | 
|
419  | 
||
420  | 
lemma smult_r_distr:  | 
|
421  | 
"(a::'a::ring) *s (p + q) = a *s p + a *s q"  | 
|
422  | 
by (rule up_eqI) simp  | 
|
423  | 
||
424  | 
lemma smult_assoc1:  | 
|
425  | 
"(a * b::'a::ring) *s p = a *s (b *s p)"  | 
|
426  | 
by (rule up_eqI) simp  | 
|
427  | 
||
428  | 
lemma smult_one [simp]:  | 
|
429  | 
"(1::'a::ring) *s p = p"  | 
|
430  | 
by (rule up_eqI) simp  | 
|
431  | 
||
432  | 
(* Polynomials form an algebra *)  | 
|
433  | 
||
434  | 
lemma smult_assoc2:  | 
|
435  | 
"(a *s p) * q = (a::'a::ring) *s (p * q)"  | 
|
| 42768 | 436  | 
using [[simproc del: ring]]  | 
| 13936 | 437  | 
by (rule up_eqI) (simp add: natsum_rdistr m_assoc)  | 
438  | 
||
439  | 
(* the following can be derived from the above ones,  | 
|
440  | 
for generality reasons, it is therefore done *)  | 
|
441  | 
||
442  | 
lemma smult_l_null [simp]:  | 
|
443  | 
"(0::'a::ring) *s p = 0"  | 
|
444  | 
proof -  | 
|
445  | 
fix a  | 
|
446  | 
have "0 *s p = (0 *s p + a *s p) + - (a *s p)" by simp  | 
|
447  | 
also have "... = (0 + a) *s p + - (a *s p)" by (simp only: smult_l_distr)  | 
|
448  | 
also have "... = 0" by simp  | 
|
449  | 
finally show ?thesis .  | 
|
450  | 
qed  | 
|
451  | 
||
452  | 
lemma smult_r_null [simp]:  | 
|
| 42768 | 453  | 
"(a::'a::ring) *s 0 = 0"  | 
| 13936 | 454  | 
proof -  | 
455  | 
fix p  | 
|
456  | 
have "a *s 0 = (a *s 0 + a *s p) + - (a *s p)" by simp  | 
|
457  | 
also have "... = a *s (0 + p) + - (a *s p)" by (simp only: smult_r_distr)  | 
|
458  | 
also have "... = 0" by simp  | 
|
459  | 
finally show ?thesis .  | 
|
460  | 
qed  | 
|
461  | 
||
462  | 
lemma smult_l_minus:  | 
|
463  | 
"(-a::'a::ring) *s p = - (a *s p)"  | 
|
464  | 
proof -  | 
|
465  | 
have "(-a) *s p = (-a *s p + a *s p) + -(a *s p)" by simp  | 
|
466  | 
also have "... = (-a + a) *s p + -(a *s p)" by (simp only: smult_l_distr)  | 
|
467  | 
also have "... = -(a *s p)" by simp  | 
|
468  | 
finally show ?thesis .  | 
|
469  | 
qed  | 
|
470  | 
||
471  | 
lemma smult_r_minus:  | 
|
472  | 
"(a::'a::ring) *s (-p) = - (a *s p)"  | 
|
473  | 
proof -  | 
|
474  | 
have "a *s (-p) = (a *s -p + a *s p) + -(a *s p)" by simp  | 
|
475  | 
also have "... = a *s (-p + p) + -(a *s p)" by (simp only: smult_r_distr)  | 
|
476  | 
also have "... = -(a *s p)" by simp  | 
|
477  | 
finally show ?thesis .  | 
|
478  | 
qed  | 
|
479  | 
||
| 35849 | 480  | 
|
| 13936 | 481  | 
section {* The degree function *}
 | 
482  | 
||
| 21423 | 483  | 
definition  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
484  | 
  deg :: "('a::zero) up => nat"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35050 
diff
changeset
 | 
485  | 
where "deg p = (LEAST n. bound n (coeff p))"  | 
| 13936 | 486  | 
|
487  | 
lemma deg_aboveI:  | 
|
488  | 
"(!!m. n < m ==> coeff p m = 0) ==> deg p <= n"  | 
|
489  | 
by (unfold deg_def) (fast intro: Least_le)  | 
|
490  | 
||
491  | 
lemma deg_aboveD:  | 
|
| 23350 | 492  | 
assumes "deg p < m" shows "coeff p m = 0"  | 
| 13936 | 493  | 
proof -  | 
494  | 
obtain n where "bound n (coeff p)" by (rule bound_coeff_obtain)  | 
|
495  | 
then have "bound (deg p) (coeff p)" by (unfold deg_def, rule LeastI)  | 
|
| 23350 | 496  | 
then show "coeff p m = 0" using `deg p < m` by (rule boundD)  | 
| 13936 | 497  | 
qed  | 
498  | 
||
499  | 
lemma deg_belowI:  | 
|
500  | 
assumes prem: "n ~= 0 ==> coeff p n ~= 0" shows "n <= deg p"  | 
|
501  | 
(* logically, this is a slightly stronger version of deg_aboveD *)  | 
|
502  | 
proof (cases "n=0")  | 
|
503  | 
case True then show ?thesis by simp  | 
|
504  | 
next  | 
|
505  | 
case False then have "coeff p n ~= 0" by (rule prem)  | 
|
506  | 
then have "~ deg p < n" by (fast dest: deg_aboveD)  | 
|
507  | 
then show ?thesis by arith  | 
|
508  | 
qed  | 
|
509  | 
||
510  | 
lemma lcoeff_nonzero_deg:  | 
|
511  | 
assumes deg: "deg p ~= 0" shows "coeff p (deg p) ~= 0"  | 
|
512  | 
proof -  | 
|
513  | 
obtain m where "deg p <= m" and m_coeff: "coeff p m ~= 0"  | 
|
514  | 
proof -  | 
|
515  | 
have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"  | 
|
516  | 
by arith (* make public?, why does proof not work with "1" *)  | 
|
517  | 
from deg have "deg p - 1 < (LEAST n. bound n (coeff p))"  | 
|
518  | 
by (unfold deg_def) arith  | 
|
519  | 
then have "~ bound (deg p - 1) (coeff p)" by (rule not_less_Least)  | 
|
520  | 
then have "EX m. deg p - 1 < m & coeff p m ~= 0"  | 
|
521  | 
by (unfold bound_def) fast  | 
|
522  | 
then have "EX m. deg p <= m & coeff p m ~= 0" by (simp add: deg minus)  | 
|
| 23350 | 523  | 
then show ?thesis by (auto intro: that)  | 
| 13936 | 524  | 
qed  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42768 
diff
changeset
 | 
525  | 
with deg_belowI have "deg p = m" by fastforce  | 
| 13936 | 526  | 
with m_coeff show ?thesis by simp  | 
527  | 
qed  | 
|
528  | 
||
529  | 
lemma lcoeff_nonzero_nonzero:  | 
|
530  | 
assumes deg: "deg p = 0" and nonzero: "p ~= 0" shows "coeff p 0 ~= 0"  | 
|
531  | 
proof -  | 
|
532  | 
have "EX m. coeff p m ~= 0"  | 
|
533  | 
proof (rule classical)  | 
|
534  | 
assume "~ ?thesis"  | 
|
535  | 
then have "p = 0" by (auto intro: up_eqI)  | 
|
536  | 
with nonzero show ?thesis by contradiction  | 
|
537  | 
qed  | 
|
538  | 
then obtain m where coeff: "coeff p m ~= 0" ..  | 
|
539  | 
then have "m <= deg p" by (rule deg_belowI)  | 
|
540  | 
then have "m = 0" by (simp add: deg)  | 
|
541  | 
with coeff show ?thesis by simp  | 
|
542  | 
qed  | 
|
543  | 
||
544  | 
lemma lcoeff_nonzero:  | 
|
545  | 
"p ~= 0 ==> coeff p (deg p) ~= 0"  | 
|
546  | 
proof (cases "deg p = 0")  | 
|
547  | 
case True  | 
|
548  | 
assume "p ~= 0"  | 
|
549  | 
with True show ?thesis by (simp add: lcoeff_nonzero_nonzero)  | 
|
550  | 
next  | 
|
551  | 
case False  | 
|
552  | 
assume "p ~= 0"  | 
|
553  | 
with False show ?thesis by (simp add: lcoeff_nonzero_deg)  | 
|
554  | 
qed  | 
|
555  | 
||
556  | 
lemma deg_eqI:  | 
|
557  | 
"[| !!m. n < m ==> coeff p m = 0;  | 
|
558  | 
!!n. n ~= 0 ==> coeff p n ~= 0|] ==> deg p = n"  | 
|
| 33657 | 559  | 
by (fast intro: le_antisym deg_aboveI deg_belowI)  | 
| 13936 | 560  | 
|
561  | 
(* Degree and polynomial operations *)  | 
|
562  | 
||
563  | 
lemma deg_add [simp]:  | 
|
564  | 
"deg ((p::'a::ring up) + q) <= max (deg p) (deg q)"  | 
|
| 
32436
 
10cd49e0c067
Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
 
nipkow 
parents: 
31021 
diff
changeset
 | 
565  | 
by (rule deg_aboveI) (simp add: deg_aboveD)  | 
| 13936 | 566  | 
|
567  | 
lemma deg_monom_ring:  | 
|
568  | 
"deg (monom a n::'a::ring up) <= n"  | 
|
569  | 
by (rule deg_aboveI) simp  | 
|
570  | 
||
571  | 
lemma deg_monom [simp]:  | 
|
572  | 
"a ~= 0 ==> deg (monom a n::'a::ring up) = n"  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42768 
diff
changeset
 | 
573  | 
by (fastforce intro: le_antisym deg_aboveI deg_belowI)  | 
| 13936 | 574  | 
|
575  | 
lemma deg_const [simp]:  | 
|
576  | 
"deg (monom (a::'a::ring) 0) = 0"  | 
|
| 33657 | 577  | 
proof (rule le_antisym)  | 
| 13936 | 578  | 
show "deg (monom a 0) <= 0" by (rule deg_aboveI) simp  | 
579  | 
next  | 
|
580  | 
show "0 <= deg (monom a 0)" by (rule deg_belowI) simp  | 
|
581  | 
qed  | 
|
582  | 
||
583  | 
lemma deg_zero [simp]:  | 
|
584  | 
"deg 0 = 0"  | 
|
| 33657 | 585  | 
proof (rule le_antisym)  | 
| 13936 | 586  | 
show "deg 0 <= 0" by (rule deg_aboveI) simp  | 
587  | 
next  | 
|
588  | 
show "0 <= deg 0" by (rule deg_belowI) simp  | 
|
589  | 
qed  | 
|
590  | 
||
591  | 
lemma deg_one [simp]:  | 
|
592  | 
"deg 1 = 0"  | 
|
| 33657 | 593  | 
proof (rule le_antisym)  | 
| 13936 | 594  | 
show "deg 1 <= 0" by (rule deg_aboveI) simp  | 
595  | 
next  | 
|
596  | 
show "0 <= deg 1" by (rule deg_belowI) simp  | 
|
597  | 
qed  | 
|
598  | 
||
599  | 
lemma uminus_monom:  | 
|
600  | 
"!!a::'a::ring. (-a = 0) = (a = 0)"  | 
|
601  | 
proof  | 
|
602  | 
fix a::"'a::ring"  | 
|
603  | 
assume "a = 0"  | 
|
604  | 
then show "-a = 0" by simp  | 
|
605  | 
next  | 
|
606  | 
fix a::"'a::ring"  | 
|
607  | 
assume "- a = 0"  | 
|
608  | 
then have "-(- a) = 0" by simp  | 
|
609  | 
then show "a = 0" by simp  | 
|
610  | 
qed  | 
|
611  | 
||
612  | 
lemma deg_uminus [simp]:  | 
|
613  | 
  "deg (-p::('a::ring) up) = deg p"
 | 
|
| 33657 | 614  | 
proof (rule le_antisym)  | 
| 13936 | 615  | 
show "deg (- p) <= deg p" by (simp add: deg_aboveI deg_aboveD)  | 
616  | 
next  | 
|
617  | 
show "deg p <= deg (- p)"  | 
|
618  | 
by (simp add: deg_belowI lcoeff_nonzero_deg uminus_monom)  | 
|
619  | 
qed  | 
|
620  | 
||
621  | 
lemma deg_smult_ring:  | 
|
622  | 
"deg ((a::'a::ring) *s p) <= (if a = 0 then 0 else deg p)"  | 
|
623  | 
proof (cases "a = 0")  | 
|
624  | 
qed (simp add: deg_aboveI deg_aboveD)+  | 
|
625  | 
||
626  | 
lemma deg_smult [simp]:  | 
|
627  | 
"deg ((a::'a::domain) *s p) = (if a = 0 then 0 else deg p)"  | 
|
| 33657 | 628  | 
proof (rule le_antisym)  | 
| 13936 | 629  | 
show "deg (a *s p) <= (if a = 0 then 0 else deg p)" by (rule deg_smult_ring)  | 
630  | 
next  | 
|
631  | 
show "(if a = 0 then 0 else deg p) <= deg (a *s p)"  | 
|
632  | 
proof (cases "a = 0")  | 
|
633  | 
qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff)  | 
|
634  | 
qed  | 
|
635  | 
||
636  | 
lemma deg_mult_ring:  | 
|
637  | 
"deg (p * q::'a::ring up) <= deg p + deg q"  | 
|
638  | 
proof (rule deg_aboveI)  | 
|
639  | 
fix m  | 
|
640  | 
assume boundm: "deg p + deg q < m"  | 
|
641  | 
  {
 | 
|
642  | 
fix k i  | 
|
643  | 
assume boundk: "deg p + deg q < k"  | 
|
644  | 
then have "coeff p i * coeff q (k - i) = 0"  | 
|
645  | 
proof (cases "deg p < i")  | 
|
646  | 
case True then show ?thesis by (simp add: deg_aboveD)  | 
|
647  | 
next  | 
|
648  | 
case False with boundk have "deg q < k - i" by arith  | 
|
649  | 
then show ?thesis by (simp add: deg_aboveD)  | 
|
650  | 
qed  | 
|
651  | 
}  | 
|
652  | 
(* This is similar to bound_mult_zero and deg_above_mult_zero in the old  | 
|
653  | 
proofs. *)  | 
|
654  | 
with boundm show "coeff (p * q) m = 0" by simp  | 
|
655  | 
qed  | 
|
656  | 
||
657  | 
lemma deg_mult [simp]:  | 
|
658  | 
"[| (p::'a::domain up) ~= 0; q ~= 0|] ==> deg (p * q) = deg p + deg q"  | 
|
| 33657 | 659  | 
proof (rule le_antisym)  | 
| 13936 | 660  | 
show "deg (p * q) <= deg p + deg q" by (rule deg_mult_ring)  | 
661  | 
next  | 
|
662  | 
let ?s = "(%i. coeff p i * coeff q (deg p + deg q - i))"  | 
|
663  | 
assume nz: "p ~= 0" "q ~= 0"  | 
|
664  | 
have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith  | 
|
665  | 
show "deg p + deg q <= deg (p * q)"  | 
|
666  | 
proof (rule deg_belowI, simp)  | 
|
667  | 
    have "setsum ?s {.. deg p + deg q}
 | 
|
| 15045 | 668  | 
      = setsum ?s ({..< deg p} Un {deg p .. deg p + deg q})"
 | 
| 13936 | 669  | 
by (simp only: ivl_disj_un_one)  | 
670  | 
    also have "... = setsum ?s {deg p .. deg p + deg q}"
 | 
|
671  | 
by (simp add: setsum_Un_disjoint ivl_disj_int_one  | 
|
672  | 
setsum_0 deg_aboveD less_add_diff)  | 
|
| 15045 | 673  | 
    also have "... = setsum ?s ({deg p} Un {deg p <.. deg p + deg q})"
 | 
| 13936 | 674  | 
by (simp only: ivl_disj_un_singleton)  | 
675  | 
also have "... = coeff p (deg p) * coeff q (deg q)"  | 
|
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
676  | 
by (simp add: setsum_Un_disjoint setsum_0 deg_aboveD)  | 
| 13936 | 677  | 
    finally have "setsum ?s {.. deg p + deg q} 
 | 
678  | 
= coeff p (deg p) * coeff q (deg q)" .  | 
|
679  | 
    with nz show "setsum ?s {.. deg p + deg q} ~= 0"
 | 
|
680  | 
by (simp add: integral_iff lcoeff_nonzero)  | 
|
681  | 
qed  | 
|
682  | 
qed  | 
|
683  | 
||
684  | 
lemma coeff_natsum:  | 
|
685  | 
"((coeff (setsum p A) k)::'a::ring) =  | 
|
686  | 
setsum (%i. coeff (p i) k) A"  | 
|
687  | 
proof (cases "finite A")  | 
|
688  | 
case True then show ?thesis by induct auto  | 
|
689  | 
next  | 
|
690  | 
case False then show ?thesis by (simp add: setsum_def)  | 
|
691  | 
qed  | 
|
692  | 
(* Instance of a more general result!!! *)  | 
|
693  | 
||
694  | 
(*  | 
|
695  | 
lemma coeff_natsum:  | 
|
696  | 
  "((coeff (setsum p {..n::nat}) k)::'a::ring) = 
 | 
|
697  | 
   setsum (%i. coeff (p i) k) {..n}"
 | 
|
698  | 
by (induct n) auto  | 
|
699  | 
*)  | 
|
700  | 
||
701  | 
lemma up_repr:  | 
|
702  | 
  "setsum (%i. monom (coeff p i) i) {..deg (p::'a::ring up)} = p"
 | 
|
703  | 
proof (rule up_eqI)  | 
|
704  | 
let ?s = "(%i. monom (coeff p i) i)"  | 
|
705  | 
fix k  | 
|
706  | 
  show "coeff (setsum ?s {..deg p}) k = coeff p k"
 | 
|
707  | 
proof (cases "k <= deg p")  | 
|
708  | 
case True  | 
|
709  | 
    hence "coeff (setsum ?s {..deg p}) k = 
 | 
|
| 15045 | 710  | 
          coeff (setsum ?s ({..k} Un {k<..deg p})) k"
 | 
| 13936 | 711  | 
by (simp only: ivl_disj_un_one)  | 
712  | 
also from True  | 
|
713  | 
    have "... = coeff (setsum ?s {..k}) k"
 | 
|
714  | 
by (simp add: setsum_Un_disjoint ivl_disj_int_one order_less_imp_not_eq2  | 
|
715  | 
setsum_0 coeff_natsum )  | 
|
716  | 
also  | 
|
| 15045 | 717  | 
    have "... = coeff (setsum ?s ({..<k} Un {k})) k"
 | 
| 13936 | 718  | 
by (simp only: ivl_disj_un_singleton)  | 
719  | 
also have "... = coeff p k"  | 
|
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
720  | 
by (simp add: setsum_Un_disjoint setsum_0 coeff_natsum deg_aboveD)  | 
| 13936 | 721  | 
finally show ?thesis .  | 
722  | 
next  | 
|
723  | 
case False  | 
|
724  | 
    hence "coeff (setsum ?s {..deg p}) k = 
 | 
|
| 15045 | 725  | 
          coeff (setsum ?s ({..<deg p} Un {deg p})) k"
 | 
| 13936 | 726  | 
by (simp only: ivl_disj_un_singleton)  | 
727  | 
also from False have "... = coeff p k"  | 
|
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
728  | 
by (simp add: setsum_Un_disjoint setsum_0 coeff_natsum deg_aboveD)  | 
| 13936 | 729  | 
finally show ?thesis .  | 
730  | 
qed  | 
|
731  | 
qed  | 
|
732  | 
||
733  | 
lemma up_repr_le:  | 
|
734  | 
  "deg (p::'a::ring up) <= n ==> setsum (%i. monom (coeff p i) i) {..n} = p"
 | 
|
735  | 
proof -  | 
|
736  | 
let ?s = "(%i. monom (coeff p i) i)"  | 
|
737  | 
assume "deg p <= n"  | 
|
| 15045 | 738  | 
  then have "setsum ?s {..n} = setsum ?s ({..deg p} Un {deg p<..n})"
 | 
| 13936 | 739  | 
by (simp only: ivl_disj_un_one)  | 
740  | 
  also have "... = setsum ?s {..deg p}"
 | 
|
741  | 
by (simp add: setsum_Un_disjoint ivl_disj_int_one  | 
|
742  | 
setsum_0 deg_aboveD)  | 
|
743  | 
also have "... = p" by (rule up_repr)  | 
|
744  | 
finally show ?thesis .  | 
|
745  | 
qed  | 
|
746  | 
||
747  | 
instance up :: ("domain") "domain"
 | 
|
748  | 
proof  | 
|
749  | 
show "1 ~= (0::'a up)"  | 
|
750  | 
proof (* notI is applied here *)  | 
|
751  | 
assume "1 = (0::'a up)"  | 
|
752  | 
hence "coeff 1 0 = (coeff 0 0::'a)" by simp  | 
|
753  | 
hence "1 = (0::'a)" by simp  | 
|
754  | 
with one_not_zero show "False" by contradiction  | 
|
755  | 
qed  | 
|
756  | 
next  | 
|
757  | 
fix p q :: "'a::domain up"  | 
|
758  | 
assume pq: "p * q = 0"  | 
|
759  | 
show "p = 0 | q = 0"  | 
|
760  | 
proof (rule classical)  | 
|
761  | 
assume c: "~ (p = 0 | q = 0)"  | 
|
762  | 
then have "deg p + deg q = deg (p * q)" by simp  | 
|
763  | 
also from pq have "... = 0" by simp  | 
|
764  | 
finally have "deg p + deg q = 0" .  | 
|
765  | 
then have f1: "deg p = 0 & deg q = 0" by simp  | 
|
766  | 
    from f1 have "p = setsum (%i. (monom (coeff p i) i)) {..0}"
 | 
|
767  | 
by (simp only: up_repr_le)  | 
|
768  | 
also have "... = monom (coeff p 0) 0" by simp  | 
|
769  | 
finally have p: "p = monom (coeff p 0) 0" .  | 
|
770  | 
    from f1 have "q = setsum (%i. (monom (coeff q i) i)) {..0}"
 | 
|
771  | 
by (simp only: up_repr_le)  | 
|
772  | 
also have "... = monom (coeff q 0) 0" by simp  | 
|
773  | 
finally have q: "q = monom (coeff q 0) 0" .  | 
|
774  | 
have "coeff p 0 * coeff q 0 = coeff (p * q) 0" by simp  | 
|
775  | 
also from pq have "... = 0" by simp  | 
|
776  | 
finally have "coeff p 0 * coeff q 0 = 0" .  | 
|
777  | 
then have "coeff p 0 = 0 | coeff q 0 = 0" by (simp only: integral_iff)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42768 
diff
changeset
 | 
778  | 
with p q show "p = 0 | q = 0" by fastforce  | 
| 13936 | 779  | 
qed  | 
780  | 
qed  | 
|
781  | 
||
782  | 
lemma monom_inj_zero:  | 
|
783  | 
"(monom a n = 0) = (a = 0)"  | 
|
784  | 
proof -  | 
|
785  | 
have "(monom a n = 0) = (monom a n = monom 0 n)" by simp  | 
|
786  | 
also have "... = (a = 0)" by (simp add: monom_inj del: monom_zero)  | 
|
787  | 
finally show ?thesis .  | 
|
788  | 
qed  | 
|
789  | 
(* term order: makes this simpler!!!  | 
|
790  | 
lemma smult_integral:  | 
|
791  | 
"(a::'a::domain) *s p = 0 ==> a = 0 | p = 0"  | 
|
792  | 
by (simp add: monom_mult_is_smult [THEN sym] integral_iff monom_inj_zero) fast  | 
|
793  | 
*)  | 
|
794  | 
lemma smult_integral:  | 
|
795  | 
"(a::'a::domain) *s p = 0 ==> a = 0 | p = 0"  | 
|
796  | 
by (simp add: monom_mult_is_smult [THEN sym] integral_iff monom_inj_zero)  | 
|
797  | 
||
| 21423 | 798  | 
|
799  | 
(* Divisibility and degree *)  | 
|
800  | 
||
801  | 
lemma "!! p::'a::domain up. [| p dvd q; q ~= 0 |] ==> deg p <= deg q"  | 
|
802  | 
apply (unfold dvd_def)  | 
|
803  | 
apply (erule exE)  | 
|
804  | 
apply hypsubst  | 
|
805  | 
apply (case_tac "p = 0")  | 
|
806  | 
apply (case_tac [2] "k = 0")  | 
|
807  | 
apply auto  | 
|
808  | 
done  | 
|
809  | 
||
| 14590 | 810  | 
end  |