src/HOLCF/Cfun2.thy
changeset 1479 21eb5e156d91
parent 1168 74be52691d62
child 2640 ee4dfce170a0
equal deleted inserted replaced
1478:2b8c2a7547ab 1479:21eb5e156d91
     1 (*  Title: 	HOLCF/cfun2.thy
     1 (*  Title:      HOLCF/cfun2.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author: 	Franz Regensburger
     3     Author:     Franz Regensburger
     4     Copyright   1993 Technische Universitaet Muenchen
     4     Copyright   1993 Technische Universitaet Muenchen
     5 
     5 
     6 Class Instance ->::(pcpo,pcpo)po
     6 Class Instance ->::(pcpo,pcpo)po
     7 
     7 
     8 *)
     8 *)
    10 Cfun2 = Cfun1 + 
    10 Cfun2 = Cfun1 + 
    11 
    11 
    12 (* Witness for the above arity axiom is cfun1.ML *)
    12 (* Witness for the above arity axiom is cfun1.ML *)
    13 arities "->" :: (pcpo,pcpo)po
    13 arities "->" :: (pcpo,pcpo)po
    14 
    14 
    15 consts	
    15 consts  
    16 	UU_cfun  :: "'a->'b"
    16         UU_cfun  :: "'a->'b"
    17 
    17 
    18 rules
    18 rules
    19 
    19 
    20 (* instance of << for type ['a -> 'b]  *)
    20 (* instance of << for type ['a -> 'b]  *)
    21 
    21 
    22 inst_cfun_po	"((op <<)::['a->'b,'a->'b]=>bool) = less_cfun"
    22 inst_cfun_po    "((op <<)::['a->'b,'a->'b]=>bool) = less_cfun"
    23 
    23 
    24 defs
    24 defs
    25 (* The least element in type 'a->'b *)
    25 (* The least element in type 'a->'b *)
    26 
    26 
    27 UU_cfun_def	"UU_cfun == fabs(% x.UU)"
    27 UU_cfun_def     "UU_cfun == fabs(% x.UU)"
    28 
    28 
    29 end
    29 end
    30 
    30