| author | paulson | 
| Wed, 28 Nov 2007 16:26:03 +0100 | |
| changeset 25492 | 4cc7976948ac | 
| parent 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 10134 | 1  | 
(* Title: HOL/AxClasses/Product.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
*)  | 
|
5  | 
||
| 16417 | 6  | 
theory Product imports Main begin  | 
| 10134 | 7  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
10681 
diff
changeset
 | 
8  | 
axclass product < type  | 
| 10134 | 9  | 
|
10  | 
consts  | 
|
11  | 
product :: "'a::product => 'a => 'a" (infixl "[*]" 70)  | 
|
12  | 
||
13  | 
||
14  | 
instance bool :: product  | 
|
15  | 
by intro_classes  | 
|
16  | 
||
17  | 
defs (overloaded)  | 
|
18  | 
product_bool_def: "x [*] y == x & y"  | 
|
19  | 
||
20  | 
end  |