| author | paulson | 
| Sun, 15 Feb 2004 10:46:37 +0100 | |
| changeset 14387 | e96d5c42c4b0 | 
| parent 12338 | de0f4a63baa5 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/Fun1.thy ID: $Id$ Author: Franz Regensburger License: GPL (GNU GENERAL PUBLIC LICENSE) 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_chfin) (* to make << defineable: *) instance fun :: (type, sq_ord) sq_ord defs less_fun_def "(op <<) == (%f1 f2.!x. f1 x << f2 x)" end