| author | wenzelm |
| Tue, 20 Oct 1998 16:29:08 +0200 | |
| changeset 5688 | 7f582495967c |
| parent 2907 | 0e272e4c7cb2 |
| child 6393 | b8dafa978382 |
| permissions | -rw-r--r-- |
| 1247 | 1 |
(* Title: HOL/AxClasses/Tutorial/Xor.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Proof the instantiation theorem bool :: agroup by hand. |
|
6 |
*) |
|
7 |
||
8 |
open AxClass; |
|
9 |
||
10 |
goal_arity Xor.thy ("bool", [], "agroup");
|
|
11 |
by (axclass_tac Xor.thy []); |
|
| 2907 | 12 |
by (rewrite_goals_tac |
13 |
[Xor.prod_bool_def,Xor.inverse_bool_def, Xor.unit_bool_def]); |
|
| 1899 | 14 |
by (ALLGOALS (Fast_tac)); |
| 1247 | 15 |
qed "bool_in_agroup"; |