| author | paulson | 
| Fri, 29 Mar 1996 10:54:44 +0100 | |
| changeset 1629 | b5e43a60443a | 
| parent 1247 | 18b1441fb603 | 
| child 8920 | af5e09b6c208 | 
| permissions | -rw-r--r-- | 
| 1247 | 1  | 
(* Title: HOL/AxClasses/Tutorial/Product.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
5  | 
Define a 'syntactic' class "product" that is logically equivalent to  | 
|
6  | 
"term".  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
Product = HOL +  | 
|
10  | 
||
11  | 
axclass  | 
|
12  | 
product < term  | 
|
13  | 
||
14  | 
consts  | 
|
15  | 
"<*>" :: "['a::product, 'a] => 'a" (infixl 70)  | 
|
16  | 
||
17  | 
end  |