src/HOLCF/Ssum3.thy
author regensbu
Fri, 06 Oct 1995 17:25:24 +0100
changeset 1274 ea0668a1c0ba
parent 1168 74be52691d62
child 1479 21eb5e156d91
permissions -rw-r--r--
added 8bit pragmas added directory ax_ops for sections axioms and ops added directory domain for sections domain and generated this is the type definition package of David Oheimb

(*  Title: 	HOLCF/ssum3.thy
    ID:         $Id$
    Author: 	Franz Regensburger
    Copyright   1993 Technische Universitaet Muenchen

Class instance of  ++ for class pcpo
*)

Ssum3 = Ssum2 +

arities "++" :: (pcpo,pcpo)pcpo			(* Witness ssum2.ML *)

consts  
	sinl	:: "'a -> ('a++'b)" 
	sinr	:: "'b -> ('a++'b)" 
	sswhen	:: "('a->'c)->('b->'c)->('a ++ 'b)-> 'c"

rules 

inst_ssum_pcpo	"(UU::'a++'b) = Isinl(UU)"


defs

sinl_def	"sinl   == (LAM x.Isinl(x))"
sinr_def	"sinr   == (LAM x.Isinr(x))"
sswhen_def	"sswhen   == (LAM f g s.Iwhen(f)(g)(s))"

translations
"case s of sinl`x => t1 | sinr`y => t2" == "sswhen`(LAM x.t1)`(LAM y.t2)`s"

(* start 8bit 1 *)
(* end 8bit 1 *)

end