author | nipkow |
Fri, 04 Apr 1997 16:03:44 +0200 | |
changeset 2907 | 0e272e4c7cb2 |
parent 1573 | 6d66b59f94a9 |
child 8920 | af5e09b6c208 |
permissions | -rw-r--r-- |
(* Title: HOL/AxClasses/Tutorial/Semigroups.thy ID: $Id$ Author: Markus Wenzel, TU Muenchen Semigroups with different 'signatures'. *) Semigroups = HOL + consts "<+>" :: "['a, 'a] => 'a" (infixl 65) "<*>" :: "['a, 'a] => 'a" (infixl 70) constdefs assoc :: "(['a, 'a] => 'a) => bool" "assoc f == ALL x y z. f (f x y) z = f x (f y z)" (* semigroups with op <+> *) axclass plus_sg < term plus_assoc "assoc (op <+>)" (* semigroups with op <*> *) axclass times_sg < term times_assoc "assoc (op <*>)" end