src/HOL/Algebra/poly/PolyHomo.thy
author paulson
Wed, 02 Jan 2002 16:07:16 +0100
changeset 12614 a65d72ddc657
parent 11779 1aa328cb273a
child 13735 7de9342aca7a
permissions -rw-r--r--
New theorems by Sidi Ehmety

(*
    Universal property and evaluation homomorphism of univariate polynomials
    $Id$
    Author: Clemens Ballarin, started 15 April 1997
*)

PolyHomo = Degree +

(* Instantiate result from Degree.ML *)

instance
  up :: (domain) domain (up_one_not_zero, up_integral)

consts
  EVAL2	:: "('a::ring => 'b) => 'b => 'a up => 'b::ring"
  EVAL	:: "'a::ring => 'a up => 'a"

defs
  EVAL2_def	"EVAL2 phi a p ==
                 setsum (%i. phi (coeff i p) * a ^ i) {..deg p}"
  EVAL_def	"EVAL == EVAL2 (%x. x)"

end