changeset 2278 | d63ffafce255 |
child 2640 | ee4dfce170a0 |
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Up2.thy Fri Nov 29 12:22:22 1996 +0100 @@ -0,0 +1,25 @@ +(* Title: HOLCF/Up2.thy + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Class Instance u::(pcpo)po + +*) + +Up2 = Up1 + + +(* Witness for the above arity axiom is up1.ML *) + +arities "u" :: (pcpo)po + +rules + +(* instance of << for type ('a)u *) + +inst_up_po "((op <<)::[('a)u,('a)u]=>bool) = less_up" + +end + + +