| author | wenzelm |
| Wed, 06 Aug 1997 11:52:16 +0200 | |
| changeset 3619 | 0fc67ad6d62a |
| parent 2907 | 0e272e4c7cb2 |
| permissions | -rw-r--r-- |
| 1247 | 1 |
(* Title: HOL/AxClasses/Tutorial/Xor.thy |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
| 2907 | 5 |
Define overloaded constants "<*>", "inverse", "1" on type "bool". |
| 1247 | 6 |
*) |
7 |
||
8 |
Xor = Group + |
|
9 |
||
10 |
defs |
|
| 2907 | 11 |
prod_bool_def "x <*> y == x ~= (y::bool)" |
12 |
inverse_bool_def "inverse x == x::bool" |
|
13 |
unit_bool_def "1 == False" |
|
| 1247 | 14 |
|
15 |
end |