src/HOLCF/ssum2.thy
changeset 243 c22b85994e17
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/ssum2.thy	Wed Jan 19 17:35:01 1994 +0100
@@ -0,0 +1,23 @@
+(*  Title: 	HOLCF/ssum2.thy
+    ID:         $Id$
+    Author: 	Franz Regensburger
+    Copyright   1993 Technische Universitaet Muenchen
+
+Class Instance ++::(pcpo,pcpo)po
+*)
+
+Ssum2 = Ssum1 + 
+
+arities "++" :: (pcpo,pcpo)po
+(* Witness for the above arity axiom is ssum1.ML *)
+
+rules
+
+(* instance of << for type ['a ++ 'b]  *)
+
+inst_ssum_po	"(op <<)::['a ++ 'b,'a ++ 'b]=>bool = less_ssum"
+
+end
+
+
+