| 1247 |      1 | (*  Title:      HOL/AxClasses/Tutorial/Semigroups.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Semigroups with different 'signatures'.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | Semigroups = HOL +
 | 
|  |      9 | 
 | 
|  |     10 | consts
 | 
|  |     11 |   "<+>"         :: "['a, 'a] => 'a"             (infixl 65)
 | 
|  |     12 |   "<*>"         :: "['a, 'a] => 'a"             (infixl 70)
 | 
|  |     13 |   assoc         :: "(['a, 'a] => 'a) => bool"
 | 
|  |     14 | 
 | 
|  |     15 | defs
 | 
|  |     16 |   assoc_def     "assoc f == ALL x y z. f (f x y) z = f x (f y z)"
 | 
|  |     17 | 
 | 
|  |     18 | 
 | 
|  |     19 | (* semigroups with op <+> *)
 | 
|  |     20 | 
 | 
|  |     21 | axclass
 | 
|  |     22 |   plus_sg < term
 | 
|  |     23 |   plus_assoc    "assoc (op <+>)"
 | 
|  |     24 | 
 | 
|  |     25 | 
 | 
|  |     26 | (* semigroups with op <*> *)
 | 
|  |     27 | 
 | 
|  |     28 | axclass
 | 
|  |     29 |   times_sg < term
 | 
|  |     30 |   times_assoc   "assoc (op <*>)"
 | 
|  |     31 | 
 | 
|  |     32 | end
 |