src/HOLCF/Cfun3.thy
author wenzelm
Sat, 03 Nov 2001 01:41:26 +0100
changeset 12030 46d57d0290a2
parent 10834 a7897aebbffc
child 14981 e73f8140af78
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 1479
diff changeset
     1
(*  Title:      HOLCF/Cfun3.thy
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
     3
    Author:     Franz Regensburger
12030
wenzelm
parents: 10834
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     6
Class instance of  -> for class pcpo
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     8
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     9
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
Cfun3 = Cfun2 +
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    11
2838
2e908f29bc3d changed continuous functions from pcpo to cpo (including instances)
slotosch
parents: 2640
diff changeset
    12
instance "->" :: (cpo,cpo)cpo              (cpo_cfun)
2e908f29bc3d changed continuous functions from pcpo to cpo (including instances)
slotosch
parents: 2640
diff changeset
    13
instance "->" :: (cpo,pcpo)pcpo            (least_cfun)
2e908f29bc3d changed continuous functions from pcpo to cpo (including instances)
slotosch
parents: 2640
diff changeset
    14
2e908f29bc3d changed continuous functions from pcpo to cpo (including instances)
slotosch
parents: 2640
diff changeset
    15
default pcpo
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    16
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    17
consts  
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    18
        Istrictify   :: "('a->'b)=>'a=>'b"
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    19
        strictify    :: "('a->'b)->'a->'b"
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    20
defs
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    21
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 3842
diff changeset
    22
Istrictify_def  "Istrictify f x == if x=UU then UU else f$x"    
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3327
diff changeset
    23
strictify_def   "strictify == (LAM f x. Istrictify f x)"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    24
3327
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    25
consts
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    26
        ID      :: "('a::cpo) -> 'a"
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    27
        cfcomp  :: "('b->'c)->(('a::cpo)->('b::cpo))->'a->('c::cpo)"
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    28
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    29
syntax  "@oo"   :: "('b->'c)=>('a->'b)=>'a->'c" ("_ oo _" [101,100] 100)
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    30
     
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 3842
diff changeset
    31
translations    "f1 oo f2" == "cfcomp$f1$f2"
3327
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    32
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    33
defs
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    34
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3327
diff changeset
    35
  ID_def        "ID ==(LAM x. x)"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 3842
diff changeset
    36
  oo_def        "cfcomp == (LAM f g x. f$(g$x))" 
3327
9b8e638f8602 Eliminated ccc1. Moved ID,oo into Cfun.
slotosch
parents: 2838
diff changeset
    37
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    38
end