| author | clasohm | 
| Mon, 11 Sep 1995 12:38:20 +0200 | |
| changeset 1253 | 131f72e2cd56 | 
| parent 1247 | 18b1441fb603 | 
| child 2907 | 0e272e4c7cb2 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/AxClasses/Tutorial/Group.thy ID: $Id$ Author: Markus Wenzel, TU Muenchen Define classes "semigroup", "group", "agroup". *) Group = Sigs + (* semigroups *) axclass semigroup < term assoc "(x <*> y) <*> z = x <*> (y <*> z)" (* groups *) axclass group < semigroup left_unit "1 <*> x = x" left_inv "inv x <*> x = 1" (* abelian groups *) axclass agroup < group commut "x <*> y = y <*> x" end