| author | clasohm | 
| Mon, 11 Sep 1995 12:38:20 +0200 | |
| changeset 1253 | 131f72e2cd56 | 
| parent 1247 | 18b1441fb603 | 
| permissions | -rw-r--r-- | 
| 1247 | 1 | (* Title: HOL/AxClasses/Tutorial/Monoid.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Define class "monoid". | |
| 6 | *) | |
| 7 | ||
| 8 | Monoid = Sigs + | |
| 9 | ||
| 10 | (* monoids *) | |
| 11 | ||
| 12 | axclass | |
| 13 | monoid < term | |
| 14 | assoc "(x <*> y) <*> z = x <*> (y <*> z)" | |
| 15 | left_unit "1 <*> x = x" | |
| 16 | right_unit "x <*> 1 = x" | |
| 17 | ||
| 18 | end |