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