| author | wenzelm |
| Thu, 23 Oct 1997 12:48:48 +0200 | |
| changeset 3979 | dac05c9341f4 |
| parent 1266 | 3ae9fe3c0f68 |
| child 4091 | 771b1f6422a8 |
| permissions | -rw-r--r-- |
(* Title: HOL/AxClasses/Tutorial/ProdGroupInsts.thy ID: $Id$ Author: Markus Wenzel, TU Muenchen Lift constant "<*>" to cartesian products, then prove that the 'functor' "*" maps semigroups into semigroups. *) ProdGroupInsts = Prod + Group + (* direct products of semigroups *) defs prod_prod_def "p <*> q == (fst p <*> fst q, snd p <*> snd q)" instance "*" :: (semigroup, semigroup) semigroup {| simp_tac (!simpset addsimps [assoc]) 1 |} end