author | paulson |
Thu, 18 Jan 1996 10:38:29 +0100 | |
changeset 1444 | 23ceb1dc9755 |
parent 1247 | 18b1441fb603 |
child 1899 | 0075a8d26a80 |
permissions | -rw-r--r-- |
(* Title: HOL/AxClasses/Tutorial/Xor.ML ID: $Id$ Author: Markus Wenzel, TU Muenchen Proof the instantiation theorem bool :: agroup by hand. *) open AxClass; goal_arity Xor.thy ("bool", [], "agroup"); by (axclass_tac Xor.thy []); by (rewrite_goals_tac [Xor.prod_bool_def, Xor.inv_bool_def, Xor.unit_bool_def]); by (ALLGOALS (fast_tac HOL_cs)); qed "bool_in_agroup";