| author | wenzelm | 
| Mon, 09 Mar 1998 16:11:28 +0100 | |
| changeset 4700 | 20ade76722d6 | 
| parent 3327 | 9b8e638f8602 | 
| child 4721 | c8a8482a8124 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/Fun1.thy ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Definition of the partial ordering for the type of all functions => (fun) REMARK: The ordering on 'a => 'b is only defined if 'b is in class po !! *) Fun1 = Pcpo + instance flat<chfin (flat_imp_chain_finite) (* to make << defineable: *) instance fun :: (term,sq_ord)sq_ord defs less_fun_def "(op <<) == (%f1 f2.!x. f1 x << f2 x)" end