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