src/HOL/Algebra/UnivPoly.thy
author ballarin
Thu, 31 Jul 2008 09:49:21 +0200
changeset 27714 27b4d7c01f8b
parent 27611 2c01c0bdb385
child 27717 21bbd410ba04
permissions -rw-r--r--
Tuned (for the sake of a meaningless log entry).
     1 (*
     2   Title:     HOL/Algebra/UnivPoly.thy
     3   Id:        $Id$
     4   Author:    Clemens Ballarin, started 9 December 1996
     5   Copyright: Clemens Ballarin
     6 *)
     7 
     8 theory UnivPoly imports Module begin
     9 
    10 
    11 section {* Univariate Polynomials *}
    12 
    13 text {*
    14   Polynomials are formalised as modules with additional operations for
    15   extracting coefficients from polynomials and for obtaining monomials
    16   from coefficients and exponents (record @{text "up_ring"}).  The
    17   carrier set is a set of bounded functions from Nat to the
    18   coefficient domain.  Bounded means that these functions return zero
    19   above a certain bound (the degree).  There is a chapter on the
    20   formalisation of polynomials in the PhD thesis \cite{Ballarin:1999},
    21   which was implemented with axiomatic type classes.  This was later
    22   ported to Locales.
    23 *}
    24 
    25 
    26 subsection {* The Constructor for Univariate Polynomials *}
    27 
    28 text {*
    29   Functions with finite support.
    30 *}
    31 
    32 locale bound =
    33   fixes z :: 'a
    34     and n :: nat
    35     and f :: "nat => 'a"
    36   assumes bound: "!!m. n < m \<Longrightarrow> f m = z"
    37 
    38 declare bound.intro [intro!]
    39   and bound.bound [dest]
    40 
    41 lemma bound_below:
    42   assumes bound: "bound z m f" and nonzero: "f n \<noteq> z" shows "n \<le> m"
    43 proof (rule classical)
    44   assume "~ ?thesis"
    45   then have "m < n" by arith
    46   with bound have "f n = z" ..
    47   with nonzero show ?thesis by contradiction
    48 qed
    49 
    50 record ('a, 'p) up_ring = "('a, 'p) module" +
    51   monom :: "['a, nat] => 'p"
    52   coeff :: "['p, nat] => 'a"
    53 
    54 constdefs (structure R)
    55   up :: "('a, 'm) ring_scheme => (nat => 'a) set"
    56   "up R == {f. f \<in> UNIV -> carrier R & (EX n. bound \<zero> n f)}"
    57   UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring"
    58   "UP R == (|
    59     carrier = up R,
    60     mult = (%p:up R. %q:up R. %n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)),
    61     one = (%i. if i=0 then \<one> else \<zero>),
    62     zero = (%i. \<zero>),
    63     add = (%p:up R. %q:up R. %i. p i \<oplus> q i),
    64     smult = (%a:carrier R. %p:up R. %i. a \<otimes> p i),
    65     monom = (%a:carrier R. %n i. if i=n then a else \<zero>),
    66     coeff = (%p:up R. %n. p n) |)"
    67 
    68 text {*
    69   Properties of the set of polynomials @{term up}.
    70 *}
    71 
    72 lemma mem_upI [intro]:
    73   "[| !!n. f n \<in> carrier R; EX n. bound (zero R) n f |] ==> f \<in> up R"
    74   by (simp add: up_def Pi_def)
    75 
    76 lemma mem_upD [dest]:
    77   "f \<in> up R ==> f n \<in> carrier R"
    78   by (simp add: up_def Pi_def)
    79 
    80 lemma (in cring) bound_upD [dest]:
    81   "f \<in> up R ==> EX n. bound \<zero> n f"
    82   by (simp add: up_def)
    83 
    84 lemma (in cring) up_one_closed:
    85    "(%n. if n = 0 then \<one> else \<zero>) \<in> up R"
    86   using up_def by force
    87 
    88 lemma (in cring) up_smult_closed:
    89   "[| a \<in> carrier R; p \<in> up R |] ==> (%i. a \<otimes> p i) \<in> up R"
    90   by force
    91 
    92 lemma (in cring) up_add_closed:
    93   "[| p \<in> up R; q \<in> up R |] ==> (%i. p i \<oplus> q i) \<in> up R"
    94 proof
    95   fix n
    96   assume "p \<in> up R" and "q \<in> up R"
    97   then show "p n \<oplus> q n \<in> carrier R"
    98     by auto
    99 next
   100   assume UP: "p \<in> up R" "q \<in> up R"
   101   show "EX n. bound \<zero> n (%i. p i \<oplus> q i)"
   102   proof -
   103     from UP obtain n where boundn: "bound \<zero> n p" by fast
   104     from UP obtain m where boundm: "bound \<zero> m q" by fast
   105     have "bound \<zero> (max n m) (%i. p i \<oplus> q i)"
   106     proof
   107       fix i
   108       assume "max n m < i"
   109       with boundn and boundm and UP show "p i \<oplus> q i = \<zero>" by fastsimp
   110     qed
   111     then show ?thesis ..
   112   qed
   113 qed
   114 
   115 lemma (in cring) up_a_inv_closed:
   116   "p \<in> up R ==> (%i. \<ominus> (p i)) \<in> up R"
   117 proof
   118   assume R: "p \<in> up R"
   119   then obtain n where "bound \<zero> n p" by auto
   120   then have "bound \<zero> n (%i. \<ominus> p i)" by auto
   121   then show "EX n. bound \<zero> n (%i. \<ominus> p i)" by auto
   122 qed auto
   123 
   124 lemma (in cring) up_mult_closed:
   125   "[| p \<in> up R; q \<in> up R |] ==>
   126   (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> up R"
   127 proof
   128   fix n
   129   assume "p \<in> up R" "q \<in> up R"
   130   then show "(\<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> carrier R"
   131     by (simp add: mem_upD  funcsetI)
   132 next
   133   assume UP: "p \<in> up R" "q \<in> up R"
   134   show "EX n. bound \<zero> n (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i))"
   135   proof -
   136     from UP obtain n where boundn: "bound \<zero> n p" by fast
   137     from UP obtain m where boundm: "bound \<zero> m q" by fast
   138     have "bound \<zero> (n + m) (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n - i))"
   139     proof
   140       fix k assume bound: "n + m < k"
   141       {
   142         fix i
   143         have "p i \<otimes> q (k-i) = \<zero>"
   144         proof (cases "n < i")
   145           case True
   146           with boundn have "p i = \<zero>" by auto
   147           moreover from UP have "q (k-i) \<in> carrier R" by auto
   148           ultimately show ?thesis by simp
   149         next
   150           case False
   151           with bound have "m < k-i" by arith
   152           with boundm have "q (k-i) = \<zero>" by auto
   153           moreover from UP have "p i \<in> carrier R" by auto
   154           ultimately show ?thesis by simp
   155         qed
   156       }
   157       then show "(\<Oplus>i \<in> {..k}. p i \<otimes> q (k-i)) = \<zero>"
   158         by (simp add: Pi_def)
   159     qed
   160     then show ?thesis by fast
   161   qed
   162 qed
   163 
   164 
   165 subsection {* Effect of Operations on Coefficients *}
   166 
   167 locale UP =
   168   fixes R (structure) and P (structure)
   169   defines P_def: "P == UP R"
   170 
   171 locale UP_cring = UP + cring R
   172 
   173 locale UP_domain = UP_cring + "domain" R
   174 
   175 text {*
   176   Temporarily declare @{thm [locale=UP] P_def} as simp rule.
   177 *}
   178 
   179 declare (in UP) P_def [simp]
   180 
   181 lemma (in UP_cring) coeff_monom [simp]:
   182   "a \<in> carrier R ==>
   183   coeff P (monom P a m) n = (if m=n then a else \<zero>)"
   184 proof -
   185   assume R: "a \<in> carrier R"
   186   then have "(%n. if n = m then a else \<zero>) \<in> up R"
   187     using up_def by force
   188   with R show ?thesis by (simp add: UP_def)
   189 qed
   190 
   191 lemma (in UP_cring) coeff_zero [simp]:
   192   "coeff P \<zero>\<^bsub>P\<^esub> n = \<zero>"
   193   by (auto simp add: UP_def)
   194 
   195 lemma (in UP_cring) coeff_one [simp]:
   196   "coeff P \<one>\<^bsub>P\<^esub> n = (if n=0 then \<one> else \<zero>)"
   197   using up_one_closed by (simp add: UP_def)
   198 
   199 lemma (in UP_cring) coeff_smult [simp]:
   200   "[| a \<in> carrier R; p \<in> carrier P |] ==>
   201   coeff P (a \<odot>\<^bsub>P\<^esub> p) n = a \<otimes> coeff P p n"
   202   by (simp add: UP_def up_smult_closed)
   203 
   204 lemma (in UP_cring) coeff_add [simp]:
   205   "[| p \<in> carrier P; q \<in> carrier P |] ==>
   206   coeff P (p \<oplus>\<^bsub>P\<^esub> q) n = coeff P p n \<oplus> coeff P q n"
   207   by (simp add: UP_def up_add_closed)
   208 
   209 lemma (in UP_cring) coeff_mult [simp]:
   210   "[| p \<in> carrier P; q \<in> carrier P |] ==>
   211   coeff P (p \<otimes>\<^bsub>P\<^esub> q) n = (\<Oplus>i \<in> {..n}. coeff P p i \<otimes> coeff P q (n-i))"
   212   by (simp add: UP_def up_mult_closed)
   213 
   214 lemma (in UP) up_eqI:
   215   assumes prem: "!!n. coeff P p n = coeff P q n"
   216     and R: "p \<in> carrier P" "q \<in> carrier P"
   217   shows "p = q"
   218 proof
   219   fix x
   220   from prem and R show "p x = q x" by (simp add: UP_def)
   221 qed
   222 
   223 
   224 subsection {* Polynomials Form a Commutative Ring. *}
   225 
   226 text {* Operations are closed over @{term P}. *}
   227 
   228 lemma (in UP_cring) UP_mult_closed [simp]:
   229   "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<otimes>\<^bsub>P\<^esub> q \<in> carrier P"
   230   by (simp add: UP_def up_mult_closed)
   231 
   232 lemma (in UP_cring) UP_one_closed [simp]:
   233   "\<one>\<^bsub>P\<^esub> \<in> carrier P"
   234   by (simp add: UP_def up_one_closed)
   235 
   236 lemma (in UP_cring) UP_zero_closed [intro, simp]:
   237   "\<zero>\<^bsub>P\<^esub> \<in> carrier P"
   238   by (auto simp add: UP_def)
   239 
   240 lemma (in UP_cring) UP_a_closed [intro, simp]:
   241   "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<oplus>\<^bsub>P\<^esub> q \<in> carrier P"
   242   by (simp add: UP_def up_add_closed)
   243 
   244 lemma (in UP_cring) monom_closed [simp]:
   245   "a \<in> carrier R ==> monom P a n \<in> carrier P"
   246   by (auto simp add: UP_def up_def Pi_def)
   247 
   248 lemma (in UP_cring) UP_smult_closed [simp]:
   249   "[| a \<in> carrier R; p \<in> carrier P |] ==> a \<odot>\<^bsub>P\<^esub> p \<in> carrier P"
   250   by (simp add: UP_def up_smult_closed)
   251 
   252 lemma (in UP) coeff_closed [simp]:
   253   "p \<in> carrier P ==> coeff P p n \<in> carrier R"
   254   by (auto simp add: UP_def)
   255 
   256 declare (in UP) P_def [simp del]
   257 
   258 text {* Algebraic ring properties *}
   259 
   260 lemma (in UP_cring) UP_a_assoc:
   261   assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
   262   shows "(p \<oplus>\<^bsub>P\<^esub> q) \<oplus>\<^bsub>P\<^esub> r = p \<oplus>\<^bsub>P\<^esub> (q \<oplus>\<^bsub>P\<^esub> r)"
   263   by (rule up_eqI, simp add: a_assoc R, simp_all add: R)
   264 
   265 lemma (in UP_cring) UP_l_zero [simp]:
   266   assumes R: "p \<in> carrier P"
   267   shows "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> p = p"
   268   by (rule up_eqI, simp_all add: R)
   269 
   270 lemma (in UP_cring) UP_l_neg_ex:
   271   assumes R: "p \<in> carrier P"
   272   shows "EX q : carrier P. q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
   273 proof -
   274   let ?q = "%i. \<ominus> (p i)"
   275   from R have closed: "?q \<in> carrier P"
   276     by (simp add: UP_def P_def up_a_inv_closed)
   277   from R have coeff: "!!n. coeff P ?q n = \<ominus> (coeff P p n)"
   278     by (simp add: UP_def P_def up_a_inv_closed)
   279   show ?thesis
   280   proof
   281     show "?q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
   282       by (auto intro!: up_eqI simp add: R closed coeff R.l_neg)
   283   qed (rule closed)
   284 qed
   285 
   286 lemma (in UP_cring) UP_a_comm:
   287   assumes R: "p \<in> carrier P" "q \<in> carrier P"
   288   shows "p \<oplus>\<^bsub>P\<^esub> q = q \<oplus>\<^bsub>P\<^esub> p"
   289   by (rule up_eqI, simp add: a_comm R, simp_all add: R)
   290 
   291 lemma (in UP_cring) UP_m_assoc:
   292   assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
   293   shows "(p \<otimes>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r = p \<otimes>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)"
   294 proof (rule up_eqI)
   295   fix n
   296   {
   297     fix k and a b c :: "nat=>'a"
   298     assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
   299       "c \<in> UNIV -> carrier R"
   300     then have "k <= n ==>
   301       (\<Oplus>j \<in> {..k}. (\<Oplus>i \<in> {..j}. a i \<otimes> b (j-i)) \<otimes> c (n-j)) =
   302       (\<Oplus>j \<in> {..k}. a j \<otimes> (\<Oplus>i \<in> {..k-j}. b i \<otimes> c (n-j-i)))"
   303       (is "_ \<Longrightarrow> ?eq k")
   304     proof (induct k)
   305       case 0 then show ?case by (simp add: Pi_def m_assoc)
   306     next
   307       case (Suc k)
   308       then have "k <= n" by arith
   309       from this R have "?eq k" by (rule Suc)
   310       with R show ?case
   311         by (simp cong: finsum_cong
   312              add: Suc_diff_le Pi_def l_distr r_distr m_assoc)
   313           (simp cong: finsum_cong add: Pi_def a_ac finsum_ldistr m_assoc)
   314     qed
   315   }
   316   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"
   317     by (simp add: Pi_def)
   318 qed (simp_all add: R)
   319 
   320 lemma (in UP_cring) UP_l_one [simp]:
   321   assumes R: "p \<in> carrier P"
   322   shows "\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p = p"
   323 proof (rule up_eqI)
   324   fix n
   325   show "coeff P (\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p) n = coeff P p n"
   326   proof (cases n)
   327     case 0 with R show ?thesis by simp
   328   next
   329     case Suc with R show ?thesis
   330       by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)
   331   qed
   332 qed (simp_all add: R)
   333 
   334 lemma (in UP_cring) UP_l_distr:
   335   assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
   336   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)"
   337   by (rule up_eqI) (simp add: l_distr R Pi_def, simp_all add: R)
   338 
   339 lemma (in UP_cring) UP_m_comm:
   340   assumes R: "p \<in> carrier P" "q \<in> carrier P"
   341   shows "p \<otimes>\<^bsub>P\<^esub> q = q \<otimes>\<^bsub>P\<^esub> p"
   342 proof (rule up_eqI)
   343   fix n
   344   {
   345     fix k and a b :: "nat=>'a"
   346     assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
   347     then have "k <= n ==>
   348       (\<Oplus>i \<in> {..k}. a i \<otimes> b (n-i)) =
   349       (\<Oplus>i \<in> {..k}. a (k-i) \<otimes> b (i+n-k))"
   350       (is "_ \<Longrightarrow> ?eq k")
   351     proof (induct k)
   352       case 0 then show ?case by (simp add: Pi_def)
   353     next
   354       case (Suc k) then show ?case
   355         by (subst (2) finsum_Suc2) (simp add: Pi_def a_comm)+
   356     qed
   357   }
   358   note l = this
   359   from R show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) n =  coeff P (q \<otimes>\<^bsub>P\<^esub> p) n"
   360     apply (simp add: Pi_def)
   361     apply (subst l)
   362     apply (auto simp add: Pi_def)
   363     apply (simp add: m_comm)
   364     done
   365 qed (simp_all add: R)
   366 
   367 theorem (in UP_cring) UP_cring:
   368   "cring P"
   369   by (auto intro!: cringI abelian_groupI comm_monoidI UP_a_assoc UP_l_zero
   370     UP_l_neg_ex UP_a_comm UP_m_assoc UP_l_one UP_m_comm UP_l_distr)
   371 
   372 lemma (in UP_cring) UP_ring:
   373   (* preliminary,
   374      we want "UP_ring R P ==> ring P", not "UP_cring R P ==> ring P" *)
   375   "ring P"
   376   by (auto intro: ring.intro cring.axioms UP_cring)
   377 
   378 lemma (in UP_cring) UP_a_inv_closed [intro, simp]:
   379   "p \<in> carrier P ==> \<ominus>\<^bsub>P\<^esub> p \<in> carrier P"
   380   by (rule abelian_group.a_inv_closed
   381     [OF ring.is_abelian_group [OF UP_ring]])
   382 
   383 lemma (in UP_cring) coeff_a_inv [simp]:
   384   assumes R: "p \<in> carrier P"
   385   shows "coeff P (\<ominus>\<^bsub>P\<^esub> p) n = \<ominus> (coeff P p n)"
   386 proof -
   387   from R coeff_closed UP_a_inv_closed have
   388     "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)"
   389     by algebra
   390   also from R have "... =  \<ominus> (coeff P p n)"
   391     by (simp del: coeff_add add: coeff_add [THEN sym]
   392       abelian_group.r_neg [OF ring.is_abelian_group [OF UP_ring]])
   393   finally show ?thesis .
   394 qed
   395 
   396 text {*
   397   Interpretation of lemmas from @{term cring}.  Saves lifting 43
   398   lemmas manually.
   399 *}
   400 
   401 interpretation UP_cring < cring P
   402   by intro_locales
   403     (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms UP_cring)+
   404 
   405 
   406 subsection {* Polynomials Form an Algebra *}
   407 
   408 lemma (in UP_cring) UP_smult_l_distr:
   409   "[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>
   410   (a \<oplus> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p \<oplus>\<^bsub>P\<^esub> b \<odot>\<^bsub>P\<^esub> p"
   411   by (rule up_eqI) (simp_all add: R.l_distr)
   412 
   413 lemma (in UP_cring) UP_smult_r_distr:
   414   "[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>
   415   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"
   416   by (rule up_eqI) (simp_all add: R.r_distr)
   417 
   418 lemma (in UP_cring) UP_smult_assoc1:
   419       "[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>
   420       (a \<otimes> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> p)"
   421   by (rule up_eqI) (simp_all add: R.m_assoc)
   422 
   423 lemma (in UP_cring) UP_smult_one [simp]:
   424       "p \<in> carrier P ==> \<one> \<odot>\<^bsub>P\<^esub> p = p"
   425   by (rule up_eqI) simp_all
   426 
   427 lemma (in UP_cring) UP_smult_assoc2:
   428   "[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>
   429   (a \<odot>\<^bsub>P\<^esub> p) \<otimes>\<^bsub>P\<^esub> q = a \<odot>\<^bsub>P\<^esub> (p \<otimes>\<^bsub>P\<^esub> q)"
   430   by (rule up_eqI) (simp_all add: R.finsum_rdistr R.m_assoc Pi_def)
   431 
   432 text {*
   433   Interpretation of lemmas from @{term algebra}.
   434 *}
   435 
   436 lemma (in cring) cring:
   437   "cring R"
   438   by unfold_locales
   439 
   440 lemma (in UP_cring) UP_algebra:
   441   "algebra R P"
   442   by (auto intro: algebraI R.cring UP_cring UP_smult_l_distr UP_smult_r_distr
   443     UP_smult_assoc1 UP_smult_assoc2)
   444 
   445 interpretation UP_cring < algebra R P
   446   by intro_locales
   447     (rule module.axioms algebra.axioms UP_algebra)+
   448 
   449 
   450 subsection {* Further Lemmas Involving Monomials *}
   451 
   452 lemma (in UP_cring) monom_zero [simp]:
   453   "monom P \<zero> n = \<zero>\<^bsub>P\<^esub>"
   454   by (simp add: UP_def P_def)
   455 
   456 lemma (in UP_cring) monom_mult_is_smult:
   457   assumes R: "a \<in> carrier R" "p \<in> carrier P"
   458   shows "monom P a 0 \<otimes>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p"
   459 proof (rule up_eqI)
   460   fix n
   461   have "coeff P (p \<otimes>\<^bsub>P\<^esub> monom P a 0) n = coeff P (a \<odot>\<^bsub>P\<^esub> p) n"
   462   proof (cases n)
   463     case 0 with R show ?thesis by (simp add: R.m_comm)
   464   next
   465     case Suc with R show ?thesis
   466       by (simp cong: R.finsum_cong add: R.r_null Pi_def)
   467         (simp add: R.m_comm)
   468   qed
   469   with R show "coeff P (monom P a 0 \<otimes>\<^bsub>P\<^esub> p) n = coeff P (a \<odot>\<^bsub>P\<^esub> p) n"
   470     by (simp add: UP_m_comm)
   471 qed (simp_all add: R)
   472 
   473 lemma (in UP_cring) monom_add [simp]:
   474   "[| a \<in> carrier R; b \<in> carrier R |] ==>
   475   monom P (a \<oplus> b) n = monom P a n \<oplus>\<^bsub>P\<^esub> monom P b n"
   476   by (rule up_eqI) simp_all
   477 
   478 lemma (in UP_cring) monom_one_Suc:
   479   "monom P \<one> (Suc n) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
   480 proof (rule up_eqI)
   481   fix k
   482   show "coeff P (monom P \<one> (Suc n)) k = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"
   483   proof (cases "k = Suc n")
   484     case True show ?thesis
   485     proof -
   486       fix m
   487       from True have less_add_diff:
   488         "!!i. [| n < i; i <= n + m |] ==> n + m - i < m" by arith
   489       from True have "coeff P (monom P \<one> (Suc n)) k = \<one>" by simp
   490       also from True
   491       have "... = (\<Oplus>i \<in> {..<n} \<union> {n}. coeff P (monom P \<one> n) i \<otimes>
   492         coeff P (monom P \<one> 1) (k - i))"
   493         by (simp cong: R.finsum_cong add: Pi_def)
   494       also have "... = (\<Oplus>i \<in>  {..n}. coeff P (monom P \<one> n) i \<otimes>
   495         coeff P (monom P \<one> 1) (k - i))"
   496         by (simp only: ivl_disj_un_singleton)
   497       also from True
   498       have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. coeff P (monom P \<one> n) i \<otimes>
   499         coeff P (monom P \<one> 1) (k - i))"
   500         by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one
   501           order_less_imp_not_eq Pi_def)
   502       also from True have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"
   503         by (simp add: ivl_disj_un_one)
   504       finally show ?thesis .
   505     qed
   506   next
   507     case False
   508     note neq = False
   509     let ?s =
   510       "\<lambda>i. (if n = i then \<one> else \<zero>) \<otimes> (if Suc 0 = k - i then \<one> else \<zero>)"
   511     from neq have "coeff P (monom P \<one> (Suc n)) k = \<zero>" by simp
   512     also have "... = (\<Oplus>i \<in> {..k}. ?s i)"
   513     proof -
   514       have f1: "(\<Oplus>i \<in> {..<n}. ?s i) = \<zero>"
   515         by (simp cong: R.finsum_cong add: Pi_def)
   516       from neq have f2: "(\<Oplus>i \<in> {n}. ?s i) = \<zero>"
   517         by (simp cong: R.finsum_cong add: Pi_def) arith
   518       have f3: "n < k ==> (\<Oplus>i \<in> {n<..k}. ?s i) = \<zero>"
   519         by (simp cong: R.finsum_cong add: order_less_imp_not_eq Pi_def)
   520       show ?thesis
   521       proof (cases "k < n")
   522         case True then show ?thesis by (simp cong: R.finsum_cong add: Pi_def)
   523       next
   524         case False then have n_le_k: "n <= k" by arith
   525         show ?thesis
   526         proof (cases "n = k")
   527           case True
   528           then have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
   529             by (simp cong: R.finsum_cong add: ivl_disj_int_singleton Pi_def)
   530           also from True have "... = (\<Oplus>i \<in> {..k}. ?s i)"
   531             by (simp only: ivl_disj_un_singleton)
   532           finally show ?thesis .
   533         next
   534           case False with n_le_k have n_less_k: "n < k" by arith
   535           with neq have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
   536             by (simp add: R.finsum_Un_disjoint f1 f2
   537               ivl_disj_int_singleton Pi_def del: Un_insert_right)
   538           also have "... = (\<Oplus>i \<in> {..n}. ?s i)"
   539             by (simp only: ivl_disj_un_singleton)
   540           also from n_less_k neq have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. ?s i)"
   541             by (simp add: R.finsum_Un_disjoint f3 ivl_disj_int_one Pi_def)
   542           also from n_less_k have "... = (\<Oplus>i \<in> {..k}. ?s i)"
   543             by (simp only: ivl_disj_un_one)
   544           finally show ?thesis .
   545         qed
   546       qed
   547     qed
   548     also have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k" by simp
   549     finally show ?thesis .
   550   qed
   551 qed (simp_all)
   552 
   553 lemma (in UP_cring) monom_mult_smult:
   554   "[| a \<in> carrier R; b \<in> carrier R |] ==> monom P (a \<otimes> b) n = a \<odot>\<^bsub>P\<^esub> monom P b n"
   555   by (rule up_eqI) simp_all
   556 
   557 lemma (in UP_cring) monom_one [simp]:
   558   "monom P \<one> 0 = \<one>\<^bsub>P\<^esub>"
   559   by (rule up_eqI) simp_all
   560 
   561 lemma (in UP_cring) monom_one_mult:
   562   "monom P \<one> (n + m) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m"
   563 proof (induct n)
   564   case 0 show ?case by simp
   565 next
   566   case Suc then show ?case
   567     by (simp only: add_Suc monom_one_Suc) (simp add: P.m_ac)
   568 qed
   569 
   570 lemma (in UP_cring) monom_mult [simp]:
   571   assumes R: "a \<in> carrier R" "b \<in> carrier R"
   572   shows "monom P (a \<otimes> b) (n + m) = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m"
   573 proof -
   574   from R have "monom P (a \<otimes> b) (n + m) = monom P (a \<otimes> b \<otimes> \<one>) (n + m)" by simp
   575   also from R have "... = a \<otimes> b \<odot>\<^bsub>P\<^esub> monom P \<one> (n + m)"
   576     by (simp add: monom_mult_smult del: R.r_one)
   577   also have "... = a \<otimes> b \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m)"
   578     by (simp only: monom_one_mult)
   579   also from R have "... = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m))"
   580     by (simp add: UP_smult_assoc1)
   581   also from R have "... = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> (monom P \<one> m \<otimes>\<^bsub>P\<^esub> monom P \<one> n))"
   582     by (simp add: P.m_comm)
   583   also from R have "... = a \<odot>\<^bsub>P\<^esub> ((b \<odot>\<^bsub>P\<^esub> monom P \<one> m) \<otimes>\<^bsub>P\<^esub> monom P \<one> n)"
   584     by (simp add: UP_smult_assoc2)
   585   also from R have "... = a \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> monom P \<one> m))"
   586     by (simp add: P.m_comm)
   587   also from R have "... = (a \<odot>\<^bsub>P\<^esub> monom P \<one> n) \<otimes>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> monom P \<one> m)"
   588     by (simp add: UP_smult_assoc2)
   589   also from R have "... = monom P (a \<otimes> \<one>) n \<otimes>\<^bsub>P\<^esub> monom P (b \<otimes> \<one>) m"
   590     by (simp add: monom_mult_smult del: R.r_one)
   591   also from R have "... = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m" by simp
   592   finally show ?thesis .
   593 qed
   594 
   595 lemma (in UP_cring) monom_a_inv [simp]:
   596   "a \<in> carrier R ==> monom P (\<ominus> a) n = \<ominus>\<^bsub>P\<^esub> monom P a n"
   597   by (rule up_eqI) simp_all
   598 
   599 lemma (in UP_cring) monom_inj:
   600   "inj_on (%a. monom P a n) (carrier R)"
   601 proof (rule inj_onI)
   602   fix x y
   603   assume R: "x \<in> carrier R" "y \<in> carrier R" and eq: "monom P x n = monom P y n"
   604   then have "coeff P (monom P x n) n = coeff P (monom P y n) n" by simp
   605   with R show "x = y" by simp
   606 qed
   607 
   608 
   609 subsection {* The Degree Function *}
   610 
   611 constdefs (structure R)
   612   deg :: "[('a, 'm) ring_scheme, nat => 'a] => nat"
   613   "deg R p == LEAST n. bound \<zero> n (coeff (UP R) p)"
   614 
   615 lemma (in UP_cring) deg_aboveI:
   616   "[| (!!m. n < m ==> coeff P p m = \<zero>); p \<in> carrier P |] ==> deg R p <= n"
   617   by (unfold deg_def P_def) (fast intro: Least_le)
   618 
   619 (*
   620 lemma coeff_bound_ex: "EX n. bound n (coeff p)"
   621 proof -
   622   have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
   623   then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
   624   then show ?thesis ..
   625 qed
   626 
   627 lemma bound_coeff_obtain:
   628   assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P"
   629 proof -
   630   have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
   631   then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
   632   with prem show P .
   633 qed
   634 *)
   635 
   636 lemma (in UP_cring) deg_aboveD:
   637   assumes "deg R p < m" and "p \<in> carrier P"
   638   shows "coeff P p m = \<zero>"
   639 proof -
   640   from `p \<in> carrier P` obtain n where "bound \<zero> n (coeff P p)"
   641     by (auto simp add: UP_def P_def)
   642   then have "bound \<zero> (deg R p) (coeff P p)"
   643     by (auto simp: deg_def P_def dest: LeastI)
   644   from this and `deg R p < m` show ?thesis ..
   645 qed
   646 
   647 lemma (in UP_cring) deg_belowI:
   648   assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"
   649     and R: "p \<in> carrier P"
   650   shows "n <= deg R p"
   651 -- {* Logically, this is a slightly stronger version of
   652    @{thm [source] deg_aboveD} *}
   653 proof (cases "n=0")
   654   case True then show ?thesis by simp
   655 next
   656   case False then have "coeff P p n ~= \<zero>" by (rule non_zero)
   657   then have "~ deg R p < n" by (fast dest: deg_aboveD intro: R)
   658   then show ?thesis by arith
   659 qed
   660 
   661 lemma (in UP_cring) lcoeff_nonzero_deg:
   662   assumes deg: "deg R p ~= 0" and R: "p \<in> carrier P"
   663   shows "coeff P p (deg R p) ~= \<zero>"
   664 proof -
   665   from R obtain m where "deg R p <= m" and m_coeff: "coeff P p m ~= \<zero>"
   666   proof -
   667     have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"
   668       by arith
   669 (* TODO: why does simplification below not work with "1" *)
   670     from deg have "deg R p - 1 < (LEAST n. bound \<zero> n (coeff P p))"
   671       by (unfold deg_def P_def) arith
   672     then have "~ bound \<zero> (deg R p - 1) (coeff P p)" by (rule not_less_Least)
   673     then have "EX m. deg R p - 1 < m & coeff P p m ~= \<zero>"
   674       by (unfold bound_def) fast
   675     then have "EX m. deg R p <= m & coeff P p m ~= \<zero>" by (simp add: deg minus)
   676     then show ?thesis by (auto intro: that)
   677   qed
   678   with deg_belowI R have "deg R p = m" by fastsimp
   679   with m_coeff show ?thesis by simp
   680 qed
   681 
   682 lemma (in UP_cring) lcoeff_nonzero_nonzero:
   683   assumes deg: "deg R p = 0" and nonzero: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
   684   shows "coeff P p 0 ~= \<zero>"
   685 proof -
   686   have "EX m. coeff P p m ~= \<zero>"
   687   proof (rule classical)
   688     assume "~ ?thesis"
   689     with R have "p = \<zero>\<^bsub>P\<^esub>" by (auto intro: up_eqI)
   690     with nonzero show ?thesis by contradiction
   691   qed
   692   then obtain m where coeff: "coeff P p m ~= \<zero>" ..
   693   from this and R have "m <= deg R p" by (rule deg_belowI)
   694   then have "m = 0" by (simp add: deg)
   695   with coeff show ?thesis by simp
   696 qed
   697 
   698 lemma (in UP_cring) lcoeff_nonzero:
   699   assumes neq: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
   700   shows "coeff P p (deg R p) ~= \<zero>"
   701 proof (cases "deg R p = 0")
   702   case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero)
   703 next
   704   case False with neq R show ?thesis by (simp add: lcoeff_nonzero_deg)
   705 qed
   706 
   707 lemma (in UP_cring) deg_eqI:
   708   "[| !!m. n < m ==> coeff P p m = \<zero>;
   709       !!n. n ~= 0 ==> coeff P p n ~= \<zero>; p \<in> carrier P |] ==> deg R p = n"
   710 by (fast intro: le_anti_sym deg_aboveI deg_belowI)
   711 
   712 text {* Degree and polynomial operations *}
   713 
   714 lemma (in UP_cring) deg_add [simp]:
   715   assumes R: "p \<in> carrier P" "q \<in> carrier P"
   716   shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
   717 proof (cases "deg R p <= deg R q")
   718   case True show ?thesis
   719     by (rule deg_aboveI) (simp_all add: True R deg_aboveD)
   720 next
   721   case False show ?thesis
   722     by (rule deg_aboveI) (simp_all add: False R deg_aboveD)
   723 qed
   724 
   725 lemma (in UP_cring) deg_monom_le:
   726   "a \<in> carrier R ==> deg R (monom P a n) <= n"
   727   by (intro deg_aboveI) simp_all
   728 
   729 lemma (in UP_cring) deg_monom [simp]:
   730   "[| a ~= \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"
   731   by (fastsimp intro: le_anti_sym deg_aboveI deg_belowI)
   732 
   733 lemma (in UP_cring) deg_const [simp]:
   734   assumes R: "a \<in> carrier R" shows "deg R (monom P a 0) = 0"
   735 proof (rule le_anti_sym)
   736   show "deg R (monom P a 0) <= 0" by (rule deg_aboveI) (simp_all add: R)
   737 next
   738   show "0 <= deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R)
   739 qed
   740 
   741 lemma (in UP_cring) deg_zero [simp]:
   742   "deg R \<zero>\<^bsub>P\<^esub> = 0"
   743 proof (rule le_anti_sym)
   744   show "deg R \<zero>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
   745 next
   746   show "0 <= deg R \<zero>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
   747 qed
   748 
   749 lemma (in UP_cring) deg_one [simp]:
   750   "deg R \<one>\<^bsub>P\<^esub> = 0"
   751 proof (rule le_anti_sym)
   752   show "deg R \<one>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
   753 next
   754   show "0 <= deg R \<one>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
   755 qed
   756 
   757 lemma (in UP_cring) deg_uminus [simp]:
   758   assumes R: "p \<in> carrier P" shows "deg R (\<ominus>\<^bsub>P\<^esub> p) = deg R p"
   759 proof (rule le_anti_sym)
   760   show "deg R (\<ominus>\<^bsub>P\<^esub> p) <= deg R p" by (simp add: deg_aboveI deg_aboveD R)
   761 next
   762   show "deg R p <= deg R (\<ominus>\<^bsub>P\<^esub> p)"
   763     by (simp add: deg_belowI lcoeff_nonzero_deg
   764       inj_on_iff [OF R.a_inv_inj, of _ "\<zero>", simplified] R)
   765 qed
   766 
   767 lemma (in UP_domain) deg_smult_ring:
   768   "[| a \<in> carrier R; p \<in> carrier P |] ==>
   769   deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
   770   by (cases "a = \<zero>") (simp add: deg_aboveI deg_aboveD)+
   771 
   772 lemma (in UP_domain) deg_smult [simp]:
   773   assumes R: "a \<in> carrier R" "p \<in> carrier P"
   774   shows "deg R (a \<odot>\<^bsub>P\<^esub> p) = (if a = \<zero> then 0 else deg R p)"
   775 proof (rule le_anti_sym)
   776   show "deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
   777     using R by (rule deg_smult_ring)
   778 next
   779   show "(if a = \<zero> then 0 else deg R p) <= deg R (a \<odot>\<^bsub>P\<^esub> p)"
   780   proof (cases "a = \<zero>")
   781   qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R)
   782 qed
   783 
   784 lemma (in UP_cring) deg_mult_cring:
   785   assumes R: "p \<in> carrier P" "q \<in> carrier P"
   786   shows "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q"
   787 proof (rule deg_aboveI)
   788   fix m
   789   assume boundm: "deg R p + deg R q < m"
   790   {
   791     fix k i
   792     assume boundk: "deg R p + deg R q < k"
   793     then have "coeff P p i \<otimes> coeff P q (k - i) = \<zero>"
   794     proof (cases "deg R p < i")
   795       case True then show ?thesis by (simp add: deg_aboveD R)
   796     next
   797       case False with boundk have "deg R q < k - i" by arith
   798       then show ?thesis by (simp add: deg_aboveD R)
   799     qed
   800   }
   801   with boundm R show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) m = \<zero>" by simp
   802 qed (simp add: R)
   803 
   804 lemma (in UP_domain) deg_mult [simp]:
   805   "[| p ~= \<zero>\<^bsub>P\<^esub>; q ~= \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==>
   806   deg R (p \<otimes>\<^bsub>P\<^esub> q) = deg R p + deg R q"
   807 proof (rule le_anti_sym)
   808   assume "p \<in> carrier P" " q \<in> carrier P"
   809   then show "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_cring)
   810 next
   811   let ?s = "(%i. coeff P p i \<otimes> coeff P q (deg R p + deg R q - i))"
   812   assume R: "p \<in> carrier P" "q \<in> carrier P" and nz: "p ~= \<zero>\<^bsub>P\<^esub>" "q ~= \<zero>\<^bsub>P\<^esub>"
   813   have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith
   814   show "deg R p + deg R q <= deg R (p \<otimes>\<^bsub>P\<^esub> q)"
   815   proof (rule deg_belowI, simp add: R)
   816     have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
   817       = (\<Oplus>i \<in> {..< deg R p} \<union> {deg R p .. deg R p + deg R q}. ?s i)"
   818       by (simp only: ivl_disj_un_one)
   819     also have "... = (\<Oplus>i \<in> {deg R p .. deg R p + deg R q}. ?s i)"
   820       by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one
   821         deg_aboveD less_add_diff R Pi_def)
   822     also have "...= (\<Oplus>i \<in> {deg R p} \<union> {deg R p <.. deg R p + deg R q}. ?s i)"
   823       by (simp only: ivl_disj_un_singleton)
   824     also have "... = coeff P p (deg R p) \<otimes> coeff P q (deg R q)"
   825       by (simp cong: R.finsum_cong
   826 	add: ivl_disj_int_singleton deg_aboveD R Pi_def)
   827     finally have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
   828       = coeff P p (deg R p) \<otimes> coeff P q (deg R q)" .
   829     with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) ~= \<zero>"
   830       by (simp add: integral_iff lcoeff_nonzero R)
   831     qed (simp add: R)
   832   qed
   833 
   834 lemma (in UP_cring) coeff_finsum:
   835   assumes fin: "finite A"
   836   shows "p \<in> A -> carrier P ==>
   837     coeff P (finsum P p A) k = (\<Oplus>i \<in> A. coeff P (p i) k)"
   838   using fin by induct (auto simp: Pi_def)
   839 
   840 lemma (in UP_cring) up_repr:
   841   assumes R: "p \<in> carrier P"
   842   shows "(\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. monom P (coeff P p i) i) = p"
   843 proof (rule up_eqI)
   844   let ?s = "(%i. monom P (coeff P p i) i)"
   845   fix k
   846   from R have RR: "!!i. (if i = k then coeff P p i else \<zero>) \<in> carrier R"
   847     by simp
   848   show "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k = coeff P p k"
   849   proof (cases "k <= deg R p")
   850     case True
   851     hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
   852           coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k} \<union> {k<..deg R p}. ?s i) k"
   853       by (simp only: ivl_disj_un_one)
   854     also from True
   855     have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k}. ?s i) k"
   856       by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint
   857         ivl_disj_int_one order_less_imp_not_eq2 coeff_finsum R RR Pi_def)
   858     also
   859     have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<k} \<union> {k}. ?s i) k"
   860       by (simp only: ivl_disj_un_singleton)
   861     also have "... = coeff P p k"
   862       by (simp cong: R.finsum_cong
   863 	add: ivl_disj_int_singleton coeff_finsum deg_aboveD R RR Pi_def)
   864     finally show ?thesis .
   865   next
   866     case False
   867     hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
   868           coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<deg R p} \<union> {deg R p}. ?s i) k"
   869       by (simp only: ivl_disj_un_singleton)
   870     also from False have "... = coeff P p k"
   871       by (simp cong: R.finsum_cong
   872 	add: ivl_disj_int_singleton coeff_finsum deg_aboveD R Pi_def)
   873     finally show ?thesis .
   874   qed
   875 qed (simp_all add: R Pi_def)
   876 
   877 lemma (in UP_cring) up_repr_le:
   878   "[| deg R p <= n; p \<in> carrier P |] ==>
   879   (\<Oplus>\<^bsub>P\<^esub> i \<in> {..n}. monom P (coeff P p i) i) = p"
   880 proof -
   881   let ?s = "(%i. monom P (coeff P p i) i)"
   882   assume R: "p \<in> carrier P" and "deg R p <= n"
   883   then have "finsum P ?s {..n} = finsum P ?s ({..deg R p} \<union> {deg R p<..n})"
   884     by (simp only: ivl_disj_un_one)
   885   also have "... = finsum P ?s {..deg R p}"
   886     by (simp cong: P.finsum_cong add: P.finsum_Un_disjoint ivl_disj_int_one
   887       deg_aboveD R Pi_def)
   888   also have "... = p" using R by (rule up_repr)
   889   finally show ?thesis .
   890 qed
   891 
   892 
   893 subsection {* Polynomials over Integral Domains *}
   894 
   895 lemma domainI:
   896   assumes cring: "cring R"
   897     and one_not_zero: "one R ~= zero R"
   898     and integral: "!!a b. [| mult R a b = zero R; a \<in> carrier R;
   899       b \<in> carrier R |] ==> a = zero R | b = zero R"
   900   shows "domain R"
   901   by (auto intro!: domain.intro domain_axioms.intro cring.axioms assms
   902     del: disjCI)
   903 
   904 lemma (in UP_domain) UP_one_not_zero:
   905   "\<one>\<^bsub>P\<^esub> ~= \<zero>\<^bsub>P\<^esub>"
   906 proof
   907   assume "\<one>\<^bsub>P\<^esub> = \<zero>\<^bsub>P\<^esub>"
   908   hence "coeff P \<one>\<^bsub>P\<^esub> 0 = (coeff P \<zero>\<^bsub>P\<^esub> 0)" by simp
   909   hence "\<one> = \<zero>" by simp
   910   with one_not_zero show "False" by contradiction
   911 qed
   912 
   913 lemma (in UP_domain) UP_integral:
   914   "[| p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==> p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
   915 proof -
   916   fix p q
   917   assume pq: "p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P" "q \<in> carrier P"
   918   show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
   919   proof (rule classical)
   920     assume c: "~ (p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>)"
   921     with R have "deg R p + deg R q = deg R (p \<otimes>\<^bsub>P\<^esub> q)" by simp
   922     also from pq have "... = 0" by simp
   923     finally have "deg R p + deg R q = 0" .
   924     then have f1: "deg R p = 0 & deg R q = 0" by simp
   925     from f1 R have "p = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P p i) i)"
   926       by (simp only: up_repr_le)
   927     also from R have "... = monom P (coeff P p 0) 0" by simp
   928     finally have p: "p = monom P (coeff P p 0) 0" .
   929     from f1 R have "q = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P q i) i)"
   930       by (simp only: up_repr_le)
   931     also from R have "... = monom P (coeff P q 0) 0" by simp
   932     finally have q: "q = monom P (coeff P q 0) 0" .
   933     from R have "coeff P p 0 \<otimes> coeff P q 0 = coeff P (p \<otimes>\<^bsub>P\<^esub> q) 0" by simp
   934     also from pq have "... = \<zero>" by simp
   935     finally have "coeff P p 0 \<otimes> coeff P q 0 = \<zero>" .
   936     with R have "coeff P p 0 = \<zero> | coeff P q 0 = \<zero>"
   937       by (simp add: R.integral_iff)
   938     with p q show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>" by fastsimp
   939   qed
   940 qed
   941 
   942 theorem (in UP_domain) UP_domain:
   943   "domain P"
   944   by (auto intro!: domainI UP_cring UP_one_not_zero UP_integral del: disjCI)
   945 
   946 text {*
   947   Interpretation of theorems from @{term domain}.
   948 *}
   949 
   950 interpretation UP_domain < "domain" P
   951   by intro_locales (rule domain.axioms UP_domain)+
   952 
   953 
   954 subsection {* The Evaluation Homomorphism and Universal Property*}
   955 
   956 (* alternative congruence rule (possibly more efficient)
   957 lemma (in abelian_monoid) finsum_cong2:
   958   "[| !!i. i \<in> A ==> f i \<in> carrier G = True; A = B;
   959   !!i. i \<in> B ==> f i = g i |] ==> finsum G f A = finsum G g B"
   960   sorry*)
   961 
   962 theorem (in cring) diagonal_sum:
   963   "[| f \<in> {..n + m::nat} -> carrier R; g \<in> {..n + m} -> carrier R |] ==>
   964   (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
   965   (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
   966 proof -
   967   assume Rf: "f \<in> {..n + m} -> carrier R" and Rg: "g \<in> {..n + m} -> carrier R"
   968   {
   969     fix j
   970     have "j <= n + m ==>
   971       (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
   972       (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..j - k}. f k \<otimes> g i)"
   973     proof (induct j)
   974       case 0 from Rf Rg show ?case by (simp add: Pi_def)
   975     next
   976       case (Suc j)
   977       have R6: "!!i k. [| k <= j; i <= Suc j - k |] ==> g i \<in> carrier R"
   978         using Suc by (auto intro!: funcset_mem [OF Rg])
   979       have R8: "!!i k. [| k <= Suc j; i <= k |] ==> g (k - i) \<in> carrier R"
   980         using Suc by (auto intro!: funcset_mem [OF Rg])
   981       have R9: "!!i k. [| k <= Suc j |] ==> f k \<in> carrier R"
   982         using Suc by (auto intro!: funcset_mem [OF Rf])
   983       have R10: "!!i k. [| k <= Suc j; i <= Suc j - k |] ==> g i \<in> carrier R"
   984         using Suc by (auto intro!: funcset_mem [OF Rg])
   985       have R11: "g 0 \<in> carrier R"
   986         using Suc by (auto intro!: funcset_mem [OF Rg])
   987       from Suc show ?case
   988         by (simp cong: finsum_cong add: Suc_diff_le a_ac
   989           Pi_def R6 R8 R9 R10 R11)
   990     qed
   991   }
   992   then show ?thesis by fast
   993 qed
   994 
   995 lemma (in abelian_monoid) boundD_carrier:
   996   "[| bound \<zero> n f; n < m |] ==> f m \<in> carrier G"
   997   by auto
   998 
   999 theorem (in cring) cauchy_product:
  1000   assumes bf: "bound \<zero> n f" and bg: "bound \<zero> m g"
  1001     and Rf: "f \<in> {..n} -> carrier R" and Rg: "g \<in> {..m} -> carrier R"
  1002   shows "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
  1003     (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"      (* State reverse direction? *)
  1004 proof -
  1005   have f: "!!x. f x \<in> carrier R"
  1006   proof -
  1007     fix x
  1008     show "f x \<in> carrier R"
  1009       using Rf bf boundD_carrier by (cases "x <= n") (auto simp: Pi_def)
  1010   qed
  1011   have g: "!!x. g x \<in> carrier R"
  1012   proof -
  1013     fix x
  1014     show "g x \<in> carrier R"
  1015       using Rg bg boundD_carrier by (cases "x <= m") (auto simp: Pi_def)
  1016   qed
  1017   from f g have "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
  1018       (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
  1019     by (simp add: diagonal_sum Pi_def)
  1020   also have "... = (\<Oplus>k \<in> {..n} \<union> {n<..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
  1021     by (simp only: ivl_disj_un_one)
  1022   also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
  1023     by (simp cong: finsum_cong
  1024       add: bound.bound [OF bf] finsum_Un_disjoint ivl_disj_int_one Pi_def)
  1025   also from f g
  1026   have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m} \<union> {m<..n + m - k}. f k \<otimes> g i)"
  1027     by (simp cong: finsum_cong add: ivl_disj_un_one le_add_diff Pi_def)
  1028   also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m}. f k \<otimes> g i)"
  1029     by (simp cong: finsum_cong
  1030       add: bound.bound [OF bg] finsum_Un_disjoint ivl_disj_int_one Pi_def)
  1031   also from f g have "... = (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"
  1032     by (simp add: finsum_ldistr diagonal_sum Pi_def,
  1033       simp cong: finsum_cong add: finsum_rdistr Pi_def)
  1034   finally show ?thesis .
  1035 qed
  1036 
  1037 lemma (in UP_cring) const_ring_hom:
  1038   "(%a. monom P a 0) \<in> ring_hom R P"
  1039   by (auto intro!: ring_hom_memI intro: up_eqI simp: monom_mult_is_smult)
  1040 
  1041 constdefs (structure S)
  1042   eval :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme,
  1043            'a => 'b, 'b, nat => 'a] => 'b"
  1044   "eval R S phi s == \<lambda>p \<in> carrier (UP R).
  1045     \<Oplus>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes> s (^) i"
  1046 
  1047 
  1048 lemma (in UP) eval_on_carrier:
  1049   fixes S (structure)
  1050   shows "p \<in> carrier P ==>
  1051   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)"
  1052   by (unfold eval_def, fold P_def) simp
  1053 
  1054 lemma (in UP) eval_extensional:
  1055   "eval R S phi p \<in> extensional (carrier P)"
  1056   by (unfold eval_def, fold P_def) simp
  1057 
  1058 
  1059 text {* The universal property of the polynomial ring *}
  1060 
  1061 locale UP_pre_univ_prop = ring_hom_cring R S h + UP_cring R P
  1062 
  1063 locale UP_univ_prop = UP_pre_univ_prop +
  1064   fixes s and Eval
  1065   assumes indet_img_carrier [simp, intro]: "s \<in> carrier S"
  1066   defines Eval_def: "Eval == eval R S h s"
  1067 
  1068 theorem (in UP_pre_univ_prop) eval_ring_hom:
  1069   assumes S: "s \<in> carrier S"
  1070   shows "eval R S h s \<in> ring_hom P S"
  1071 proof (rule ring_hom_memI)
  1072   fix p
  1073   assume R: "p \<in> carrier P"
  1074   then show "eval R S h s p \<in> carrier S"
  1075     by (simp only: eval_on_carrier) (simp add: S Pi_def)
  1076 next
  1077   fix p q
  1078   assume R: "p \<in> carrier P" "q \<in> carrier P"
  1079   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"
  1080   proof (simp only: eval_on_carrier UP_mult_closed)
  1081     from R S have
  1082       "(\<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) =
  1083       (\<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}.
  1084         h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1085       by (simp cong: S.finsum_cong
  1086         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def
  1087         del: coeff_mult)
  1088     also from R have "... =
  1089       (\<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)"
  1090       by (simp only: ivl_disj_un_one deg_mult_cring)
  1091     also from R S have "... =
  1092       (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}.
  1093          \<Oplus>\<^bsub>S\<^esub> k \<in> {..i}.
  1094            h (coeff P p k) \<otimes>\<^bsub>S\<^esub> h (coeff P q (i - k)) \<otimes>\<^bsub>S\<^esub>
  1095            (s (^)\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> (i - k)))"
  1096       by (simp cong: S.finsum_cong add: S.nat_pow_mult Pi_def
  1097         S.m_ac S.finsum_rdistr)
  1098     also from R S have "... =
  1099       (\<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>
  1100       (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1101       by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac
  1102         Pi_def)
  1103     finally show
  1104       "(\<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) =
  1105       (\<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>
  1106       (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
  1107   qed
  1108 next
  1109   fix p q
  1110   assume R: "p \<in> carrier P" "q \<in> carrier P"
  1111   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"
  1112   proof (simp only: eval_on_carrier P.a_closed)
  1113     from S R have
  1114       "(\<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) =
  1115       (\<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)}.
  1116         h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1117       by (simp cong: S.finsum_cong
  1118         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def
  1119         del: coeff_add)
  1120     also from R have "... =
  1121         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..max (deg R p) (deg R q)}.
  1122           h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1123       by (simp add: ivl_disj_un_one)
  1124     also from R S have "... =
  1125       (\<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>
  1126       (\<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)"
  1127       by (simp cong: S.finsum_cong
  1128         add: S.l_distr deg_aboveD ivl_disj_int_one Pi_def)
  1129     also have "... =
  1130         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p} \<union> {deg R p<..max (deg R p) (deg R q)}.
  1131           h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
  1132         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q} \<union> {deg R q<..max (deg R p) (deg R q)}.
  1133           h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1134       by (simp only: ivl_disj_un_one le_maxI1 le_maxI2)
  1135     also from R S have "... =
  1136       (\<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>
  1137       (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1138       by (simp cong: S.finsum_cong
  1139         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
  1140     finally show
  1141       "(\<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) =
  1142       (\<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>
  1143       (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
  1144   qed
  1145 next
  1146   show "eval R S h s \<one>\<^bsub>P\<^esub> = \<one>\<^bsub>S\<^esub>"
  1147     by (simp only: eval_on_carrier UP_one_closed) simp
  1148 qed
  1149 
  1150 text {* Interpretation of ring homomorphism lemmas. *}
  1151 
  1152 interpretation UP_univ_prop < ring_hom_cring P S Eval
  1153   apply (unfold Eval_def)
  1154   apply intro_locales
  1155   apply (rule ring_hom_cring.axioms)
  1156   apply (rule ring_hom_cring.intro)
  1157   apply unfold_locales
  1158   apply (rule eval_ring_hom)
  1159   apply rule
  1160   done
  1161 
  1162 
  1163 text {* Further properties of the evaluation homomorphism. *}
  1164 
  1165 text {*
  1166   The following lemma could be proved in @{text UP_cring} with the additional
  1167   assumption that @{text h} is closed. *}
  1168 
  1169 lemma (in UP_pre_univ_prop) eval_const:
  1170   "[| s \<in> carrier S; r \<in> carrier R |] ==> eval R S h s (monom P r 0) = h r"
  1171   by (simp only: eval_on_carrier monom_closed) simp
  1172 
  1173 text {* The following proof is complicated by the fact that in arbitrary
  1174   rings one might have @{term "one R = zero R"}. *}
  1175 
  1176 (* TODO: simplify by cases "one R = zero R" *)
  1177 
  1178 lemma (in UP_pre_univ_prop) eval_monom1:
  1179   assumes S: "s \<in> carrier S"
  1180   shows "eval R S h s (monom P \<one> 1) = s"
  1181 proof (simp only: eval_on_carrier monom_closed R.one_closed)
  1182    from S have
  1183     "(\<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) =
  1184     (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)} \<union> {deg R (monom P \<one> 1)<..1}.
  1185       h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1186     by (simp cong: S.finsum_cong del: coeff_monom
  1187       add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
  1188   also have "... =
  1189     (\<Oplus>\<^bsub>S\<^esub> i \<in> {..1}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
  1190     by (simp only: ivl_disj_un_one deg_monom_le R.one_closed)
  1191   also have "... = s"
  1192   proof (cases "s = \<zero>\<^bsub>S\<^esub>")
  1193     case True then show ?thesis by (simp add: Pi_def)
  1194   next
  1195     case False then show ?thesis by (simp add: S Pi_def)
  1196   qed
  1197   finally show "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (monom P \<one> 1)}.
  1198     h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) = s" .
  1199 qed
  1200 
  1201 lemma (in UP_cring) monom_pow:
  1202   assumes R: "a \<in> carrier R"
  1203   shows "(monom P a n) (^)\<^bsub>P\<^esub> m = monom P (a (^) m) (n * m)"
  1204 proof (induct m)
  1205   case 0 from R show ?case by simp
  1206 next
  1207   case Suc with R show ?case
  1208     by (simp del: monom_mult add: monom_mult [THEN sym] add_commute)
  1209 qed
  1210 
  1211 lemma (in ring_hom_cring) hom_pow [simp]:
  1212   "x \<in> carrier R ==> h (x (^) n) = h x (^)\<^bsub>S\<^esub> (n::nat)"
  1213   by (induct n) simp_all
  1214 
  1215 lemma (in UP_univ_prop) Eval_monom:
  1216   "r \<in> carrier R ==> Eval (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
  1217 proof -
  1218   assume R: "r \<in> carrier R"
  1219   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)"
  1220     by (simp del: monom_mult add: monom_mult [THEN sym] monom_pow)
  1221   also
  1222   from R eval_monom1 [where s = s, folded Eval_def]
  1223   have "... = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
  1224     by (simp add: eval_const [where s = s, folded Eval_def])
  1225   finally show ?thesis .
  1226 qed
  1227 
  1228 lemma (in UP_pre_univ_prop) eval_monom:
  1229   assumes R: "r \<in> carrier R" and S: "s \<in> carrier S"
  1230   shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
  1231 proof -
  1232   interpret UP_univ_prop [R S h P s _]
  1233     using UP_pre_univ_prop_axioms P_def R S
  1234     by (auto intro: UP_univ_prop.intro UP_univ_prop_axioms.intro)
  1235   from R
  1236   show ?thesis by (rule Eval_monom)
  1237 qed
  1238 
  1239 lemma (in UP_univ_prop) Eval_smult:
  1240   "[| r \<in> carrier R; p \<in> carrier P |] ==> Eval (r \<odot>\<^bsub>P\<^esub> p) = h r \<otimes>\<^bsub>S\<^esub> Eval p"
  1241 proof -
  1242   assume R: "r \<in> carrier R" and P: "p \<in> carrier P"
  1243   then show ?thesis
  1244     by (simp add: monom_mult_is_smult [THEN sym]
  1245       eval_const [where s = s, folded Eval_def])
  1246 qed
  1247 
  1248 lemma ring_hom_cringI:
  1249   assumes "cring R"
  1250     and "cring S"
  1251     and "h \<in> ring_hom R S"
  1252   shows "ring_hom_cring R S h"
  1253   by (fast intro: ring_hom_cring.intro ring_hom_cring_axioms.intro
  1254     cring.axioms assms)
  1255 
  1256 lemma (in UP_pre_univ_prop) UP_hom_unique:
  1257   assumes "ring_hom_cring P S Phi"
  1258   assumes Phi: "Phi (monom P \<one> (Suc 0)) = s"
  1259       "!!r. r \<in> carrier R ==> Phi (monom P r 0) = h r"
  1260   assumes "ring_hom_cring P S Psi"
  1261   assumes Psi: "Psi (monom P \<one> (Suc 0)) = s"
  1262       "!!r. r \<in> carrier R ==> Psi (monom P r 0) = h r"
  1263     and P: "p \<in> carrier P" and S: "s \<in> carrier S"
  1264   shows "Phi p = Psi p"
  1265 proof -
  1266   interpret ring_hom_cring [P S Phi] by fact
  1267   interpret ring_hom_cring [P S Psi] by fact
  1268   have "Phi p =
  1269       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)"
  1270     by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
  1271   also
  1272   have "... =
  1273       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)"
  1274     by (simp add: Phi Psi P Pi_def comp_def)
  1275   also have "... = Psi p"
  1276     by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
  1277   finally show ?thesis .
  1278 qed
  1279 
  1280 lemma (in UP_pre_univ_prop) ring_homD:
  1281   assumes Phi: "Phi \<in> ring_hom P S"
  1282   shows "ring_hom_cring P S Phi"
  1283 proof (rule ring_hom_cring.intro)
  1284   show "ring_hom_cring_axioms P S Phi"
  1285   by (rule ring_hom_cring_axioms.intro) (rule Phi)
  1286 qed unfold_locales
  1287 
  1288 theorem (in UP_pre_univ_prop) UP_universal_property:
  1289   assumes S: "s \<in> carrier S"
  1290   shows "EX! Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) &
  1291     Phi (monom P \<one> 1) = s &
  1292     (ALL r : carrier R. Phi (monom P r 0) = h r)"
  1293   using S eval_monom1
  1294   apply (auto intro: eval_ring_hom eval_const eval_extensional)
  1295   apply (rule extensionalityI)
  1296   apply (auto intro: UP_hom_unique ring_homD)
  1297   done
  1298 
  1299 
  1300 subsection {* Sample Application of Evaluation Homomorphism *}
  1301 
  1302 lemma UP_pre_univ_propI:
  1303   assumes "cring R"
  1304     and "cring S"
  1305     and "h \<in> ring_hom R S"
  1306   shows "UP_pre_univ_prop R S h"
  1307   using assms
  1308   by (auto intro!: UP_pre_univ_prop.intro ring_hom_cring.intro
  1309     ring_hom_cring_axioms.intro UP_cring.intro)
  1310 
  1311 constdefs
  1312   INTEG :: "int ring"
  1313   "INTEG == (| carrier = UNIV, mult = op *, one = 1, zero = 0, add = op + |)"
  1314 
  1315 lemma INTEG_cring:
  1316   "cring INTEG"
  1317   by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI
  1318     zadd_zminus_inverse2 zadd_zmult_distrib)
  1319 
  1320 lemma INTEG_id_eval:
  1321   "UP_pre_univ_prop INTEG INTEG id"
  1322   by (fast intro: UP_pre_univ_propI INTEG_cring id_ring_hom)
  1323 
  1324 text {*
  1325   Interpretation now enables to import all theorems and lemmas
  1326   valid in the context of homomorphisms between @{term INTEG} and @{term
  1327   "UP INTEG"} globally.
  1328 *}
  1329 
  1330 interpretation INTEG: UP_pre_univ_prop [INTEG INTEG id]
  1331   apply simp
  1332   using INTEG_id_eval
  1333   apply simp
  1334   done
  1335 
  1336 lemma INTEG_closed [intro, simp]:
  1337   "z \<in> carrier INTEG"
  1338   by (unfold INTEG_def) simp
  1339 
  1340 lemma INTEG_mult [simp]:
  1341   "mult INTEG z w = z * w"
  1342   by (unfold INTEG_def) simp
  1343 
  1344 lemma INTEG_pow [simp]:
  1345   "pow INTEG z n = z ^ n"
  1346   by (induct n) (simp_all add: INTEG_def nat_pow_def)
  1347 
  1348 lemma "eval INTEG INTEG id 10 (monom (UP INTEG) 5 2) = 500"
  1349   by (simp add: INTEG.eval_monom)
  1350 
  1351 end