| 1247 |      1 | (*  Title:      GroupDefs.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | GroupDefs = MonoidGroupInsts + Prod + Fun +
 | 
|  |      7 | 
 | 
|  |      8 | 
 | 
|  |      9 | (* bool *)
 | 
|  |     10 | 
 | 
|  |     11 | instance
 | 
| 2907 |     12 |   bool :: {times, inverse, one}
 | 
| 1247 |     13 | 
 | 
|  |     14 | defs
 | 
|  |     15 |   times_bool_def        "x * y == (x ~= y)"
 | 
| 2907 |     16 |   inverse_bool_def      "inverse x == (x::bool)"
 | 
| 1247 |     17 |   one_bool_def          "1 == False"
 | 
|  |     18 | 
 | 
|  |     19 | 
 | 
|  |     20 | (* cartesian products *)
 | 
|  |     21 | 
 | 
|  |     22 | instance
 | 
| 2907 |     23 |   "*" :: (term, term) {times, inverse, one}
 | 
| 1247 |     24 | 
 | 
|  |     25 | defs
 | 
|  |     26 |   times_prod_def        "p * q == (fst p * fst q, snd p * snd q)"
 | 
| 2907 |     27 |   inverse_prod_def      "inverse p == (inverse (fst p), inverse (snd p))"
 | 
| 1247 |     28 |   one_prod_def          "1 == (1, 1)"
 | 
|  |     29 | 
 | 
|  |     30 | 
 | 
|  |     31 | (* function spaces *)
 | 
|  |     32 | 
 | 
|  |     33 | instance
 | 
| 2907 |     34 |   fun :: (term, term) {times, inverse, one}
 | 
| 1247 |     35 | 
 | 
|  |     36 | defs
 | 
|  |     37 |   times_fun_def         "f * g == (%x. f x * g x)"
 | 
| 2907 |     38 |   inverse_fun_def       "inverse f == (%x. inverse (f x))"
 | 
| 1247 |     39 |   one_fun_def           "1 == (%x. 1)"
 | 
|  |     40 | 
 | 
|  |     41 | end
 |