author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
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 |