src/HOLCF/fun1.thy
changeset 243 c22b85994e17
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/fun1.thy	Wed Jan 19 17:35:01 1994 +0100
@@ -0,0 +1,30 @@
+(*  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 +
+
+(* default class is still term *)
+
+consts
+  less_fun	:: "['a=>'b::po,'a=>'b] => bool"	
+
+rules
+   (* definition of the ordering less_fun            *)
+   (* in fun1.ML it is proved that less_fun is a po *)
+   
+  less_fun_def "less_fun(f1,f2) == ! x. f1(x) << f2(x)"  
+
+end
+
+
+
+