src/HOLCF/Sprod3.thy
author nipkow
Thu, 14 Feb 2002 11:50:52 +0100
changeset 12887 d25b43743e10
parent 12030 46d57d0290a2
child 14981 e73f8140af78
permissions -rw-r--r--
nodups -> distinct

(*  Title:      HOLCF/sprod3.thy
    ID:         $Id$
    Author:     Franz Regensburger
    License:    GPL (GNU GENERAL PUBLIC LICENSE)

Class instance of  ** for class pcpo
*)

Sprod3 = Sprod2 +

instance "**" :: (pcpo,pcpo)pcpo  (least_sprod,cpo_sprod)

consts  
  spair		:: "'a -> 'b -> ('a**'b)" (* continuous strict pairing *)
  sfst		:: "('a**'b)->'a"
  ssnd		:: "('a**'b)->'b"
  ssplit	:: "('a->'b->'c)->('a**'b)->'c"

syntax  
  "@stuple"	:: "['a, args] => 'a ** 'b"	("(1'(:_,/ _:'))")

translations
        "(:x, y, z:)"   == "(:x, (:y, z:):)"
        "(:x, y:)"      == "spair$x$y"

defs
spair_def       "spair  == (LAM x y. Ispair x y)"
sfst_def        "sfst   == (LAM p. Isfst p)"
ssnd_def        "ssnd   == (LAM p. Issnd p)"     
ssplit_def      "ssplit == (LAM f. strictify$(LAM p. f$(sfst$p)$(ssnd$p)))"

end