| author | wenzelm |
| Thu, 03 Aug 2000 18:44:24 +0200 | |
| changeset 9513 | 8531c18d9181 |
| parent 8920 | af5e09b6c208 |
| child 10007 | 64bf7da1994a |
| permissions | -rw-r--r-- |
(* Title: HOL/AxClasses/Tutorial/Semigroups.thy ID: $Id$ Author: Markus Wenzel, TU Muenchen *) theory Semigroups = Main:; consts times :: "'a => 'a => 'a" (infixl "[*]" 70); axclass semigroup < "term" assoc: "(x [*] y) [*] z = x [*] (y [*] z)"; consts plus :: "'a => 'a => 'a" (infixl "[+]" 70); axclass plus_semigroup < "term" assoc: "(x [+] y) [+] z = x [+] (y [+] z)"; end;