| author | haftmann | 
| Mon, 12 Jul 2010 08:58:12 +0200 | |
| changeset 37764 | 3489daf839d5 | 
| parent 36096 | abc6a2ea4b88 | 
| child 38131 | df8fc03995a4 | 
| permissions | -rw-r--r-- | 
| 35849 | 1  | 
(* Title: HOL/Algebra/UnivPoly.thy  | 
2  | 
Author: Clemens Ballarin, started 9 December 1996  | 
|
3  | 
Copyright: Clemens Ballarin  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
4  | 
|
| 27933 | 5  | 
Contributions, in particular on long division, by Jesus Aransay.  | 
| 13940 | 6  | 
*)  | 
7  | 
||
| 28823 | 8  | 
theory UnivPoly  | 
9  | 
imports Module RingHom  | 
|
10  | 
begin  | 
|
| 13940 | 11  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
12  | 
section {* Univariate Polynomials *}
 | 
| 13940 | 13  | 
|
| 14553 | 14  | 
text {*
 | 
| 14666 | 15  | 
Polynomials are formalised as modules with additional operations for  | 
16  | 
extracting coefficients from polynomials and for obtaining monomials  | 
|
17  | 
  from coefficients and exponents (record @{text "up_ring"}).  The
 | 
|
18  | 
carrier set is a set of bounded functions from Nat to the  | 
|
19  | 
coefficient domain. Bounded means that these functions return zero  | 
|
20  | 
above a certain bound (the degree). There is a chapter on the  | 
|
| 14706 | 21  | 
  formalisation of polynomials in the PhD thesis \cite{Ballarin:1999},
 | 
22  | 
which was implemented with axiomatic type classes. This was later  | 
|
23  | 
ported to Locales.  | 
|
| 14553 | 24  | 
*}  | 
25  | 
||
| 14666 | 26  | 
|
| 
13949
 
0ce528cd6f19
HOL-Algebra complete for release Isabelle2003 (modulo section headers).
 
ballarin 
parents: 
13940 
diff
changeset
 | 
27  | 
subsection {* The Constructor for Univariate Polynomials *}
 | 
| 13940 | 28  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
29  | 
text {*
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
30  | 
Functions with finite support.  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
31  | 
*}  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
32  | 
|
| 14666 | 33  | 
locale bound =  | 
34  | 
fixes z :: 'a  | 
|
35  | 
and n :: nat  | 
|
36  | 
and f :: "nat => 'a"  | 
|
37  | 
assumes bound: "!!m. n < m \<Longrightarrow> f m = z"  | 
|
| 13940 | 38  | 
|
| 14666 | 39  | 
declare bound.intro [intro!]  | 
40  | 
and bound.bound [dest]  | 
|
| 13940 | 41  | 
|
42  | 
lemma bound_below:  | 
|
| 14666 | 43  | 
assumes bound: "bound z m f" and nonzero: "f n \<noteq> z" shows "n \<le> m"  | 
| 13940 | 44  | 
proof (rule classical)  | 
45  | 
assume "~ ?thesis"  | 
|
46  | 
then have "m < n" by arith  | 
|
47  | 
with bound have "f n = z" ..  | 
|
48  | 
with nonzero show ?thesis by contradiction  | 
|
49  | 
qed  | 
|
50  | 
||
51  | 
record ('a, 'p) up_ring = "('a, 'p) module" +
 | 
|
52  | 
monom :: "['a, nat] => 'p"  | 
|
53  | 
coeff :: "['p, nat] => 'a"  | 
|
54  | 
||
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
55  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
56  | 
  up :: "('a, 'm) ring_scheme => (nat => 'a) set"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
57  | 
  where "up R = {f. f \<in> UNIV -> carrier R & (EX n. bound \<zero>\<^bsub>R\<^esub> n f)}"
 | 
| 27933 | 58  | 
|
59  | 
definition UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring"
 | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
60  | 
where "UP R = (|  | 
| 27933 | 61  | 
carrier = up R,  | 
62  | 
   mult = (%p:up R. %q:up R. %n. \<Oplus>\<^bsub>R\<^esub>i \<in> {..n}. p i \<otimes>\<^bsub>R\<^esub> q (n-i)),
 | 
|
63  | 
one = (%i. if i=0 then \<one>\<^bsub>R\<^esub> else \<zero>\<^bsub>R\<^esub>),  | 
|
64  | 
zero = (%i. \<zero>\<^bsub>R\<^esub>),  | 
|
65  | 
add = (%p:up R. %q:up R. %i. p i \<oplus>\<^bsub>R\<^esub> q i),  | 
|
66  | 
smult = (%a:carrier R. %p:up R. %i. a \<otimes>\<^bsub>R\<^esub> p i),  | 
|
67  | 
monom = (%a:carrier R. %n i. if i=n then a else \<zero>\<^bsub>R\<^esub>),  | 
|
68  | 
coeff = (%p:up R. %n. p n) |)"  | 
|
| 13940 | 69  | 
|
70  | 
text {*
 | 
|
71  | 
  Properties of the set of polynomials @{term up}.
 | 
|
72  | 
*}  | 
|
73  | 
||
74  | 
lemma mem_upI [intro]:  | 
|
75  | 
"[| !!n. f n \<in> carrier R; EX n. bound (zero R) n f |] ==> f \<in> up R"  | 
|
76  | 
by (simp add: up_def Pi_def)  | 
|
77  | 
||
78  | 
lemma mem_upD [dest]:  | 
|
79  | 
"f \<in> up R ==> f n \<in> carrier R"  | 
|
80  | 
by (simp add: up_def Pi_def)  | 
|
81  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
82  | 
context ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
83  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
84  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
85  | 
lemma bound_upD [dest]: "f \<in> up R ==> EX n. bound \<zero> n f" by (simp add: up_def)  | 
| 13940 | 86  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
87  | 
lemma up_one_closed: "(%n. if n = 0 then \<one> else \<zero>) \<in> up R" using up_def by force  | 
| 13940 | 88  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
89  | 
lemma up_smult_closed: "[| a \<in> carrier R; p \<in> up R |] ==> (%i. a \<otimes> p i) \<in> up R" by force  | 
| 13940 | 90  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
91  | 
lemma up_add_closed:  | 
| 13940 | 92  | 
"[| p \<in> up R; q \<in> up R |] ==> (%i. p i \<oplus> q i) \<in> up R"  | 
93  | 
proof  | 
|
94  | 
fix n  | 
|
95  | 
assume "p \<in> up R" and "q \<in> up R"  | 
|
96  | 
then show "p n \<oplus> q n \<in> carrier R"  | 
|
97  | 
by auto  | 
|
98  | 
next  | 
|
99  | 
assume UP: "p \<in> up R" "q \<in> up R"  | 
|
100  | 
show "EX n. bound \<zero> n (%i. p i \<oplus> q i)"  | 
|
101  | 
proof -  | 
|
102  | 
from UP obtain n where boundn: "bound \<zero> n p" by fast  | 
|
103  | 
from UP obtain m where boundm: "bound \<zero> m q" by fast  | 
|
104  | 
have "bound \<zero> (max n m) (%i. p i \<oplus> q i)"  | 
|
105  | 
proof  | 
|
106  | 
fix i  | 
|
107  | 
assume "max n m < i"  | 
|
108  | 
with boundn and boundm and UP show "p i \<oplus> q i = \<zero>" by fastsimp  | 
|
109  | 
qed  | 
|
110  | 
then show ?thesis ..  | 
|
111  | 
qed  | 
|
112  | 
qed  | 
|
113  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
114  | 
lemma up_a_inv_closed:  | 
| 13940 | 115  | 
"p \<in> up R ==> (%i. \<ominus> (p i)) \<in> up R"  | 
116  | 
proof  | 
|
117  | 
assume R: "p \<in> up R"  | 
|
118  | 
then obtain n where "bound \<zero> n p" by auto  | 
|
119  | 
then have "bound \<zero> n (%i. \<ominus> p i)" by auto  | 
|
120  | 
then show "EX n. bound \<zero> n (%i. \<ominus> p i)" by auto  | 
|
121  | 
qed auto  | 
|
122  | 
||
| 27933 | 123  | 
lemma up_minus_closed:  | 
124  | 
"[| p \<in> up R; q \<in> up R |] ==> (%i. p i \<ominus> q i) \<in> up R"  | 
|
125  | 
using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed a_minus_def [of _ R]  | 
|
126  | 
by auto  | 
|
127  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
128  | 
lemma up_mult_closed:  | 
| 13940 | 129  | 
"[| p \<in> up R; q \<in> up R |] ==>  | 
| 14666 | 130  | 
  (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> up R"
 | 
| 13940 | 131  | 
proof  | 
132  | 
fix n  | 
|
133  | 
assume "p \<in> up R" "q \<in> up R"  | 
|
| 14666 | 134  | 
  then show "(\<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> carrier R"
 | 
| 13940 | 135  | 
by (simp add: mem_upD funcsetI)  | 
136  | 
next  | 
|
137  | 
assume UP: "p \<in> up R" "q \<in> up R"  | 
|
| 14666 | 138  | 
  show "EX n. bound \<zero> n (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i))"
 | 
| 13940 | 139  | 
proof -  | 
140  | 
from UP obtain n where boundn: "bound \<zero> n p" by fast  | 
|
141  | 
from UP obtain m where boundm: "bound \<zero> m q" by fast  | 
|
| 14666 | 142  | 
    have "bound \<zero> (n + m) (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n - i))"
 | 
| 13940 | 143  | 
proof  | 
| 14666 | 144  | 
fix k assume bound: "n + m < k"  | 
| 13940 | 145  | 
      {
 | 
| 14666 | 146  | 
fix i  | 
147  | 
have "p i \<otimes> q (k-i) = \<zero>"  | 
|
148  | 
proof (cases "n < i")  | 
|
149  | 
case True  | 
|
150  | 
with boundn have "p i = \<zero>" by auto  | 
|
| 13940 | 151  | 
moreover from UP have "q (k-i) \<in> carrier R" by auto  | 
| 14666 | 152  | 
ultimately show ?thesis by simp  | 
153  | 
next  | 
|
154  | 
case False  | 
|
155  | 
with bound have "m < k-i" by arith  | 
|
156  | 
with boundm have "q (k-i) = \<zero>" by auto  | 
|
157  | 
moreover from UP have "p i \<in> carrier R" by auto  | 
|
158  | 
ultimately show ?thesis by simp  | 
|
159  | 
qed  | 
|
| 13940 | 160  | 
}  | 
| 14666 | 161  | 
      then show "(\<Oplus>i \<in> {..k}. p i \<otimes> q (k-i)) = \<zero>"
 | 
162  | 
by (simp add: Pi_def)  | 
|
| 13940 | 163  | 
qed  | 
164  | 
then show ?thesis by fast  | 
|
165  | 
qed  | 
|
166  | 
qed  | 
|
167  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
168  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
169  | 
|
| 14666 | 170  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
171  | 
subsection {* Effect of Operations on Coefficients *}
 | 
| 13940 | 172  | 
|
| 19783 | 173  | 
locale UP =  | 
174  | 
fixes R (structure) and P (structure)  | 
|
| 13940 | 175  | 
defines P_def: "P == UP R"  | 
176  | 
||
| 29240 | 177  | 
locale UP_ring = UP + R: ring R  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
178  | 
|
| 29240 | 179  | 
locale UP_cring = UP + R: cring R  | 
| 13940 | 180  | 
|
| 29237 | 181  | 
sublocale UP_cring < UP_ring  | 
| 29240 | 182  | 
by intro_locales [1] (rule P_def)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
183  | 
|
| 29240 | 184  | 
locale UP_domain = UP + R: "domain" R  | 
| 13940 | 185  | 
|
| 29237 | 186  | 
sublocale UP_domain < UP_cring  | 
| 29240 | 187  | 
by intro_locales [1] (rule P_def)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
188  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
189  | 
context UP  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
190  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
191  | 
|
| 
30363
 
9b8d9b6ef803
proper local context for text with antiquotations;
 
wenzelm 
parents: 
29246 
diff
changeset
 | 
192  | 
text {*Temporarily declare @{thm P_def} as simp rule.*}
 | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
193  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
194  | 
declare P_def [simp]  | 
| 13940 | 195  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
196  | 
lemma up_eqI:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
197  | 
assumes prem: "!!n. coeff P p n = coeff P q n" and R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
198  | 
shows "p = q"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
199  | 
proof  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
200  | 
fix x  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
201  | 
from prem and R show "p x = q x" by (simp add: UP_def)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
202  | 
qed  | 
| 13940 | 203  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
204  | 
lemma coeff_closed [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
205  | 
"p \<in> carrier P ==> coeff P p n \<in> carrier R" by (auto simp add: UP_def)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
206  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
207  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
208  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
209  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
210  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
211  | 
|
| 27933 | 212  | 
(* Theorems generalised from commutative rings to rings by Jesus Aransay. *)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
213  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
214  | 
lemma coeff_monom [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
215  | 
"a \<in> carrier R ==> coeff P (monom P a m) n = (if m=n then a else \<zero>)"  | 
| 13940 | 216  | 
proof -  | 
217  | 
assume R: "a \<in> carrier R"  | 
|
218  | 
then have "(%n. if n = m then a else \<zero>) \<in> up R"  | 
|
219  | 
using up_def by force  | 
|
220  | 
with R show ?thesis by (simp add: UP_def)  | 
|
221  | 
qed  | 
|
222  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
223  | 
lemma coeff_zero [simp]: "coeff P \<zero>\<^bsub>P\<^esub> n = \<zero>" by (auto simp add: UP_def)  | 
| 13940 | 224  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
225  | 
lemma coeff_one [simp]: "coeff P \<one>\<^bsub>P\<^esub> n = (if n=0 then \<one> else \<zero>)"  | 
| 13940 | 226  | 
using up_one_closed by (simp add: UP_def)  | 
227  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
228  | 
lemma coeff_smult [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
229  | 
"[| a \<in> carrier R; p \<in> carrier P |] ==> coeff P (a \<odot>\<^bsub>P\<^esub> p) n = a \<otimes> coeff P p n"  | 
| 13940 | 230  | 
by (simp add: UP_def up_smult_closed)  | 
231  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
232  | 
lemma coeff_add [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
233  | 
"[| p \<in> carrier P; q \<in> carrier P |] ==> coeff P (p \<oplus>\<^bsub>P\<^esub> q) n = coeff P p n \<oplus> coeff P q n"  | 
| 13940 | 234  | 
by (simp add: UP_def up_add_closed)  | 
235  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
236  | 
lemma coeff_mult [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
237  | 
  "[| p \<in> carrier P; q \<in> carrier P |] ==> coeff P (p \<otimes>\<^bsub>P\<^esub> q) n = (\<Oplus>i \<in> {..n}. coeff P p i \<otimes> coeff P q (n-i))"
 | 
| 13940 | 238  | 
by (simp add: UP_def up_mult_closed)  | 
239  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
240  | 
end  | 
| 14666 | 241  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
242  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
243  | 
subsection {* Polynomials Form a Ring. *}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
244  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
245  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
246  | 
begin  | 
| 13940 | 247  | 
|
| 14666 | 248  | 
text {* Operations are closed over @{term P}. *}
 | 
| 13940 | 249  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
250  | 
lemma UP_mult_closed [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
251  | 
"[| p \<in> carrier P; q \<in> carrier P |] ==> p \<otimes>\<^bsub>P\<^esub> q \<in> carrier P" by (simp add: UP_def up_mult_closed)  | 
| 13940 | 252  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
253  | 
lemma UP_one_closed [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
254  | 
"\<one>\<^bsub>P\<^esub> \<in> carrier P" by (simp add: UP_def up_one_closed)  | 
| 13940 | 255  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
256  | 
lemma UP_zero_closed [intro, simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
257  | 
"\<zero>\<^bsub>P\<^esub> \<in> carrier P" by (auto simp add: UP_def)  | 
| 13940 | 258  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
259  | 
lemma UP_a_closed [intro, simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
260  | 
"[| p \<in> carrier P; q \<in> carrier P |] ==> p \<oplus>\<^bsub>P\<^esub> q \<in> carrier P" by (simp add: UP_def up_add_closed)  | 
| 13940 | 261  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
262  | 
lemma monom_closed [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
263  | 
"a \<in> carrier R ==> monom P a n \<in> carrier P" by (auto simp add: UP_def up_def Pi_def)  | 
| 13940 | 264  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
265  | 
lemma UP_smult_closed [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
266  | 
"[| a \<in> carrier R; p \<in> carrier P |] ==> a \<odot>\<^bsub>P\<^esub> p \<in> carrier P" by (simp add: UP_def up_smult_closed)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
267  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
268  | 
end  | 
| 13940 | 269  | 
|
270  | 
declare (in UP) P_def [simp del]  | 
|
271  | 
||
272  | 
text {* Algebraic ring properties *}
 | 
|
273  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
274  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
275  | 
begin  | 
| 13940 | 276  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
277  | 
lemma UP_a_assoc:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
278  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
279  | 
shows "(p \<oplus>\<^bsub>P\<^esub> q) \<oplus>\<^bsub>P\<^esub> r = p \<oplus>\<^bsub>P\<^esub> (q \<oplus>\<^bsub>P\<^esub> r)" by (rule up_eqI, simp add: a_assoc R, simp_all add: R)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
280  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
281  | 
lemma UP_l_zero [simp]:  | 
| 13940 | 282  | 
assumes R: "p \<in> carrier P"  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
283  | 
shows "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> p = p" by (rule up_eqI, simp_all add: R)  | 
| 13940 | 284  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
285  | 
lemma UP_l_neg_ex:  | 
| 13940 | 286  | 
assumes R: "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
287  | 
shows "EX q : carrier P. q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 288  | 
proof -  | 
289  | 
let ?q = "%i. \<ominus> (p i)"  | 
|
290  | 
from R have closed: "?q \<in> carrier P"  | 
|
291  | 
by (simp add: UP_def P_def up_a_inv_closed)  | 
|
292  | 
from R have coeff: "!!n. coeff P ?q n = \<ominus> (coeff P p n)"  | 
|
293  | 
by (simp add: UP_def P_def up_a_inv_closed)  | 
|
294  | 
show ?thesis  | 
|
295  | 
proof  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
296  | 
show "?q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 297  | 
by (auto intro!: up_eqI simp add: R closed coeff R.l_neg)  | 
298  | 
qed (rule closed)  | 
|
299  | 
qed  | 
|
300  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
301  | 
lemma UP_a_comm:  | 
| 13940 | 302  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
303  | 
shows "p \<oplus>\<^bsub>P\<^esub> q = q \<oplus>\<^bsub>P\<^esub> p" by (rule up_eqI, simp add: a_comm R, simp_all add: R)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
304  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
305  | 
lemma UP_m_assoc:  | 
| 13940 | 306  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
307  | 
shows "(p \<otimes>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r = p \<otimes>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)"  | 
| 13940 | 308  | 
proof (rule up_eqI)  | 
309  | 
fix n  | 
|
310  | 
  {
 | 
|
311  | 
fix k and a b c :: "nat=>'a"  | 
|
312  | 
assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"  | 
|
313  | 
"c \<in> UNIV -> carrier R"  | 
|
314  | 
then have "k <= n ==>  | 
|
| 14666 | 315  | 
      (\<Oplus>j \<in> {..k}. (\<Oplus>i \<in> {..j}. a i \<otimes> b (j-i)) \<otimes> c (n-j)) =
 | 
316  | 
      (\<Oplus>j \<in> {..k}. a j \<otimes> (\<Oplus>i \<in> {..k-j}. b i \<otimes> c (n-j-i)))"
 | 
|
| 19582 | 317  | 
(is "_ \<Longrightarrow> ?eq k")  | 
| 13940 | 318  | 
proof (induct k)  | 
319  | 
case 0 then show ?case by (simp add: Pi_def m_assoc)  | 
|
320  | 
next  | 
|
321  | 
case (Suc k)  | 
|
322  | 
then have "k <= n" by arith  | 
|
| 23350 | 323  | 
from this R have "?eq k" by (rule Suc)  | 
| 13940 | 324  | 
with R show ?case  | 
| 14666 | 325  | 
by (simp cong: finsum_cong  | 
| 13940 | 326  | 
add: Suc_diff_le Pi_def l_distr r_distr m_assoc)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
327  | 
(simp cong: finsum_cong add: Pi_def a_ac finsum_ldistr m_assoc)  | 
| 13940 | 328  | 
qed  | 
329  | 
}  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
330  | 
with R show "coeff P ((p \<otimes>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r) n = coeff P (p \<otimes>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)) n"  | 
| 13940 | 331  | 
by (simp add: Pi_def)  | 
332  | 
qed (simp_all add: R)  | 
|
333  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
334  | 
lemma UP_r_one [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
335  | 
assumes R: "p \<in> carrier P" shows "p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub> = p"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
336  | 
proof (rule up_eqI)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
337  | 
fix n  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
338  | 
show "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) n = coeff P p n"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
339  | 
proof (cases n)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
340  | 
case 0  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
341  | 
    {
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
342  | 
with R show ?thesis by simp  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
343  | 
}  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
344  | 
next  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
345  | 
case Suc  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
346  | 
    {
 | 
| 27933 | 347  | 
(*JE: in the locale UP_cring the proof was solved only with "by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)", but I did not get it to work here*)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
348  | 
fix nn assume Succ: "n = Suc nn"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
349  | 
have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
350  | 
proof -  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
351  | 
        have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
352  | 
        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
353  | 
using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
354  | 
also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
355  | 
proof -  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
356  | 
          have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
357  | 
            using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
358  | 
unfolding Pi_def by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
359  | 
also have "\<dots> = \<zero>" by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
360  | 
finally show ?thesis using r_zero R by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
361  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
362  | 
also have "\<dots> = coeff P p (Suc nn)" using R by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
363  | 
finally show ?thesis by simp  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
364  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
365  | 
then show ?thesis using Succ by simp  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
366  | 
}  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
367  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
368  | 
qed (simp_all add: R)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
369  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
370  | 
lemma UP_l_one [simp]:  | 
| 13940 | 371  | 
assumes R: "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
372  | 
shows "\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p = p"  | 
| 13940 | 373  | 
proof (rule up_eqI)  | 
374  | 
fix n  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
375  | 
show "coeff P (\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p) n = coeff P p n"  | 
| 13940 | 376  | 
proof (cases n)  | 
377  | 
case 0 with R show ?thesis by simp  | 
|
378  | 
next  | 
|
379  | 
case Suc with R show ?thesis  | 
|
380  | 
by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)  | 
|
381  | 
qed  | 
|
382  | 
qed (simp_all add: R)  | 
|
383  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
384  | 
lemma UP_l_distr:  | 
| 13940 | 385  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
386  | 
shows "(p \<oplus>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r = (p \<otimes>\<^bsub>P\<^esub> r) \<oplus>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)"  | 
| 13940 | 387  | 
by (rule up_eqI) (simp add: l_distr R Pi_def, simp_all add: R)  | 
388  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
389  | 
lemma UP_r_distr:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
390  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
391  | 
shows "r \<otimes>\<^bsub>P\<^esub> (p \<oplus>\<^bsub>P\<^esub> q) = (r \<otimes>\<^bsub>P\<^esub> p) \<oplus>\<^bsub>P\<^esub> (r \<otimes>\<^bsub>P\<^esub> q)"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
392  | 
by (rule up_eqI) (simp add: r_distr R Pi_def, simp_all add: R)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
393  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
394  | 
theorem UP_ring: "ring P"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
395  | 
by (auto intro!: ringI abelian_groupI monoidI UP_a_assoc)  | 
| 27933 | 396  | 
(auto intro: UP_a_comm UP_l_neg_ex UP_m_assoc UP_l_distr UP_r_distr)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
397  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
398  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
399  | 
|
| 27933 | 400  | 
|
401  | 
subsection {* Polynomials Form a Commutative Ring. *}
 | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
402  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
403  | 
context UP_cring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
404  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
405  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
406  | 
lemma UP_m_comm:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
407  | 
assumes R1: "p \<in> carrier P" and R2: "q \<in> carrier P" shows "p \<otimes>\<^bsub>P\<^esub> q = q \<otimes>\<^bsub>P\<^esub> p"  | 
| 13940 | 408  | 
proof (rule up_eqI)  | 
| 14666 | 409  | 
fix n  | 
| 13940 | 410  | 
  {
 | 
411  | 
fix k and a b :: "nat=>'a"  | 
|
412  | 
assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"  | 
|
| 14666 | 413  | 
then have "k <= n ==>  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
414  | 
      (\<Oplus>i \<in> {..k}. a i \<otimes> b (n-i)) = (\<Oplus>i \<in> {..k}. a (k-i) \<otimes> b (i+n-k))"
 | 
| 19582 | 415  | 
(is "_ \<Longrightarrow> ?eq k")  | 
| 13940 | 416  | 
proof (induct k)  | 
417  | 
case 0 then show ?case by (simp add: Pi_def)  | 
|
418  | 
next  | 
|
419  | 
case (Suc k) then show ?case  | 
|
| 15944 | 420  | 
by (subst (2) finsum_Suc2) (simp add: Pi_def a_comm)+  | 
| 13940 | 421  | 
qed  | 
422  | 
}  | 
|
423  | 
note l = this  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
424  | 
from R1 R2 show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) n = coeff P (q \<otimes>\<^bsub>P\<^esub> p) n"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
425  | 
unfolding coeff_mult [OF R1 R2, of n]  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
426  | 
unfolding coeff_mult [OF R2 R1, of n]  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
427  | 
using l [of "(\<lambda>i. coeff P p i)" "(\<lambda>i. coeff P q i)" "n"] by (simp add: Pi_def m_comm)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
428  | 
qed (simp_all add: R1 R2)  | 
| 13940 | 429  | 
|
| 35849 | 430  | 
|
431  | 
subsection {*Polynomials over a commutative ring for a commutative ring*}
 | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
432  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
433  | 
theorem UP_cring:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
434  | 
"cring P" using UP_ring unfolding cring_def by (auto intro!: comm_monoidI UP_m_assoc UP_m_comm)  | 
| 13940 | 435  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
436  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
437  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
438  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
439  | 
begin  | 
| 
14399
 
dc677b35e54f
New lemmas about inversion of restricted functions.
 
ballarin 
parents: 
13975 
diff
changeset
 | 
440  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
441  | 
lemma UP_a_inv_closed [intro, simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
442  | 
"p \<in> carrier P ==> \<ominus>\<^bsub>P\<^esub> p \<in> carrier P"  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
443  | 
by (rule abelian_group.a_inv_closed [OF ring.is_abelian_group [OF UP_ring]])  | 
| 13940 | 444  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
445  | 
lemma coeff_a_inv [simp]:  | 
| 13940 | 446  | 
assumes R: "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
447  | 
shows "coeff P (\<ominus>\<^bsub>P\<^esub> p) n = \<ominus> (coeff P p n)"  | 
| 13940 | 448  | 
proof -  | 
449  | 
from R coeff_closed UP_a_inv_closed have  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
450  | 
"coeff P (\<ominus>\<^bsub>P\<^esub> p) n = \<ominus> coeff P p n \<oplus> (coeff P p n \<oplus> coeff P (\<ominus>\<^bsub>P\<^esub> p) n)"  | 
| 13940 | 451  | 
by algebra  | 
452  | 
also from R have "... = \<ominus> (coeff P p n)"  | 
|
453  | 
by (simp del: coeff_add add: coeff_add [THEN sym]  | 
|
| 
14399
 
dc677b35e54f
New lemmas about inversion of restricted functions.
 
ballarin 
parents: 
13975 
diff
changeset
 | 
454  | 
abelian_group.r_neg [OF ring.is_abelian_group [OF UP_ring]])  | 
| 13940 | 455  | 
finally show ?thesis .  | 
456  | 
qed  | 
|
457  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
458  | 
end  | 
| 13940 | 459  | 
|
| 29240 | 460  | 
sublocale UP_ring < P: ring P using UP_ring .  | 
461  | 
sublocale UP_cring < P: cring P using UP_cring .  | 
|
| 13940 | 462  | 
|
| 14666 | 463  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
464  | 
subsection {* Polynomials Form an Algebra *}
 | 
| 13940 | 465  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
466  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
467  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
468  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
469  | 
lemma UP_smult_l_distr:  | 
| 13940 | 470  | 
"[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
471  | 
(a \<oplus> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p \<oplus>\<^bsub>P\<^esub> b \<odot>\<^bsub>P\<^esub> p"  | 
| 13940 | 472  | 
by (rule up_eqI) (simp_all add: R.l_distr)  | 
473  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
474  | 
lemma UP_smult_r_distr:  | 
| 13940 | 475  | 
"[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
476  | 
a \<odot>\<^bsub>P\<^esub> (p \<oplus>\<^bsub>P\<^esub> q) = a \<odot>\<^bsub>P\<^esub> p \<oplus>\<^bsub>P\<^esub> a \<odot>\<^bsub>P\<^esub> q"  | 
| 13940 | 477  | 
by (rule up_eqI) (simp_all add: R.r_distr)  | 
478  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
479  | 
lemma UP_smult_assoc1:  | 
| 13940 | 480  | 
"[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
481  | 
(a \<otimes> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> p)"  | 
| 13940 | 482  | 
by (rule up_eqI) (simp_all add: R.m_assoc)  | 
483  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
484  | 
lemma UP_smult_zero [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
485  | 
"p \<in> carrier P ==> \<zero> \<odot>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
486  | 
by (rule up_eqI) simp_all  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
487  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
488  | 
lemma UP_smult_one [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
489  | 
"p \<in> carrier P ==> \<one> \<odot>\<^bsub>P\<^esub> p = p"  | 
| 13940 | 490  | 
by (rule up_eqI) simp_all  | 
491  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
492  | 
lemma UP_smult_assoc2:  | 
| 13940 | 493  | 
"[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
494  | 
(a \<odot>\<^bsub>P\<^esub> p) \<otimes>\<^bsub>P\<^esub> q = a \<odot>\<^bsub>P\<^esub> (p \<otimes>\<^bsub>P\<^esub> q)"  | 
| 13940 | 495  | 
by (rule up_eqI) (simp_all add: R.finsum_rdistr R.m_assoc Pi_def)  | 
496  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
497  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
498  | 
|
| 13940 | 499  | 
text {*
 | 
| 17094 | 500  | 
  Interpretation of lemmas from @{term algebra}.
 | 
| 13940 | 501  | 
*}  | 
502  | 
||
503  | 
lemma (in cring) cring:  | 
|
| 28823 | 504  | 
"cring R" ..  | 
| 13940 | 505  | 
|
506  | 
lemma (in UP_cring) UP_algebra:  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
507  | 
"algebra R P" by (auto intro!: algebraI R.cring UP_cring UP_smult_l_distr UP_smult_r_distr  | 
| 13940 | 508  | 
UP_smult_assoc1 UP_smult_assoc2)  | 
509  | 
||
| 29237 | 510  | 
sublocale UP_cring < algebra R P using UP_algebra .  | 
| 13940 | 511  | 
|
512  | 
||
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
513  | 
subsection {* Further Lemmas Involving Monomials *}
 | 
| 13940 | 514  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
515  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
516  | 
begin  | 
| 13940 | 517  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
518  | 
lemma monom_zero [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
519  | 
"monom P \<zero> n = \<zero>\<^bsub>P\<^esub>" by (simp add: UP_def P_def)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
520  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
521  | 
lemma monom_mult_is_smult:  | 
| 13940 | 522  | 
assumes R: "a \<in> carrier R" "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
523  | 
shows "monom P a 0 \<otimes>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p"  | 
| 13940 | 524  | 
proof (rule up_eqI)  | 
525  | 
fix n  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
526  | 
show "coeff P (monom P a 0 \<otimes>\<^bsub>P\<^esub> p) n = coeff P (a \<odot>\<^bsub>P\<^esub> p) n"  | 
| 13940 | 527  | 
proof (cases n)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
528  | 
case 0 with R show ?thesis by simp  | 
| 13940 | 529  | 
next  | 
530  | 
case Suc with R show ?thesis  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
531  | 
using R.finsum_Suc2 by (simp del: R.finsum_Suc add: R.r_null Pi_def)  | 
| 13940 | 532  | 
qed  | 
533  | 
qed (simp_all add: R)  | 
|
534  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
535  | 
lemma monom_one [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
536  | 
"monom P \<one> 0 = \<one>\<^bsub>P\<^esub>"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
537  | 
by (rule up_eqI) simp_all  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
538  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
539  | 
lemma monom_add [simp]:  | 
| 13940 | 540  | 
"[| a \<in> carrier R; b \<in> carrier R |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
541  | 
monom P (a \<oplus> b) n = monom P a n \<oplus>\<^bsub>P\<^esub> monom P b n"  | 
| 13940 | 542  | 
by (rule up_eqI) simp_all  | 
543  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
544  | 
lemma monom_one_Suc:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
545  | 
"monom P \<one> (Suc n) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"  | 
| 13940 | 546  | 
proof (rule up_eqI)  | 
547  | 
fix k  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
548  | 
show "coeff P (monom P \<one> (Suc n)) k = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"  | 
| 13940 | 549  | 
proof (cases "k = Suc n")  | 
550  | 
case True show ?thesis  | 
|
551  | 
proof -  | 
|
| 26934 | 552  | 
fix m  | 
| 14666 | 553  | 
from True have less_add_diff:  | 
554  | 
"!!i. [| n < i; i <= n + m |] ==> n + m - i < m" by arith  | 
|
| 13940 | 555  | 
from True have "coeff P (monom P \<one> (Suc n)) k = \<one>" by simp  | 
556  | 
also from True  | 
|
| 15045 | 557  | 
      have "... = (\<Oplus>i \<in> {..<n} \<union> {n}. coeff P (monom P \<one> n) i \<otimes>
 | 
| 14666 | 558  | 
coeff P (monom P \<one> 1) (k - i))"  | 
| 17094 | 559  | 
by (simp cong: R.finsum_cong add: Pi_def)  | 
| 14666 | 560  | 
      also have "... = (\<Oplus>i \<in>  {..n}. coeff P (monom P \<one> n) i \<otimes>
 | 
561  | 
coeff P (monom P \<one> 1) (k - i))"  | 
|
562  | 
by (simp only: ivl_disj_un_singleton)  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
563  | 
also from True  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
564  | 
      have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. coeff P (monom P \<one> n) i \<otimes>
 | 
| 14666 | 565  | 
coeff P (monom P \<one> 1) (k - i))"  | 
| 17094 | 566  | 
by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one  | 
| 14666 | 567  | 
order_less_imp_not_eq Pi_def)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
568  | 
also from True have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"  | 
| 14666 | 569  | 
by (simp add: ivl_disj_un_one)  | 
| 13940 | 570  | 
finally show ?thesis .  | 
571  | 
qed  | 
|
572  | 
next  | 
|
573  | 
case False  | 
|
574  | 
note neq = False  | 
|
575  | 
let ?s =  | 
|
| 14666 | 576  | 
"\<lambda>i. (if n = i then \<one> else \<zero>) \<otimes> (if Suc 0 = k - i then \<one> else \<zero>)"  | 
| 13940 | 577  | 
from neq have "coeff P (monom P \<one> (Suc n)) k = \<zero>" by simp  | 
| 14666 | 578  | 
    also have "... = (\<Oplus>i \<in> {..k}. ?s i)"
 | 
| 13940 | 579  | 
proof -  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
580  | 
      have f1: "(\<Oplus>i \<in> {..<n}. ?s i) = \<zero>"
 | 
| 17094 | 581  | 
by (simp cong: R.finsum_cong add: Pi_def)  | 
| 14666 | 582  | 
      from neq have f2: "(\<Oplus>i \<in> {n}. ?s i) = \<zero>"
 | 
| 
20432
 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 
webertj 
parents: 
20318 
diff
changeset
 | 
583  | 
by (simp cong: R.finsum_cong add: Pi_def) arith  | 
| 15045 | 584  | 
      have f3: "n < k ==> (\<Oplus>i \<in> {n<..k}. ?s i) = \<zero>"
 | 
| 17094 | 585  | 
by (simp cong: R.finsum_cong add: order_less_imp_not_eq Pi_def)  | 
| 13940 | 586  | 
show ?thesis  | 
587  | 
proof (cases "k < n")  | 
|
| 17094 | 588  | 
case True then show ?thesis by (simp cong: R.finsum_cong add: Pi_def)  | 
| 13940 | 589  | 
next  | 
| 14666 | 590  | 
case False then have n_le_k: "n <= k" by arith  | 
591  | 
show ?thesis  | 
|
592  | 
proof (cases "n = k")  | 
|
593  | 
case True  | 
|
| 15045 | 594  | 
          then have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
 | 
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
595  | 
by (simp cong: R.finsum_cong add: Pi_def)  | 
| 14666 | 596  | 
          also from True have "... = (\<Oplus>i \<in> {..k}. ?s i)"
 | 
597  | 
by (simp only: ivl_disj_un_singleton)  | 
|
598  | 
finally show ?thesis .  | 
|
599  | 
next  | 
|
600  | 
case False with n_le_k have n_less_k: "n < k" by arith  | 
|
| 15045 | 601  | 
          with neq have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
 | 
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
602  | 
by (simp add: R.finsum_Un_disjoint f1 f2 Pi_def del: Un_insert_right)  | 
| 14666 | 603  | 
          also have "... = (\<Oplus>i \<in> {..n}. ?s i)"
 | 
604  | 
by (simp only: ivl_disj_un_singleton)  | 
|
| 15045 | 605  | 
          also from n_less_k neq have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. ?s i)"
 | 
| 17094 | 606  | 
by (simp add: R.finsum_Un_disjoint f3 ivl_disj_int_one Pi_def)  | 
| 14666 | 607  | 
          also from n_less_k have "... = (\<Oplus>i \<in> {..k}. ?s i)"
 | 
608  | 
by (simp only: ivl_disj_un_one)  | 
|
609  | 
finally show ?thesis .  | 
|
610  | 
qed  | 
|
| 13940 | 611  | 
qed  | 
612  | 
qed  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
613  | 
also have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k" by simp  | 
| 13940 | 614  | 
finally show ?thesis .  | 
615  | 
qed  | 
|
616  | 
qed (simp_all)  | 
|
617  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
618  | 
lemma monom_one_Suc2:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
619  | 
"monom P \<one> (Suc n) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> n"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
620  | 
proof (induct n)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
621  | 
case 0 show ?case by simp  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
622  | 
next  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
623  | 
case Suc  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
624  | 
  {
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
625  | 
fix k:: nat  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
626  | 
assume hypo: "monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
627  | 
then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
628  | 
proof -  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
629  | 
have lhs: "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
630  | 
unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
631  | 
note cl = monom_closed [OF R.one_closed, of 1]  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
632  | 
note clk = monom_closed [OF R.one_closed, of k]  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
633  | 
have rhs: "monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
634  | 
unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] ..  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
635  | 
from lhs rhs show ?thesis by simp  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
636  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
637  | 
}  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
638  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
639  | 
|
| 
30363
 
9b8d9b6ef803
proper local context for text with antiquotations;
 
wenzelm 
parents: 
29246 
diff
changeset
 | 
640  | 
text{*The following corollary follows from lemmas @{thm "monom_one_Suc"} 
 | 
| 
 
9b8d9b6ef803
proper local context for text with antiquotations;
 
wenzelm 
parents: 
29246 
diff
changeset
 | 
641  | 
  and @{thm "monom_one_Suc2"}, and is trivial in @{term UP_cring}*}
 | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
642  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
643  | 
corollary monom_one_comm: shows "monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
644  | 
unfolding monom_one_Suc [symmetric] monom_one_Suc2 [symmetric] ..  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
645  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
646  | 
lemma monom_mult_smult:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
647  | 
"[| a \<in> carrier R; b \<in> carrier R |] ==> monom P (a \<otimes> b) n = a \<odot>\<^bsub>P\<^esub> monom P b n"  | 
| 13940 | 648  | 
by (rule up_eqI) simp_all  | 
649  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
650  | 
lemma monom_one_mult:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
651  | 
"monom P \<one> (n + m) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m"  | 
| 13940 | 652  | 
proof (induct n)  | 
653  | 
case 0 show ?case by simp  | 
|
654  | 
next  | 
|
655  | 
case Suc then show ?case  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
656  | 
unfolding add_Suc unfolding monom_one_Suc unfolding Suc.hyps  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
657  | 
using m_assoc monom_one_comm [of m] by simp  | 
| 13940 | 658  | 
qed  | 
659  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
660  | 
lemma monom_one_mult_comm: "monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m = monom P \<one> m \<otimes>\<^bsub>P\<^esub> monom P \<one> n"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
661  | 
unfolding monom_one_mult [symmetric] by (rule up_eqI) simp_all  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
662  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
663  | 
lemma monom_mult [simp]:  | 
| 27933 | 664  | 
assumes a_in_R: "a \<in> carrier R" and b_in_R: "b \<in> carrier R"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
665  | 
shows "monom P (a \<otimes> b) (n + m) = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m"  | 
| 27933 | 666  | 
proof (rule up_eqI)  | 
667  | 
fix k  | 
|
668  | 
show "coeff P (monom P (a \<otimes> b) (n + m)) k = coeff P (monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m) k"  | 
|
669  | 
proof (cases "n + m = k")  | 
|
670  | 
case True  | 
|
671  | 
    {
 | 
|
672  | 
show ?thesis  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
673  | 
unfolding True [symmetric]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
674  | 
coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
675  | 
coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
676  | 
        using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
677  | 
"(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
678  | 
a_in_R b_in_R  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
679  | 
unfolding simp_implies_def  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
680  | 
        using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
681  | 
unfolding Pi_def by auto  | 
| 27933 | 682  | 
}  | 
683  | 
next  | 
|
684  | 
case False  | 
|
685  | 
    {
 | 
|
686  | 
show ?thesis  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
687  | 
unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
688  | 
unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
689  | 
unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
690  | 
        using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
691  | 
unfolding Pi_def simp_implies_def using a_in_R b_in_R by force  | 
| 27933 | 692  | 
}  | 
693  | 
qed  | 
|
694  | 
qed (simp_all add: a_in_R b_in_R)  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
695  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
696  | 
lemma monom_a_inv [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
697  | 
"a \<in> carrier R ==> monom P (\<ominus> a) n = \<ominus>\<^bsub>P\<^esub> monom P a n"  | 
| 13940 | 698  | 
by (rule up_eqI) simp_all  | 
699  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
700  | 
lemma monom_inj:  | 
| 13940 | 701  | 
"inj_on (%a. monom P a n) (carrier R)"  | 
702  | 
proof (rule inj_onI)  | 
|
703  | 
fix x y  | 
|
704  | 
assume R: "x \<in> carrier R" "y \<in> carrier R" and eq: "monom P x n = monom P y n"  | 
|
705  | 
then have "coeff P (monom P x n) n = coeff P (monom P y n) n" by simp  | 
|
706  | 
with R show "x = y" by simp  | 
|
707  | 
qed  | 
|
708  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
709  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
710  | 
|
| 17094 | 711  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
712  | 
subsection {* The Degree Function *}
 | 
| 13940 | 713  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
714  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
715  | 
  deg :: "[('a, 'm) ring_scheme, nat => 'a] => nat"
 | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
716  | 
where "deg R p = (LEAST n. bound \<zero>\<^bsub>R\<^esub> n (coeff (UP R) p))"  | 
| 13940 | 717  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
718  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
719  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
720  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
721  | 
lemma deg_aboveI:  | 
| 14666 | 722  | 
"[| (!!m. n < m ==> coeff P p m = \<zero>); p \<in> carrier P |] ==> deg R p <= n"  | 
| 13940 | 723  | 
by (unfold deg_def P_def) (fast intro: Least_le)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
724  | 
|
| 13940 | 725  | 
(*  | 
726  | 
lemma coeff_bound_ex: "EX n. bound n (coeff p)"  | 
|
727  | 
proof -  | 
|
728  | 
have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)  | 
|
729  | 
then obtain n where "bound n (coeff p)" by (unfold UP_def) fast  | 
|
730  | 
then show ?thesis ..  | 
|
731  | 
qed  | 
|
| 14666 | 732  | 
|
| 13940 | 733  | 
lemma bound_coeff_obtain:  | 
734  | 
assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P"  | 
|
735  | 
proof -  | 
|
736  | 
have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)  | 
|
737  | 
then obtain n where "bound n (coeff p)" by (unfold UP_def) fast  | 
|
738  | 
with prem show P .  | 
|
739  | 
qed  | 
|
740  | 
*)  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
741  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
742  | 
lemma deg_aboveD:  | 
| 23350 | 743  | 
assumes "deg R p < m" and "p \<in> carrier P"  | 
744  | 
shows "coeff P p m = \<zero>"  | 
|
| 13940 | 745  | 
proof -  | 
| 23350 | 746  | 
from `p \<in> carrier P` obtain n where "bound \<zero> n (coeff P p)"  | 
| 13940 | 747  | 
by (auto simp add: UP_def P_def)  | 
748  | 
then have "bound \<zero> (deg R p) (coeff P p)"  | 
|
749  | 
by (auto simp: deg_def P_def dest: LeastI)  | 
|
| 23350 | 750  | 
from this and `deg R p < m` show ?thesis ..  | 
| 13940 | 751  | 
qed  | 
752  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
753  | 
lemma deg_belowI:  | 
| 13940 | 754  | 
assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"  | 
755  | 
and R: "p \<in> carrier P"  | 
|
756  | 
shows "n <= deg R p"  | 
|
| 14666 | 757  | 
-- {* Logically, this is a slightly stronger version of
 | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
758  | 
   @{thm [source] deg_aboveD} *}
 | 
| 13940 | 759  | 
proof (cases "n=0")  | 
760  | 
case True then show ?thesis by simp  | 
|
761  | 
next  | 
|
762  | 
case False then have "coeff P p n ~= \<zero>" by (rule non_zero)  | 
|
763  | 
then have "~ deg R p < n" by (fast dest: deg_aboveD intro: R)  | 
|
764  | 
then show ?thesis by arith  | 
|
765  | 
qed  | 
|
766  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
767  | 
lemma lcoeff_nonzero_deg:  | 
| 13940 | 768  | 
assumes deg: "deg R p ~= 0" and R: "p \<in> carrier P"  | 
769  | 
shows "coeff P p (deg R p) ~= \<zero>"  | 
|
770  | 
proof -  | 
|
771  | 
from R obtain m where "deg R p <= m" and m_coeff: "coeff P p m ~= \<zero>"  | 
|
772  | 
proof -  | 
|
773  | 
have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"  | 
|
774  | 
by arith  | 
|
775  | 
from deg have "deg R p - 1 < (LEAST n. bound \<zero> n (coeff P p))"  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
776  | 
by (unfold deg_def P_def) simp  | 
| 13940 | 777  | 
then have "~ bound \<zero> (deg R p - 1) (coeff P p)" by (rule not_less_Least)  | 
778  | 
then have "EX m. deg R p - 1 < m & coeff P p m ~= \<zero>"  | 
|
779  | 
by (unfold bound_def) fast  | 
|
780  | 
then have "EX m. deg R p <= m & coeff P p m ~= \<zero>" by (simp add: deg minus)  | 
|
| 23350 | 781  | 
then show ?thesis by (auto intro: that)  | 
| 13940 | 782  | 
qed  | 
783  | 
with deg_belowI R have "deg R p = m" by fastsimp  | 
|
784  | 
with m_coeff show ?thesis by simp  | 
|
785  | 
qed  | 
|
786  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
787  | 
lemma lcoeff_nonzero_nonzero:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
788  | 
assumes deg: "deg R p = 0" and nonzero: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"  | 
| 13940 | 789  | 
shows "coeff P p 0 ~= \<zero>"  | 
790  | 
proof -  | 
|
791  | 
have "EX m. coeff P p m ~= \<zero>"  | 
|
792  | 
proof (rule classical)  | 
|
793  | 
assume "~ ?thesis"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
794  | 
with R have "p = \<zero>\<^bsub>P\<^esub>" by (auto intro: up_eqI)  | 
| 13940 | 795  | 
with nonzero show ?thesis by contradiction  | 
796  | 
qed  | 
|
797  | 
then obtain m where coeff: "coeff P p m ~= \<zero>" ..  | 
|
| 23350 | 798  | 
from this and R have "m <= deg R p" by (rule deg_belowI)  | 
| 13940 | 799  | 
then have "m = 0" by (simp add: deg)  | 
800  | 
with coeff show ?thesis by simp  | 
|
801  | 
qed  | 
|
802  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
803  | 
lemma lcoeff_nonzero:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
804  | 
assumes neq: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"  | 
| 13940 | 805  | 
shows "coeff P p (deg R p) ~= \<zero>"  | 
806  | 
proof (cases "deg R p = 0")  | 
|
807  | 
case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero)  | 
|
808  | 
next  | 
|
809  | 
case False with neq R show ?thesis by (simp add: lcoeff_nonzero_deg)  | 
|
810  | 
qed  | 
|
811  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
812  | 
lemma deg_eqI:  | 
| 13940 | 813  | 
"[| !!m. n < m ==> coeff P p m = \<zero>;  | 
814  | 
!!n. n ~= 0 ==> coeff P p n ~= \<zero>; p \<in> carrier P |] ==> deg R p = n"  | 
|
| 33657 | 815  | 
by (fast intro: le_antisym deg_aboveI deg_belowI)  | 
| 13940 | 816  | 
|
| 17094 | 817  | 
text {* Degree and polynomial operations *}
 | 
| 13940 | 818  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
819  | 
lemma deg_add [simp]:  | 
| 
32436
 
10cd49e0c067
Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
 
nipkow 
parents: 
30729 
diff
changeset
 | 
820  | 
"p \<in> carrier P \<Longrightarrow> q \<in> carrier P \<Longrightarrow>  | 
| 
 
10cd49e0c067
Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
 
nipkow 
parents: 
30729 
diff
changeset
 | 
821  | 
deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"  | 
| 
 
10cd49e0c067
Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
 
nipkow 
parents: 
30729 
diff
changeset
 | 
822  | 
by(rule deg_aboveI)(simp_all add: deg_aboveD)  | 
| 13940 | 823  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
824  | 
lemma deg_monom_le:  | 
| 13940 | 825  | 
"a \<in> carrier R ==> deg R (monom P a n) <= n"  | 
826  | 
by (intro deg_aboveI) simp_all  | 
|
827  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
828  | 
lemma deg_monom [simp]:  | 
| 13940 | 829  | 
"[| a ~= \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"  | 
| 33657 | 830  | 
by (fastsimp intro: le_antisym deg_aboveI deg_belowI)  | 
| 13940 | 831  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
832  | 
lemma deg_const [simp]:  | 
| 13940 | 833  | 
assumes R: "a \<in> carrier R" shows "deg R (monom P a 0) = 0"  | 
| 33657 | 834  | 
proof (rule le_antisym)  | 
| 13940 | 835  | 
show "deg R (monom P a 0) <= 0" by (rule deg_aboveI) (simp_all add: R)  | 
836  | 
next  | 
|
837  | 
show "0 <= deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R)  | 
|
838  | 
qed  | 
|
839  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
840  | 
lemma deg_zero [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
841  | 
"deg R \<zero>\<^bsub>P\<^esub> = 0"  | 
| 33657 | 842  | 
proof (rule le_antisym)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
843  | 
show "deg R \<zero>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all  | 
| 13940 | 844  | 
next  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
845  | 
show "0 <= deg R \<zero>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all  | 
| 13940 | 846  | 
qed  | 
847  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
848  | 
lemma deg_one [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
849  | 
"deg R \<one>\<^bsub>P\<^esub> = 0"  | 
| 33657 | 850  | 
proof (rule le_antisym)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
851  | 
show "deg R \<one>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all  | 
| 13940 | 852  | 
next  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
853  | 
show "0 <= deg R \<one>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all  | 
| 13940 | 854  | 
qed  | 
855  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
856  | 
lemma deg_uminus [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
857  | 
assumes R: "p \<in> carrier P" shows "deg R (\<ominus>\<^bsub>P\<^esub> p) = deg R p"  | 
| 33657 | 858  | 
proof (rule le_antisym)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
859  | 
show "deg R (\<ominus>\<^bsub>P\<^esub> p) <= deg R p" by (simp add: deg_aboveI deg_aboveD R)  | 
| 13940 | 860  | 
next  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
861  | 
show "deg R p <= deg R (\<ominus>\<^bsub>P\<^esub> p)"  | 
| 13940 | 862  | 
by (simp add: deg_belowI lcoeff_nonzero_deg  | 
| 17094 | 863  | 
inj_on_iff [OF R.a_inv_inj, of _ "\<zero>", simplified] R)  | 
| 13940 | 864  | 
qed  | 
865  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
866  | 
text{*The following lemma is later \emph{overwritten} by the most
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
867  | 
  specific one for domains, @{text deg_smult}.*}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
868  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
869  | 
lemma deg_smult_ring [simp]:  | 
| 13940 | 870  | 
"[| a \<in> carrier R; p \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
871  | 
deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"  | 
| 13940 | 872  | 
by (cases "a = \<zero>") (simp add: deg_aboveI deg_aboveD)+  | 
873  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
874  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
875  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
876  | 
context UP_domain  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
877  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
878  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
879  | 
lemma deg_smult [simp]:  | 
| 13940 | 880  | 
assumes R: "a \<in> carrier R" "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
881  | 
shows "deg R (a \<odot>\<^bsub>P\<^esub> p) = (if a = \<zero> then 0 else deg R p)"  | 
| 33657 | 882  | 
proof (rule le_antisym)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
883  | 
show "deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"  | 
| 23350 | 884  | 
using R by (rule deg_smult_ring)  | 
| 13940 | 885  | 
next  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
886  | 
show "(if a = \<zero> then 0 else deg R p) <= deg R (a \<odot>\<^bsub>P\<^esub> p)"  | 
| 13940 | 887  | 
proof (cases "a = \<zero>")  | 
888  | 
qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R)  | 
|
889  | 
qed  | 
|
890  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
891  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
892  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
893  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
894  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
895  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
896  | 
lemma deg_mult_ring:  | 
| 13940 | 897  | 
assumes R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
898  | 
shows "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q"  | 
| 13940 | 899  | 
proof (rule deg_aboveI)  | 
900  | 
fix m  | 
|
901  | 
assume boundm: "deg R p + deg R q < m"  | 
|
902  | 
  {
 | 
|
903  | 
fix k i  | 
|
904  | 
assume boundk: "deg R p + deg R q < k"  | 
|
905  | 
then have "coeff P p i \<otimes> coeff P q (k - i) = \<zero>"  | 
|
906  | 
proof (cases "deg R p < i")  | 
|
907  | 
case True then show ?thesis by (simp add: deg_aboveD R)  | 
|
908  | 
next  | 
|
909  | 
case False with boundk have "deg R q < k - i" by arith  | 
|
910  | 
then show ?thesis by (simp add: deg_aboveD R)  | 
|
911  | 
qed  | 
|
912  | 
}  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
913  | 
with boundm R show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) m = \<zero>" by simp  | 
| 13940 | 914  | 
qed (simp add: R)  | 
915  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
916  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
917  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
918  | 
context UP_domain  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
919  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
920  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
921  | 
lemma deg_mult [simp]:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
922  | 
"[| p ~= \<zero>\<^bsub>P\<^esub>; q ~= \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==>  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
923  | 
deg R (p \<otimes>\<^bsub>P\<^esub> q) = deg R p + deg R q"  | 
| 33657 | 924  | 
proof (rule le_antisym)  | 
| 13940 | 925  | 
assume "p \<in> carrier P" " q \<in> carrier P"  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
926  | 
then show "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_ring)  | 
| 13940 | 927  | 
next  | 
928  | 
let ?s = "(%i. coeff P p i \<otimes> coeff P q (deg R p + deg R q - i))"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
929  | 
assume R: "p \<in> carrier P" "q \<in> carrier P" and nz: "p ~= \<zero>\<^bsub>P\<^esub>" "q ~= \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 930  | 
have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
931  | 
show "deg R p + deg R q <= deg R (p \<otimes>\<^bsub>P\<^esub> q)"  | 
| 13940 | 932  | 
proof (rule deg_belowI, simp add: R)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
933  | 
    have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
934  | 
      = (\<Oplus>i \<in> {..< deg R p} \<union> {deg R p .. deg R p + deg R q}. ?s i)"
 | 
| 13940 | 935  | 
by (simp only: ivl_disj_un_one)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
936  | 
    also have "... = (\<Oplus>i \<in> {deg R p .. deg R p + deg R q}. ?s i)"
 | 
| 17094 | 937  | 
by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one  | 
| 13940 | 938  | 
deg_aboveD less_add_diff R Pi_def)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
939  | 
    also have "...= (\<Oplus>i \<in> {deg R p} \<union> {deg R p <.. deg R p + deg R q}. ?s i)"
 | 
| 13940 | 940  | 
by (simp only: ivl_disj_un_singleton)  | 
| 14666 | 941  | 
also have "... = coeff P p (deg R p) \<otimes> coeff P q (deg R q)"  | 
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
942  | 
by (simp cong: R.finsum_cong add: deg_aboveD R Pi_def)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
943  | 
    finally have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
 | 
| 13940 | 944  | 
= coeff P p (deg R p) \<otimes> coeff P q (deg R q)" .  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
945  | 
    with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) ~= \<zero>"
 | 
| 13940 | 946  | 
by (simp add: integral_iff lcoeff_nonzero R)  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
947  | 
qed (simp add: R)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
948  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
949  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
950  | 
end  | 
| 13940 | 951  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
952  | 
text{*The following lemmas also can be lifted to @{term UP_ring}.*}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
953  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
954  | 
context UP_ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
955  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
956  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
957  | 
lemma coeff_finsum:  | 
| 13940 | 958  | 
assumes fin: "finite A"  | 
959  | 
shows "p \<in> A -> carrier P ==>  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
960  | 
coeff P (finsum P p A) k = (\<Oplus>i \<in> A. coeff P (p i) k)"  | 
| 13940 | 961  | 
using fin by induct (auto simp: Pi_def)  | 
962  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
963  | 
lemma up_repr:  | 
| 13940 | 964  | 
assumes R: "p \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
965  | 
  shows "(\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. monom P (coeff P p i) i) = p"
 | 
| 13940 | 966  | 
proof (rule up_eqI)  | 
967  | 
let ?s = "(%i. monom P (coeff P p i) i)"  | 
|
968  | 
fix k  | 
|
969  | 
from R have RR: "!!i. (if i = k then coeff P p i else \<zero>) \<in> carrier R"  | 
|
970  | 
by simp  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
971  | 
  show "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k = coeff P p k"
 | 
| 13940 | 972  | 
proof (cases "k <= deg R p")  | 
973  | 
case True  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
974  | 
    hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
975  | 
          coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k} \<union> {k<..deg R p}. ?s i) k"
 | 
| 13940 | 976  | 
by (simp only: ivl_disj_un_one)  | 
977  | 
also from True  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
978  | 
    have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k}. ?s i) k"
 | 
| 17094 | 979  | 
by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint  | 
| 14666 | 980  | 
ivl_disj_int_one order_less_imp_not_eq2 coeff_finsum R RR Pi_def)  | 
| 13940 | 981  | 
also  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
982  | 
    have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<k} \<union> {k}. ?s i) k"
 | 
| 13940 | 983  | 
by (simp only: ivl_disj_un_singleton)  | 
984  | 
also have "... = coeff P p k"  | 
|
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
985  | 
by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R RR Pi_def)  | 
| 13940 | 986  | 
finally show ?thesis .  | 
987  | 
next  | 
|
988  | 
case False  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
989  | 
    hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
990  | 
          coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<deg R p} \<union> {deg R p}. ?s i) k"
 | 
| 13940 | 991  | 
by (simp only: ivl_disj_un_singleton)  | 
992  | 
also from False have "... = coeff P p k"  | 
|
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
32436 
diff
changeset
 | 
993  | 
by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R Pi_def)  | 
| 13940 | 994  | 
finally show ?thesis .  | 
995  | 
qed  | 
|
996  | 
qed (simp_all add: R Pi_def)  | 
|
997  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
998  | 
lemma up_repr_le:  | 
| 13940 | 999  | 
"[| deg R p <= n; p \<in> carrier P |] ==>  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1000  | 
  (\<Oplus>\<^bsub>P\<^esub> i \<in> {..n}. monom P (coeff P p i) i) = p"
 | 
| 13940 | 1001  | 
proof -  | 
1002  | 
let ?s = "(%i. monom P (coeff P p i) i)"  | 
|
1003  | 
assume R: "p \<in> carrier P" and "deg R p <= n"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1004  | 
  then have "finsum P ?s {..n} = finsum P ?s ({..deg R p} \<union> {deg R p<..n})"
 | 
| 13940 | 1005  | 
by (simp only: ivl_disj_un_one)  | 
1006  | 
  also have "... = finsum P ?s {..deg R p}"
 | 
|
| 17094 | 1007  | 
by (simp cong: P.finsum_cong add: P.finsum_Un_disjoint ivl_disj_int_one  | 
| 13940 | 1008  | 
deg_aboveD R Pi_def)  | 
| 23350 | 1009  | 
also have "... = p" using R by (rule up_repr)  | 
| 13940 | 1010  | 
finally show ?thesis .  | 
1011  | 
qed  | 
|
1012  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1013  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1014  | 
|
| 17094 | 1015  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
1016  | 
subsection {* Polynomials over Integral Domains *}
 | 
| 13940 | 1017  | 
|
1018  | 
lemma domainI:  | 
|
1019  | 
assumes cring: "cring R"  | 
|
1020  | 
and one_not_zero: "one R ~= zero R"  | 
|
1021  | 
and integral: "!!a b. [| mult R a b = zero R; a \<in> carrier R;  | 
|
1022  | 
b \<in> carrier R |] ==> a = zero R | b = zero R"  | 
|
1023  | 
shows "domain R"  | 
|
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27611 
diff
changeset
 | 
1024  | 
by (auto intro!: domain.intro domain_axioms.intro cring.axioms assms  | 
| 13940 | 1025  | 
del: disjCI)  | 
1026  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1027  | 
context UP_domain  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1028  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1029  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1030  | 
lemma UP_one_not_zero:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1031  | 
"\<one>\<^bsub>P\<^esub> ~= \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 1032  | 
proof  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1033  | 
assume "\<one>\<^bsub>P\<^esub> = \<zero>\<^bsub>P\<^esub>"  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1034  | 
hence "coeff P \<one>\<^bsub>P\<^esub> 0 = (coeff P \<zero>\<^bsub>P\<^esub> 0)" by simp  | 
| 13940 | 1035  | 
hence "\<one> = \<zero>" by simp  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1036  | 
with R.one_not_zero show "False" by contradiction  | 
| 13940 | 1037  | 
qed  | 
1038  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1039  | 
lemma UP_integral:  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1040  | 
"[| p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==> p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 1041  | 
proof -  | 
1042  | 
fix p q  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1043  | 
assume pq: "p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1044  | 
show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"  | 
| 13940 | 1045  | 
proof (rule classical)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1046  | 
assume c: "~ (p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>)"  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1047  | 
with R have "deg R p + deg R q = deg R (p \<otimes>\<^bsub>P\<^esub> q)" by simp  | 
| 13940 | 1048  | 
also from pq have "... = 0" by simp  | 
1049  | 
finally have "deg R p + deg R q = 0" .  | 
|
1050  | 
then have f1: "deg R p = 0 & deg R q = 0" by simp  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1051  | 
    from f1 R have "p = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P p i) i)"
 | 
| 13940 | 1052  | 
by (simp only: up_repr_le)  | 
1053  | 
also from R have "... = monom P (coeff P p 0) 0" by simp  | 
|
1054  | 
finally have p: "p = monom P (coeff P p 0) 0" .  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1055  | 
    from f1 R have "q = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P q i) i)"
 | 
| 13940 | 1056  | 
by (simp only: up_repr_le)  | 
1057  | 
also from R have "... = monom P (coeff P q 0) 0" by simp  | 
|
1058  | 
finally have q: "q = monom P (coeff P q 0) 0" .  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1059  | 
from R have "coeff P p 0 \<otimes> coeff P q 0 = coeff P (p \<otimes>\<^bsub>P\<^esub> q) 0" by simp  | 
| 13940 | 1060  | 
also from pq have "... = \<zero>" by simp  | 
1061  | 
finally have "coeff P p 0 \<otimes> coeff P q 0 = \<zero>" .  | 
|
1062  | 
with R have "coeff P p 0 = \<zero> | coeff P q 0 = \<zero>"  | 
|
1063  | 
by (simp add: R.integral_iff)  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1064  | 
with p q show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>" by fastsimp  | 
| 13940 | 1065  | 
qed  | 
1066  | 
qed  | 
|
1067  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1068  | 
theorem UP_domain:  | 
| 13940 | 1069  | 
"domain P"  | 
1070  | 
by (auto intro!: domainI UP_cring UP_one_not_zero UP_integral del: disjCI)  | 
|
1071  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1072  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1073  | 
|
| 13940 | 1074  | 
text {*
 | 
| 17094 | 1075  | 
  Interpretation of theorems from @{term domain}.
 | 
| 13940 | 1076  | 
*}  | 
1077  | 
||
| 29237 | 1078  | 
sublocale UP_domain < "domain" P  | 
| 
19984
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19931 
diff
changeset
 | 
1079  | 
by intro_locales (rule domain.axioms UP_domain)+  | 
| 13940 | 1080  | 
|
| 14666 | 1081  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
1082  | 
subsection {* The Evaluation Homomorphism and Universal Property*}
 | 
| 13940 | 1083  | 
|
| 14666 | 1084  | 
(* alternative congruence rule (possibly more efficient)  | 
1085  | 
lemma (in abelian_monoid) finsum_cong2:  | 
|
1086  | 
"[| !!i. i \<in> A ==> f i \<in> carrier G = True; A = B;  | 
|
1087  | 
!!i. i \<in> B ==> f i = g i |] ==> finsum G f A = finsum G g B"  | 
|
1088  | 
sorry*)  | 
|
1089  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1090  | 
lemma (in abelian_monoid) boundD_carrier:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1091  | 
"[| bound \<zero> n f; n < m |] ==> f m \<in> carrier G"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1092  | 
by auto  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1093  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1094  | 
context ring  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1095  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1096  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1097  | 
theorem diagonal_sum:  | 
| 13940 | 1098  | 
  "[| f \<in> {..n + m::nat} -> carrier R; g \<in> {..n + m} -> carrier R |] ==>
 | 
| 14666 | 1099  | 
  (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
 | 
1100  | 
  (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
 | 
|
| 13940 | 1101  | 
proof -  | 
1102  | 
  assume Rf: "f \<in> {..n + m} -> carrier R" and Rg: "g \<in> {..n + m} -> carrier R"
 | 
|
1103  | 
  {
 | 
|
1104  | 
fix j  | 
|
1105  | 
have "j <= n + m ==>  | 
|
| 14666 | 1106  | 
      (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
 | 
1107  | 
      (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..j - k}. f k \<otimes> g i)"
 | 
|
| 13940 | 1108  | 
proof (induct j)  | 
1109  | 
case 0 from Rf Rg show ?case by (simp add: Pi_def)  | 
|
1110  | 
next  | 
|
| 14666 | 1111  | 
case (Suc j)  | 
| 13940 | 1112  | 
have R6: "!!i k. [| k <= j; i <= Suc j - k |] ==> g i \<in> carrier R"  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
19984 
diff
changeset
 | 
1113  | 
using Suc by (auto intro!: funcset_mem [OF Rg])  | 
| 13940 | 1114  | 
have R8: "!!i k. [| k <= Suc j; i <= k |] ==> g (k - i) \<in> carrier R"  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
19984 
diff
changeset
 | 
1115  | 
using Suc by (auto intro!: funcset_mem [OF Rg])  | 
| 13940 | 1116  | 
have R9: "!!i k. [| k <= Suc j |] ==> f k \<in> carrier R"  | 
| 14666 | 1117  | 
using Suc by (auto intro!: funcset_mem [OF Rf])  | 
| 13940 | 1118  | 
have R10: "!!i k. [| k <= Suc j; i <= Suc j - k |] ==> g i \<in> carrier R"  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
19984 
diff
changeset
 | 
1119  | 
using Suc by (auto intro!: funcset_mem [OF Rg])  | 
| 13940 | 1120  | 
have R11: "g 0 \<in> carrier R"  | 
| 14666 | 1121  | 
using Suc by (auto intro!: funcset_mem [OF Rg])  | 
| 13940 | 1122  | 
from Suc show ?case  | 
| 14666 | 1123  | 
by (simp cong: finsum_cong add: Suc_diff_le a_ac  | 
1124  | 
Pi_def R6 R8 R9 R10 R11)  | 
|
| 13940 | 1125  | 
qed  | 
1126  | 
}  | 
|
1127  | 
then show ?thesis by fast  | 
|
1128  | 
qed  | 
|
1129  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1130  | 
theorem cauchy_product:  | 
| 13940 | 1131  | 
assumes bf: "bound \<zero> n f" and bg: "bound \<zero> m g"  | 
1132  | 
    and Rf: "f \<in> {..n} -> carrier R" and Rg: "g \<in> {..m} -> carrier R"
 | 
|
| 14666 | 1133  | 
  shows "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
 | 
| 17094 | 1134  | 
    (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"      (* State reverse direction? *)
 | 
| 13940 | 1135  | 
proof -  | 
1136  | 
have f: "!!x. f x \<in> carrier R"  | 
|
1137  | 
proof -  | 
|
1138  | 
fix x  | 
|
1139  | 
show "f x \<in> carrier R"  | 
|
1140  | 
using Rf bf boundD_carrier by (cases "x <= n") (auto simp: Pi_def)  | 
|
1141  | 
qed  | 
|
1142  | 
have g: "!!x. g x \<in> carrier R"  | 
|
1143  | 
proof -  | 
|
1144  | 
fix x  | 
|
1145  | 
show "g x \<in> carrier R"  | 
|
1146  | 
using Rg bg boundD_carrier by (cases "x <= m") (auto simp: Pi_def)  | 
|
1147  | 
qed  | 
|
| 14666 | 1148  | 
  from f g have "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
 | 
1149  | 
      (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
 | 
|
| 13940 | 1150  | 
by (simp add: diagonal_sum Pi_def)  | 
| 15045 | 1151  | 
  also have "... = (\<Oplus>k \<in> {..n} \<union> {n<..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
 | 
| 13940 | 1152  | 
by (simp only: ivl_disj_un_one)  | 
| 14666 | 1153  | 
  also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
 | 
| 13940 | 1154  | 
by (simp cong: finsum_cong  | 
| 14666 | 1155  | 
add: bound.bound [OF bf] finsum_Un_disjoint ivl_disj_int_one Pi_def)  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1156  | 
also from f g  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1157  | 
  have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m} \<union> {m<..n + m - k}. f k \<otimes> g i)"
 | 
| 13940 | 1158  | 
by (simp cong: finsum_cong add: ivl_disj_un_one le_add_diff Pi_def)  | 
| 14666 | 1159  | 
  also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m}. f k \<otimes> g i)"
 | 
| 13940 | 1160  | 
by (simp cong: finsum_cong  | 
| 14666 | 1161  | 
add: bound.bound [OF bg] finsum_Un_disjoint ivl_disj_int_one Pi_def)  | 
1162  | 
  also from f g have "... = (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"
 | 
|
| 13940 | 1163  | 
by (simp add: finsum_ldistr diagonal_sum Pi_def,  | 
1164  | 
simp cong: finsum_cong add: finsum_rdistr Pi_def)  | 
|
1165  | 
finally show ?thesis .  | 
|
1166  | 
qed  | 
|
1167  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1168  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1169  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1170  | 
lemma (in UP_ring) const_ring_hom:  | 
| 13940 | 1171  | 
"(%a. monom P a 0) \<in> ring_hom R P"  | 
1172  | 
by (auto intro!: ring_hom_memI intro: up_eqI simp: monom_mult_is_smult)  | 
|
1173  | 
||
| 27933 | 1174  | 
definition  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1175  | 
  eval :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme,
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1176  | 
'a => 'b, 'b, nat => 'a] => 'b"  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1177  | 
where "eval R S phi s = (\<lambda>p \<in> carrier (UP R).  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1178  | 
    \<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1179  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1180  | 
context UP  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1181  | 
begin  | 
| 14666 | 1182  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1183  | 
lemma eval_on_carrier:  | 
| 19783 | 1184  | 
fixes S (structure)  | 
| 17094 | 1185  | 
shows "p \<in> carrier P ==>  | 
1186  | 
  eval R S phi s p = (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. phi (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
|
| 13940 | 1187  | 
by (unfold eval_def, fold P_def) simp  | 
1188  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1189  | 
lemma eval_extensional:  | 
| 17094 | 1190  | 
"eval R S phi p \<in> extensional (carrier P)"  | 
| 13940 | 1191  | 
by (unfold eval_def, fold P_def) simp  | 
1192  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1193  | 
end  | 
| 17094 | 1194  | 
|
1195  | 
text {* The universal property of the polynomial ring *}
 | 
|
1196  | 
||
| 29240 | 1197  | 
locale UP_pre_univ_prop = ring_hom_cring + UP_cring  | 
1198  | 
||
1199  | 
(* FIXME print_locale ring_hom_cring fails *)  | 
|
| 17094 | 1200  | 
|
| 19783 | 1201  | 
locale UP_univ_prop = UP_pre_univ_prop +  | 
1202  | 
fixes s and Eval  | 
|
| 17094 | 1203  | 
assumes indet_img_carrier [simp, intro]: "s \<in> carrier S"  | 
1204  | 
defines Eval_def: "Eval == eval R S h s"  | 
|
1205  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1206  | 
text{*JE: I have moved the following lemma from Ring.thy and lifted then to the locale @{term ring_hom_ring} from @{term ring_hom_cring}.*}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1207  | 
text{*JE: I was considering using it in @{text eval_ring_hom}, but that property does not hold for non commutative rings, so 
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1208  | 
maybe it is not that necessary.*}  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1209  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1210  | 
lemma (in ring_hom_ring) hom_finsum [simp]:  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1211  | 
"[| finite A; f \<in> A -> carrier R |] ==>  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1212  | 
h (finsum R f A) = finsum S (h o f) A"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1213  | 
proof (induct set: finite)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1214  | 
case empty then show ?case by simp  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1215  | 
next  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1216  | 
case insert then show ?case by (simp add: Pi_def)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1217  | 
qed  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1218  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1219  | 
context UP_pre_univ_prop  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1220  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1221  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1222  | 
theorem eval_ring_hom:  | 
| 17094 | 1223  | 
assumes S: "s \<in> carrier S"  | 
1224  | 
shows "eval R S h s \<in> ring_hom P S"  | 
|
| 13940 | 1225  | 
proof (rule ring_hom_memI)  | 
1226  | 
fix p  | 
|
| 17094 | 1227  | 
assume R: "p \<in> carrier P"  | 
| 13940 | 1228  | 
then show "eval R S h s p \<in> carrier S"  | 
| 17094 | 1229  | 
by (simp only: eval_on_carrier) (simp add: S Pi_def)  | 
| 13940 | 1230  | 
next  | 
1231  | 
fix p q  | 
|
| 17094 | 1232  | 
assume R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1233  | 
then show "eval R S h s (p \<oplus>\<^bsub>P\<^esub> q) = eval R S h s p \<oplus>\<^bsub>S\<^esub> eval R S h s q"  | 
| 17094 | 1234  | 
proof (simp only: eval_on_carrier P.a_closed)  | 
1235  | 
from S R have  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1236  | 
      "(\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1237  | 
      (\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<oplus>\<^bsub>P\<^esub> q)<..max (deg R p) (deg R q)}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1238  | 
h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"  | 
| 17094 | 1239  | 
by (simp cong: S.finsum_cong  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1240  | 
add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def del: coeff_add)  | 
| 17094 | 1241  | 
also from R have "... =  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1242  | 
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..max (deg R p) (deg R q)}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1243  | 
h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"  | 
| 13940 | 1244  | 
by (simp add: ivl_disj_un_one)  | 
| 17094 | 1245  | 
also from R S have "... =  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1246  | 
      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1247  | 
      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 17094 | 1248  | 
by (simp cong: S.finsum_cong  | 
1249  | 
add: S.l_distr deg_aboveD ivl_disj_int_one Pi_def)  | 
|
| 13940 | 1250  | 
also have "... =  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1251  | 
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p} \<union> {deg R p<..max (deg R p) (deg R q)}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1252  | 
h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1253  | 
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q} \<union> {deg R q<..max (deg R p) (deg R q)}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1254  | 
h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"  | 
| 13940 | 1255  | 
by (simp only: ivl_disj_un_one le_maxI1 le_maxI2)  | 
| 17094 | 1256  | 
also from R S have "... =  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1257  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1258  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 17094 | 1259  | 
by (simp cong: S.finsum_cong  | 
1260  | 
add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)  | 
|
| 13940 | 1261  | 
finally show  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1262  | 
      "(\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1263  | 
      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1264  | 
      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
 | 
| 13940 | 1265  | 
qed  | 
1266  | 
next  | 
|
| 17094 | 1267  | 
show "eval R S h s \<one>\<^bsub>P\<^esub> = \<one>\<^bsub>S\<^esub>"  | 
| 13940 | 1268  | 
by (simp only: eval_on_carrier UP_one_closed) simp  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1269  | 
next  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1270  | 
fix p q  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1271  | 
assume R: "p \<in> carrier P" "q \<in> carrier P"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1272  | 
then show "eval R S h s (p \<otimes>\<^bsub>P\<^esub> q) = eval R S h s p \<otimes>\<^bsub>S\<^esub> eval R S h s q"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1273  | 
proof (simp only: eval_on_carrier UP_mult_closed)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1274  | 
from R S have  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1275  | 
      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1276  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<otimes>\<^bsub>P\<^esub> q)<..deg R p + deg R q}.
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1277  | 
h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1278  | 
by (simp cong: S.finsum_cong  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1279  | 
add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1280  | 
del: coeff_mult)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1281  | 
also from R have "... =  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1282  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1283  | 
by (simp only: ivl_disj_un_one deg_mult_ring)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1284  | 
also from R S have "... =  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1285  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}.
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1286  | 
         \<Oplus>\<^bsub>S\<^esub> k \<in> {..i}.
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1287  | 
h (coeff P p k) \<otimes>\<^bsub>S\<^esub> h (coeff P q (i - k)) \<otimes>\<^bsub>S\<^esub>  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1288  | 
(s (^)\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> (i - k)))"  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1289  | 
by (simp cong: S.finsum_cong add: S.nat_pow_mult Pi_def  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1290  | 
S.m_ac S.finsum_rdistr)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1291  | 
also from R S have "... =  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1292  | 
      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1293  | 
      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1294  | 
by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1295  | 
Pi_def)  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1296  | 
finally show  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1297  | 
      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1298  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1299  | 
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1300  | 
qed  | 
| 13940 | 1301  | 
qed  | 
1302  | 
||
| 21502 | 1303  | 
text {*
 | 
1304  | 
  The following lemma could be proved in @{text UP_cring} with the additional
 | 
|
1305  | 
  assumption that @{text h} is closed. *}
 | 
|
| 13940 | 1306  | 
|
| 17094 | 1307  | 
lemma (in UP_pre_univ_prop) eval_const:  | 
| 13940 | 1308  | 
"[| s \<in> carrier S; r \<in> carrier R |] ==> eval R S h s (monom P r 0) = h r"  | 
1309  | 
by (simp only: eval_on_carrier monom_closed) simp  | 
|
1310  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1311  | 
text {* Further properties of the evaluation homomorphism. *}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1312  | 
|
| 13940 | 1313  | 
text {* The following proof is complicated by the fact that in arbitrary
 | 
1314  | 
  rings one might have @{term "one R = zero R"}. *}
 | 
|
1315  | 
||
1316  | 
(* TODO: simplify by cases "one R = zero R" *)  | 
|
1317  | 
||
| 17094 | 1318  | 
lemma (in UP_pre_univ_prop) eval_monom1:  | 
1319  | 
assumes S: "s \<in> carrier S"  | 
|
1320  | 
shows "eval R S h s (monom P \<one> 1) = s"  | 
|
| 13940 | 1321  | 
proof (simp only: eval_on_carrier monom_closed R.one_closed)  | 
| 17094 | 1322  | 
from S have  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1323  | 
    "(\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1324  | 
    (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)} \<union> {deg R (monom P \<one> 1)<..1}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1325  | 
h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"  | 
| 17094 | 1326  | 
by (simp cong: S.finsum_cong del: coeff_monom  | 
1327  | 
add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)  | 
|
| 14666 | 1328  | 
also have "... =  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1329  | 
    (\<Oplus>\<^bsub>S\<^esub> i \<in> {..1}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
 | 
| 13940 | 1330  | 
by (simp only: ivl_disj_un_one deg_monom_le R.one_closed)  | 
1331  | 
also have "... = s"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1332  | 
proof (cases "s = \<zero>\<^bsub>S\<^esub>")  | 
| 13940 | 1333  | 
case True then show ?thesis by (simp add: Pi_def)  | 
1334  | 
next  | 
|
| 17094 | 1335  | 
case False then show ?thesis by (simp add: S Pi_def)  | 
| 13940 | 1336  | 
qed  | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1337  | 
  finally show "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (monom P \<one> 1)}.
 | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1338  | 
h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) = s" .  | 
| 13940 | 1339  | 
qed  | 
1340  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1341  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1342  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1343  | 
text {* Interpretation of ring homomorphism lemmas. *}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1344  | 
|
| 29237 | 1345  | 
sublocale UP_univ_prop < ring_hom_cring P S Eval  | 
| 36092 | 1346  | 
unfolding Eval_def  | 
1347  | 
by unfold_locales (fast intro: eval_ring_hom)  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1348  | 
|
| 13940 | 1349  | 
lemma (in UP_cring) monom_pow:  | 
1350  | 
assumes R: "a \<in> carrier R"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1351  | 
shows "(monom P a n) (^)\<^bsub>P\<^esub> m = monom P (a (^) m) (n * m)"  | 
| 13940 | 1352  | 
proof (induct m)  | 
1353  | 
case 0 from R show ?case by simp  | 
|
1354  | 
next  | 
|
1355  | 
case Suc with R show ?case  | 
|
1356  | 
by (simp del: monom_mult add: monom_mult [THEN sym] add_commute)  | 
|
1357  | 
qed  | 
|
1358  | 
||
1359  | 
lemma (in ring_hom_cring) hom_pow [simp]:  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1360  | 
"x \<in> carrier R ==> h (x (^) n) = h x (^)\<^bsub>S\<^esub> (n::nat)"  | 
| 13940 | 1361  | 
by (induct n) simp_all  | 
1362  | 
||
| 17094 | 1363  | 
lemma (in UP_univ_prop) Eval_monom:  | 
1364  | 
"r \<in> carrier R ==> Eval (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"  | 
|
| 13940 | 1365  | 
proof -  | 
| 17094 | 1366  | 
assume R: "r \<in> carrier R"  | 
1367  | 
from R have "Eval (monom P r n) = Eval (monom P r 0 \<otimes>\<^bsub>P\<^esub> (monom P \<one> 1) (^)\<^bsub>P\<^esub> n)"  | 
|
1368  | 
by (simp del: monom_mult add: monom_mult [THEN sym] monom_pow)  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1369  | 
also  | 
| 17094 | 1370  | 
from R eval_monom1 [where s = s, folded Eval_def]  | 
1371  | 
have "... = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"  | 
|
1372  | 
by (simp add: eval_const [where s = s, folded Eval_def])  | 
|
| 13940 | 1373  | 
finally show ?thesis .  | 
1374  | 
qed  | 
|
1375  | 
||
| 17094 | 1376  | 
lemma (in UP_pre_univ_prop) eval_monom:  | 
1377  | 
assumes R: "r \<in> carrier R" and S: "s \<in> carrier S"  | 
|
1378  | 
shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1379  | 
proof -  | 
| 29237 | 1380  | 
interpret UP_univ_prop R S h P s "eval R S h s"  | 
| 26202 | 1381  | 
using UP_pre_univ_prop_axioms P_def R S  | 
| 22931 | 1382  | 
by (auto intro: UP_univ_prop.intro UP_univ_prop_axioms.intro)  | 
| 17094 | 1383  | 
from R  | 
1384  | 
show ?thesis by (rule Eval_monom)  | 
|
1385  | 
qed  | 
|
1386  | 
||
1387  | 
lemma (in UP_univ_prop) Eval_smult:  | 
|
1388  | 
"[| r \<in> carrier R; p \<in> carrier P |] ==> Eval (r \<odot>\<^bsub>P\<^esub> p) = h r \<otimes>\<^bsub>S\<^esub> Eval p"  | 
|
1389  | 
proof -  | 
|
1390  | 
assume R: "r \<in> carrier R" and P: "p \<in> carrier P"  | 
|
1391  | 
then show ?thesis  | 
|
1392  | 
by (simp add: monom_mult_is_smult [THEN sym]  | 
|
1393  | 
eval_const [where s = s, folded Eval_def])  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1394  | 
qed  | 
| 13940 | 1395  | 
|
1396  | 
lemma ring_hom_cringI:  | 
|
1397  | 
assumes "cring R"  | 
|
1398  | 
and "cring S"  | 
|
1399  | 
and "h \<in> ring_hom R S"  | 
|
1400  | 
shows "ring_hom_cring R S h"  | 
|
1401  | 
by (fast intro: ring_hom_cring.intro ring_hom_cring_axioms.intro  | 
|
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27611 
diff
changeset
 | 
1402  | 
cring.axioms assms)  | 
| 13940 | 1403  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1404  | 
context UP_pre_univ_prop  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1405  | 
begin  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1406  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1407  | 
lemma UP_hom_unique:  | 
| 27611 | 1408  | 
assumes "ring_hom_cring P S Phi"  | 
| 17094 | 1409  | 
assumes Phi: "Phi (monom P \<one> (Suc 0)) = s"  | 
| 13940 | 1410  | 
"!!r. r \<in> carrier R ==> Phi (monom P r 0) = h r"  | 
| 27611 | 1411  | 
assumes "ring_hom_cring P S Psi"  | 
| 17094 | 1412  | 
assumes Psi: "Psi (monom P \<one> (Suc 0)) = s"  | 
| 13940 | 1413  | 
"!!r. r \<in> carrier R ==> Psi (monom P r 0) = h r"  | 
| 17094 | 1414  | 
and P: "p \<in> carrier P" and S: "s \<in> carrier S"  | 
| 13940 | 1415  | 
shows "Phi p = Psi p"  | 
1416  | 
proof -  | 
|
| 29237 | 1417  | 
interpret ring_hom_cring P S Phi by fact  | 
1418  | 
interpret ring_hom_cring P S Psi by fact  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1419  | 
have "Phi p =  | 
| 
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1420  | 
      Phi (\<Oplus>\<^bsub>P \<^esub>i \<in> {..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
 | 
| 17094 | 1421  | 
by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)  | 
| 15696 | 1422  | 
also  | 
1423  | 
have "... =  | 
|
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1424  | 
      Psi (\<Oplus>\<^bsub>P \<^esub>i\<in>{..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
 | 
| 17094 | 1425  | 
by (simp add: Phi Psi P Pi_def comp_def)  | 
| 13940 | 1426  | 
also have "... = Psi p"  | 
| 17094 | 1427  | 
by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)  | 
| 13940 | 1428  | 
finally show ?thesis .  | 
1429  | 
qed  | 
|
1430  | 
||
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1431  | 
lemma ring_homD:  | 
| 17094 | 1432  | 
assumes Phi: "Phi \<in> ring_hom P S"  | 
1433  | 
shows "ring_hom_cring P S Phi"  | 
|
| 36092 | 1434  | 
by unfold_locales (rule Phi)  | 
| 17094 | 1435  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1436  | 
theorem UP_universal_property:  | 
| 17094 | 1437  | 
assumes S: "s \<in> carrier S"  | 
1438  | 
shows "EX! Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) &  | 
|
| 14666 | 1439  | 
Phi (monom P \<one> 1) = s &  | 
| 13940 | 1440  | 
(ALL r : carrier R. Phi (monom P r 0) = h r)"  | 
| 17094 | 1441  | 
using S eval_monom1  | 
| 13940 | 1442  | 
apply (auto intro: eval_ring_hom eval_const eval_extensional)  | 
| 14666 | 1443  | 
apply (rule extensionalityI)  | 
| 17094 | 1444  | 
apply (auto intro: UP_hom_unique ring_homD)  | 
| 14666 | 1445  | 
done  | 
| 13940 | 1446  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1447  | 
end  | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
1448  | 
|
| 27933 | 1449  | 
text{*JE: The following lemma was added by me; it might be even lifted to a simpler locale*}
 | 
1450  | 
||
1451  | 
context monoid  | 
|
1452  | 
begin  | 
|
1453  | 
||
1454  | 
lemma nat_pow_eone[simp]: assumes x_in_G: "x \<in> carrier G" shows "x (^) (1::nat) = x"  | 
|
1455  | 
using nat_pow_Suc [of x 0] unfolding nat_pow_0 [of x] unfolding l_one [OF x_in_G] by simp  | 
|
1456  | 
||
1457  | 
end  | 
|
1458  | 
||
1459  | 
context UP_ring  | 
|
1460  | 
begin  | 
|
1461  | 
||
1462  | 
abbreviation lcoeff :: "(nat =>'a) => 'a" where "lcoeff p == coeff P p (deg R p)"  | 
|
1463  | 
||
1464  | 
lemma lcoeff_nonzero2: assumes p_in_R: "p \<in> carrier P" and p_not_zero: "p \<noteq> \<zero>\<^bsub>P\<^esub>" shows "lcoeff p \<noteq> \<zero>"  | 
|
1465  | 
using lcoeff_nonzero [OF p_not_zero p_in_R] .  | 
|
1466  | 
||
| 35849 | 1467  | 
|
| 27933 | 1468  | 
subsection{*The long division algorithm: some previous facts.*}
 | 
1469  | 
||
1470  | 
lemma coeff_minus [simp]:  | 
|
1471  | 
assumes p: "p \<in> carrier P" and q: "q \<in> carrier P" shows "coeff P (p \<ominus>\<^bsub>P\<^esub> q) n = coeff P p n \<ominus> coeff P q n"  | 
|
1472  | 
unfolding a_minus_def [OF p q] unfolding coeff_add [OF p a_inv_closed [OF q]] unfolding coeff_a_inv [OF q]  | 
|
1473  | 
using coeff_closed [OF p, of n] using coeff_closed [OF q, of n] by algebra  | 
|
1474  | 
||
1475  | 
lemma lcoeff_closed [simp]: assumes p: "p \<in> carrier P" shows "lcoeff p \<in> carrier R"  | 
|
1476  | 
using coeff_closed [OF p, of "deg R p"] by simp  | 
|
1477  | 
||
1478  | 
lemma deg_smult_decr: assumes a_in_R: "a \<in> carrier R" and f_in_P: "f \<in> carrier P" shows "deg R (a \<odot>\<^bsub>P\<^esub> f) \<le> deg R f"  | 
|
1479  | 
using deg_smult_ring [OF a_in_R f_in_P] by (cases "a = \<zero>", auto)  | 
|
1480  | 
||
1481  | 
lemma coeff_monom_mult: assumes R: "c \<in> carrier R" and P: "p \<in> carrier P"  | 
|
1482  | 
shows "coeff P (monom P c n \<otimes>\<^bsub>P\<^esub> p) (m + n) = c \<otimes> (coeff P p m)"  | 
|
1483  | 
proof -  | 
|
1484  | 
  have "coeff P (monom P c n \<otimes>\<^bsub>P\<^esub> p) (m + n) = (\<Oplus>i\<in>{..m + n}. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i))"
 | 
|
1485  | 
unfolding coeff_mult [OF monom_closed [OF R, of n] P, of "m + n"] unfolding coeff_monom [OF R, of n] by simp  | 
|
1486  | 
  also have "(\<Oplus>i\<in>{..m + n}. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i)) = 
 | 
|
1487  | 
    (\<Oplus>i\<in>{..m + n}. (if n = i then c \<otimes> coeff P p (m + n - i) else \<zero>))"
 | 
|
1488  | 
    using  R.finsum_cong [of "{..m + n}" "{..m + n}" "(\<lambda>i::nat. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i))" 
 | 
|
1489  | 
"(\<lambda>i::nat. (if n = i then c \<otimes> coeff P p (m + n - i) else \<zero>))"]  | 
|
1490  | 
using coeff_closed [OF P] unfolding Pi_def simp_implies_def using R by auto  | 
|
1491  | 
  also have "\<dots> = c \<otimes> coeff P p m" using R.finsum_singleton [of n "{..m + n}" "(\<lambda>i. c \<otimes> coeff P p (m + n - i))"]
 | 
|
1492  | 
unfolding Pi_def using coeff_closed [OF P] using P R by auto  | 
|
1493  | 
finally show ?thesis by simp  | 
|
1494  | 
qed  | 
|
1495  | 
||
1496  | 
lemma deg_lcoeff_cancel:  | 
|
1497  | 
assumes p_in_P: "p \<in> carrier P" and q_in_P: "q \<in> carrier P" and r_in_P: "r \<in> carrier P"  | 
|
1498  | 
and deg_r_nonzero: "deg R r \<noteq> 0"  | 
|
1499  | 
and deg_R_p: "deg R p \<le> deg R r" and deg_R_q: "deg R q \<le> deg R r"  | 
|
1500  | 
and coeff_R_p_eq_q: "coeff P p (deg R r) = \<ominus>\<^bsub>R\<^esub> (coeff P q (deg R r))"  | 
|
1501  | 
shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) < deg R r"  | 
|
1502  | 
proof -  | 
|
1503  | 
have deg_le: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<le> deg R r"  | 
|
1504  | 
proof (rule deg_aboveI)  | 
|
1505  | 
fix m  | 
|
1506  | 
assume deg_r_le: "deg R r < m"  | 
|
1507  | 
show "coeff P (p \<oplus>\<^bsub>P\<^esub> q) m = \<zero>"  | 
|
1508  | 
proof -  | 
|
1509  | 
have slp: "deg R p < m" and "deg R q < m" using deg_R_p deg_R_q using deg_r_le by auto  | 
|
1510  | 
then have max_sl: "max (deg R p) (deg R q) < m" by simp  | 
|
1511  | 
then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith  | 
|
1512  | 
with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1513  | 
using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp  | 
| 27933 | 1514  | 
qed  | 
1515  | 
qed (simp add: p_in_P q_in_P)  | 
|
1516  | 
moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"  | 
|
1517  | 
proof (rule ccontr)  | 
|
1518  | 
assume nz: "\<not> deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r" then have deg_eq: "deg R (p \<oplus>\<^bsub>P\<^esub> q) = deg R r" by simp  | 
|
1519  | 
from deg_r_nonzero have r_nonzero: "r \<noteq> \<zero>\<^bsub>P\<^esub>" by (cases "r = \<zero>\<^bsub>P\<^esub>", simp_all)  | 
|
1520  | 
have "coeff P (p \<oplus>\<^bsub>P\<^esub> q) (deg R r) = \<zero>\<^bsub>R\<^esub>" using coeff_add [OF p_in_P q_in_P, of "deg R r"] using coeff_R_p_eq_q  | 
|
1521  | 
using coeff_closed [OF p_in_P, of "deg R r"] coeff_closed [OF q_in_P, of "deg R r"] by algebra  | 
|
1522  | 
with lcoeff_nonzero [OF r_nonzero r_in_P] and deg_eq show False using lcoeff_nonzero [of "p \<oplus>\<^bsub>P\<^esub> q"] using p_in_P q_in_P  | 
|
1523  | 
using deg_r_nonzero by (cases "p \<oplus>\<^bsub>P\<^esub> q \<noteq> \<zero>\<^bsub>P\<^esub>", auto)  | 
|
1524  | 
qed  | 
|
1525  | 
ultimately show ?thesis by simp  | 
|
1526  | 
qed  | 
|
1527  | 
||
1528  | 
lemma monom_deg_mult:  | 
|
1529  | 
assumes f_in_P: "f \<in> carrier P" and g_in_P: "g \<in> carrier P" and deg_le: "deg R g \<le> deg R f"  | 
|
1530  | 
and a_in_R: "a \<in> carrier R"  | 
|
1531  | 
shows "deg R (g \<otimes>\<^bsub>P\<^esub> monom P a (deg R f - deg R g)) \<le> deg R f"  | 
|
1532  | 
using deg_mult_ring [OF g_in_P monom_closed [OF a_in_R, of "deg R f - deg R g"]]  | 
|
1533  | 
apply (cases "a = \<zero>") using g_in_P apply simp  | 
|
1534  | 
using deg_monom [OF _ a_in_R, of "deg R f - deg R g"] using deg_le by simp  | 
|
1535  | 
||
1536  | 
lemma deg_zero_impl_monom:  | 
|
1537  | 
assumes f_in_P: "f \<in> carrier P" and deg_f: "deg R f = 0"  | 
|
1538  | 
shows "f = monom P (coeff P f 0) 0"  | 
|
1539  | 
apply (rule up_eqI) using coeff_monom [OF coeff_closed [OF f_in_P], of 0 0]  | 
|
1540  | 
using f_in_P deg_f using deg_aboveD [of f _] by auto  | 
|
1541  | 
||
1542  | 
end  | 
|
1543  | 
||
1544  | 
||
1545  | 
subsection {* The long division proof for commutative rings *}
 | 
|
1546  | 
||
1547  | 
context UP_cring  | 
|
1548  | 
begin  | 
|
1549  | 
||
1550  | 
lemma exI3: assumes exist: "Pred x y z"  | 
|
1551  | 
shows "\<exists> x y z. Pred x y z"  | 
|
1552  | 
using exist by blast  | 
|
1553  | 
||
1554  | 
text {* Jacobson's Theorem 2.14 *}
 | 
|
1555  | 
||
1556  | 
lemma long_div_theorem:  | 
|
1557  | 
assumes g_in_P [simp]: "g \<in> carrier P" and f_in_P [simp]: "f \<in> carrier P"  | 
|
1558  | 
and g_not_zero: "g \<noteq> \<zero>\<^bsub>P\<^esub>"  | 
|
1559  | 
shows "\<exists> q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"  | 
|
1560  | 
proof -  | 
|
1561  | 
let ?pred = "(\<lambda> q r (k::nat).  | 
|
1562  | 
(q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"  | 
|
1563  | 
and ?lg = "lcoeff g"  | 
|
1564  | 
show ?thesis  | 
|
1565  | 
(*JE: we distinguish some particular cases where the solution is almost direct.*)  | 
|
1566  | 
proof (cases "deg R f < deg R g")  | 
|
1567  | 
case True  | 
|
1568  | 
(*JE: if the degree of f is smaller than the one of g the solution is straightforward.*)  | 
|
1569  | 
(* CB: avoid exI3 *)  | 
|
1570  | 
have "?pred \<zero>\<^bsub>P\<^esub> f 0" using True by force  | 
|
1571  | 
then show ?thesis by fast  | 
|
1572  | 
next  | 
|
1573  | 
case False then have deg_g_le_deg_f: "deg R g \<le> deg R f" by simp  | 
|
1574  | 
    {
 | 
|
1575  | 
(*JE: we now apply the induction hypothesis with some additional facts required*)  | 
|
1576  | 
from f_in_P deg_g_le_deg_f show ?thesis  | 
|
| 34915 | 1577  | 
proof (induct "deg R f" arbitrary: "f" rule: less_induct)  | 
1578  | 
case less  | 
|
1579  | 
note f_in_P [simp] = `f \<in> carrier P`  | 
|
1580  | 
and deg_g_le_deg_f = `deg R g \<le> deg R f`  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1581  | 
let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1582  | 
and ?q = "monom P (lcoeff f) (deg R f - deg R g)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1583  | 
show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1584  | 
proof -  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1585  | 
(*JE: we first extablish the existence of a triple satisfying the previous equation.  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1586  | 
Then we will have to prove the second part of the predicate.*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1587  | 
have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1588  | 
using minus_add  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1589  | 
using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1590  | 
using r_neg by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1591  | 
show ?thesis  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1592  | 
proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1593  | 
(*JE: if the degree of the remainder satisfies the statement property we are done*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1594  | 
case True  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1595  | 
            {
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1596  | 
show ?thesis  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1597  | 
proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1598  | 
show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1599  | 
show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1600  | 
qed (simp_all)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1601  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1602  | 
next  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1603  | 
case False note n_deg_r_l_deg_g = False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1604  | 
            {
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1605  | 
(*JE: otherwise, we verify the conditions of the induction hypothesis.*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1606  | 
show ?thesis  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1607  | 
proof (cases "deg R f = 0")  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1608  | 
(*JE: the solutions are different if the degree of f is zero or not*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1609  | 
case True  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1610  | 
                {
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1611  | 
have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1612  | 
have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1613  | 
unfolding deg_g apply simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1614  | 
unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1615  | 
using deg_zero_impl_monom [OF g_in_P deg_g] by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1616  | 
then show ?thesis using f_in_P by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1617  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1618  | 
next  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1619  | 
case False note deg_f_nzero = False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1620  | 
                {
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1621  | 
(*JE: now it only remains the case where the induction hypothesis can be used.*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1622  | 
(*JE: we first prove that the degree of the remainder is smaller than the one of f*)  | 
| 34915 | 1623  | 
have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R f"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1624  | 
proof -  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1625  | 
have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1626  | 
also have "\<dots> < deg R f"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1627  | 
proof (rule deg_lcoeff_cancel)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1628  | 
show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"  | 
| 34915 | 1629  | 
using deg_smult_ring [of "lcoeff g" f]  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1630  | 
using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1631  | 
show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1632  | 
using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1633  | 
by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1634  | 
show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1635  | 
unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1636  | 
unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1637  | 
                        using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1638  | 
"(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1639  | 
"(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1640  | 
                        using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1641  | 
unfolding Pi_def using deg_g_le_deg_f by force  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1642  | 
qed (simp_all add: deg_f_nzero)  | 
| 34915 | 1643  | 
finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R f" .  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1644  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1645  | 
moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1646  | 
moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1647  | 
moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1648  | 
(*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1649  | 
ultimately obtain q' r' k'  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1650  | 
where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1651  | 
and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"  | 
| 34915 | 1652  | 
using less by blast  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1653  | 
(*JE: we now prove that the new quotient, remainder and exponent can be used to get  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1654  | 
the quotient, remainder and exponent of the long division theorem*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1655  | 
show ?thesis  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1656  | 
proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1657  | 
show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1658  | 
proof -  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1659  | 
have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1660  | 
using smult_assoc1 exist by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1661  | 
also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1662  | 
using UP_smult_r_distr by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1663  | 
also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1664  | 
using rem_desc by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1665  | 
also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1666  | 
using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1667  | 
using q'_in_carrier r'_in_carrier by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1668  | 
also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1669  | 
using q'_in_carrier by (auto simp add: m_comm)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1670  | 
also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1671  | 
using smult_assoc2 q'_in_carrier by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1672  | 
also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1673  | 
using sym [OF l_distr] and q'_in_carrier by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1674  | 
finally show ?thesis using m_comm q'_in_carrier by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1675  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1676  | 
qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1677  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1678  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1679  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1680  | 
qed  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1681  | 
qed  | 
| 27933 | 1682  | 
qed  | 
1683  | 
}  | 
|
1684  | 
qed  | 
|
1685  | 
qed  | 
|
1686  | 
||
1687  | 
end  | 
|
1688  | 
||
1689  | 
||
1690  | 
text {*The remainder theorem as corollary of the long division theorem.*}
 | 
|
1691  | 
||
1692  | 
context UP_cring  | 
|
1693  | 
begin  | 
|
1694  | 
||
1695  | 
lemma deg_minus_monom:  | 
|
1696  | 
assumes a: "a \<in> carrier R"  | 
|
1697  | 
  and R_not_trivial: "(carrier R \<noteq> {\<zero>})"
 | 
|
1698  | 
shows "deg R (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = 1"  | 
|
1699  | 
(is "deg R ?g = 1")  | 
|
1700  | 
proof -  | 
|
1701  | 
have "deg R ?g \<le> 1"  | 
|
1702  | 
proof (rule deg_aboveI)  | 
|
1703  | 
fix m  | 
|
1704  | 
assume "(1::nat) < m"  | 
|
1705  | 
then show "coeff P ?g m = \<zero>"  | 
|
1706  | 
using coeff_minus using a by auto algebra  | 
|
1707  | 
qed (simp add: a)  | 
|
1708  | 
moreover have "deg R ?g \<ge> 1"  | 
|
1709  | 
proof (rule deg_belowI)  | 
|
1710  | 
show "coeff P ?g 1 \<noteq> \<zero>"  | 
|
1711  | 
using a using R.carrier_one_not_zero R_not_trivial by simp algebra  | 
|
1712  | 
qed (simp add: a)  | 
|
1713  | 
ultimately show ?thesis by simp  | 
|
1714  | 
qed  | 
|
1715  | 
||
1716  | 
lemma lcoeff_monom:  | 
|
1717  | 
  assumes a: "a \<in> carrier R" and R_not_trivial: "(carrier R \<noteq> {\<zero>})"
 | 
|
1718  | 
shows "lcoeff (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = \<one>"  | 
|
1719  | 
using deg_minus_monom [OF a R_not_trivial]  | 
|
1720  | 
using coeff_minus a by auto algebra  | 
|
1721  | 
||
1722  | 
lemma deg_nzero_nzero:  | 
|
1723  | 
assumes deg_p_nzero: "deg R p \<noteq> 0"  | 
|
1724  | 
shows "p \<noteq> \<zero>\<^bsub>P\<^esub>"  | 
|
1725  | 
using deg_zero deg_p_nzero by auto  | 
|
1726  | 
||
1727  | 
lemma deg_monom_minus:  | 
|
1728  | 
assumes a: "a \<in> carrier R"  | 
|
1729  | 
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
 | 
|
1730  | 
shows "deg R (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = 1"  | 
|
1731  | 
(is "deg R ?g = 1")  | 
|
1732  | 
proof -  | 
|
1733  | 
have "deg R ?g \<le> 1"  | 
|
1734  | 
proof (rule deg_aboveI)  | 
|
1735  | 
fix m::nat assume "1 < m" then show "coeff P ?g m = \<zero>"  | 
|
1736  | 
using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of m]  | 
|
1737  | 
using coeff_monom [OF R.one_closed, of 1 m] using coeff_monom [OF a, of 0 m] by auto algebra  | 
|
1738  | 
qed (simp add: a)  | 
|
1739  | 
moreover have "1 \<le> deg R ?g"  | 
|
1740  | 
proof (rule deg_belowI)  | 
|
1741  | 
show "coeff P ?g 1 \<noteq> \<zero>"  | 
|
1742  | 
using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of 1]  | 
|
1743  | 
using coeff_monom [OF R.one_closed, of 1 1] using coeff_monom [OF a, of 0 1]  | 
|
1744  | 
using R_not_trivial using R.carrier_one_not_zero  | 
|
1745  | 
by auto algebra  | 
|
1746  | 
qed (simp add: a)  | 
|
1747  | 
ultimately show ?thesis by simp  | 
|
1748  | 
qed  | 
|
1749  | 
||
1750  | 
lemma eval_monom_expr:  | 
|
1751  | 
assumes a: "a \<in> carrier R"  | 
|
1752  | 
shows "eval R R id a (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = \<zero>"  | 
|
1753  | 
(is "eval R R id a ?g = _")  | 
|
1754  | 
proof -  | 
|
| 36092 | 1755  | 
interpret UP_pre_univ_prop R R id by unfold_locales simp  | 
| 27933 | 1756  | 
have eval_ring_hom: "eval R R id a \<in> ring_hom P R" using eval_ring_hom [OF a] by simp  | 
| 36092 | 1757  | 
interpret ring_hom_cring P R "eval R R id a" by unfold_locales (rule eval_ring_hom)  | 
| 27933 | 1758  | 
have mon1_closed: "monom P \<one>\<^bsub>R\<^esub> 1 \<in> carrier P"  | 
1759  | 
and mon0_closed: "monom P a 0 \<in> carrier P"  | 
|
1760  | 
and min_mon0_closed: "\<ominus>\<^bsub>P\<^esub> monom P a 0 \<in> carrier P"  | 
|
1761  | 
using a R.a_inv_closed by auto  | 
|
1762  | 
have "eval R R id a ?g = eval R R id a (monom P \<one> 1) \<ominus> eval R R id a (monom P a 0)"  | 
|
1763  | 
unfolding P.minus_eq [OF mon1_closed mon0_closed]  | 
|
| 29246 | 1764  | 
unfolding hom_add [OF mon1_closed min_mon0_closed]  | 
1765  | 
unfolding hom_a_inv [OF mon0_closed]  | 
|
| 27933 | 1766  | 
using R.minus_eq [symmetric] mon1_closed mon0_closed by auto  | 
1767  | 
also have "\<dots> = a \<ominus> a"  | 
|
1768  | 
using eval_monom [OF R.one_closed a, of 1] using eval_monom [OF a a, of 0] using a by simp  | 
|
1769  | 
also have "\<dots> = \<zero>"  | 
|
1770  | 
using a by algebra  | 
|
1771  | 
finally show ?thesis by simp  | 
|
1772  | 
qed  | 
|
1773  | 
||
1774  | 
lemma remainder_theorem_exist:  | 
|
1775  | 
assumes f: "f \<in> carrier P" and a: "a \<in> carrier R"  | 
|
1776  | 
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
 | 
|
1777  | 
shows "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (deg R r = 0)"  | 
|
1778  | 
(is "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (deg R r = 0)")  | 
|
1779  | 
proof -  | 
|
1780  | 
let ?g = "monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0"  | 
|
1781  | 
from deg_minus_monom [OF a R_not_trivial]  | 
|
1782  | 
have deg_g_nzero: "deg R ?g \<noteq> 0" by simp  | 
|
1783  | 
have "\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and>  | 
|
1784  | 
lcoeff ?g (^) k \<odot>\<^bsub>P\<^esub> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R ?g)"  | 
|
1785  | 
using long_div_theorem [OF _ f deg_nzero_nzero [OF deg_g_nzero]] a  | 
|
1786  | 
by auto  | 
|
1787  | 
then show ?thesis  | 
|
1788  | 
unfolding lcoeff_monom [OF a R_not_trivial]  | 
|
1789  | 
unfolding deg_monom_minus [OF a R_not_trivial]  | 
|
1790  | 
using smult_one [OF f] using deg_zero by force  | 
|
1791  | 
qed  | 
|
1792  | 
||
1793  | 
lemma remainder_theorem_expression:  | 
|
1794  | 
assumes f [simp]: "f \<in> carrier P" and a [simp]: "a \<in> carrier R"  | 
|
1795  | 
and q [simp]: "q \<in> carrier P" and r [simp]: "r \<in> carrier P"  | 
|
1796  | 
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
 | 
|
1797  | 
and f_expr: "f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r"  | 
|
1798  | 
(is "f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r" is "f = ?gq \<oplus>\<^bsub>P\<^esub> r")  | 
|
1799  | 
and deg_r_0: "deg R r = 0"  | 
|
1800  | 
shows "r = monom P (eval R R id a f) 0"  | 
|
1801  | 
proof -  | 
|
| 29237 | 1802  | 
interpret UP_pre_univ_prop R R id P proof qed simp  | 
| 27933 | 1803  | 
have eval_ring_hom: "eval R R id a \<in> ring_hom P R"  | 
1804  | 
using eval_ring_hom [OF a] by simp  | 
|
1805  | 
have "eval R R id a f = eval R R id a ?gq \<oplus>\<^bsub>R\<^esub> eval R R id a r"  | 
|
1806  | 
unfolding f_expr using ring_hom_add [OF eval_ring_hom] by auto  | 
|
1807  | 
also have "\<dots> = ((eval R R id a ?g) \<otimes> (eval R R id a q)) \<oplus>\<^bsub>R\<^esub> eval R R id a r"  | 
|
1808  | 
using ring_hom_mult [OF eval_ring_hom] by auto  | 
|
1809  | 
also have "\<dots> = \<zero> \<oplus> eval R R id a r"  | 
|
1810  | 
unfolding eval_monom_expr [OF a] using eval_ring_hom  | 
|
1811  | 
unfolding ring_hom_def using q unfolding Pi_def by simp  | 
|
1812  | 
also have "\<dots> = eval R R id a r"  | 
|
1813  | 
using eval_ring_hom unfolding ring_hom_def using r unfolding Pi_def by simp  | 
|
1814  | 
finally have eval_eq: "eval R R id a f = eval R R id a r" by simp  | 
|
1815  | 
from deg_zero_impl_monom [OF r deg_r_0]  | 
|
1816  | 
have "r = monom P (coeff P r 0) 0" by simp  | 
|
1817  | 
with eval_const [OF a, of "coeff P r 0"] eval_eq  | 
|
1818  | 
show ?thesis by auto  | 
|
1819  | 
qed  | 
|
1820  | 
||
1821  | 
corollary remainder_theorem:  | 
|
1822  | 
assumes f [simp]: "f \<in> carrier P" and a [simp]: "a \<in> carrier R"  | 
|
1823  | 
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
 | 
|
1824  | 
shows "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and>  | 
|
1825  | 
f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> monom P (eval R R id a f) 0"  | 
|
1826  | 
(is "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> monom P (eval R R id a f) 0")  | 
|
1827  | 
proof -  | 
|
1828  | 
from remainder_theorem_exist [OF f a R_not_trivial]  | 
|
1829  | 
obtain q r  | 
|
1830  | 
where q_r: "q \<in> carrier P \<and> r \<in> carrier P \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r"  | 
|
1831  | 
and deg_r: "deg R r = 0" by force  | 
|
1832  | 
with remainder_theorem_expression [OF f a _ _ R_not_trivial, of q r]  | 
|
1833  | 
show ?thesis by auto  | 
|
1834  | 
qed  | 
|
1835  | 
||
1836  | 
end  | 
|
1837  | 
||
| 17094 | 1838  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents: 
20282 
diff
changeset
 | 
1839  | 
subsection {* Sample Application of Evaluation Homomorphism *}
 | 
| 13940 | 1840  | 
|
| 17094 | 1841  | 
lemma UP_pre_univ_propI:  | 
| 13940 | 1842  | 
assumes "cring R"  | 
1843  | 
and "cring S"  | 
|
1844  | 
and "h \<in> ring_hom R S"  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19783 
diff
changeset
 | 
1845  | 
shows "UP_pre_univ_prop R S h"  | 
| 23350 | 1846  | 
using assms  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19783 
diff
changeset
 | 
1847  | 
by (auto intro!: UP_pre_univ_prop.intro ring_hom_cring.intro  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19783 
diff
changeset
 | 
1848  | 
ring_hom_cring_axioms.intro UP_cring.intro)  | 
| 13940 | 1849  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1850  | 
definition  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1851  | 
INTEG :: "int ring"  | 
| 
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1852  | 
where "INTEG = (| carrier = UNIV, mult = op *, one = 1, zero = 0, add = op + |)"  | 
| 13975 | 1853  | 
|
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
1854  | 
lemma INTEG_cring: "cring INTEG"  | 
| 13975 | 1855  | 
by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI  | 
1856  | 
zadd_zminus_inverse2 zadd_zmult_distrib)  | 
|
1857  | 
||
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1858  | 
lemma INTEG_id_eval:  | 
| 17094 | 1859  | 
"UP_pre_univ_prop INTEG INTEG id"  | 
1860  | 
by (fast intro: UP_pre_univ_propI INTEG_cring id_ring_hom)  | 
|
| 13940 | 1861  | 
|
1862  | 
text {*
 | 
|
| 17094 | 1863  | 
Interpretation now enables to import all theorems and lemmas  | 
| 13940 | 1864  | 
  valid in the context of homomorphisms between @{term INTEG} and @{term
 | 
| 
15095
 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 
ballarin 
parents: 
15076 
diff
changeset
 | 
1865  | 
"UP INTEG"} globally.  | 
| 14666 | 1866  | 
*}  | 
| 13940 | 1867  | 
|
| 
30729
 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
1868  | 
interpretation INTEG: UP_pre_univ_prop INTEG INTEG id "UP INTEG"  | 
| 28823 | 1869  | 
using INTEG_id_eval by simp_all  | 
| 
15763
 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 
ballarin 
parents: 
15696 
diff
changeset
 | 
1870  | 
|
| 13940 | 1871  | 
lemma INTEG_closed [intro, simp]:  | 
1872  | 
"z \<in> carrier INTEG"  | 
|
1873  | 
by (unfold INTEG_def) simp  | 
|
1874  | 
||
1875  | 
lemma INTEG_mult [simp]:  | 
|
1876  | 
"mult INTEG z w = z * w"  | 
|
1877  | 
by (unfold INTEG_def) simp  | 
|
1878  | 
||
1879  | 
lemma INTEG_pow [simp]:  | 
|
1880  | 
"pow INTEG z n = z ^ n"  | 
|
1881  | 
by (induct n) (simp_all add: INTEG_def nat_pow_def)  | 
|
1882  | 
||
1883  | 
lemma "eval INTEG INTEG id 10 (monom (UP INTEG) 5 2) = 500"  | 
|
| 
15763
 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 
ballarin 
parents: 
15696 
diff
changeset
 | 
1884  | 
by (simp add: INTEG.eval_monom)  | 
| 13940 | 1885  | 
|
| 14590 | 1886  | 
end  |