| author | oheimb | 
| Sat, 25 Oct 1997 14:43:55 +0200 | |
| changeset 4005 | 8858c472691a | 
| parent 3842 | b55686a7b22c | 
| child 10834 | a7897aebbffc | 
| permissions | -rw-r--r-- | 
| 2640 | 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 | 3  | 
Author: Franz Regensburger  | 
| 
243
 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 
nipkow 
parents:  
diff
changeset
 | 
4  | 
Copyright 1993 Technische Universitaet Muenchen  | 
| 
 
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 | 18  | 
        Istrictify   :: "('a->'b)=>'a=>'b"
 | 
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  | 
|
| 1479 | 22  | 
Istrictify_def "Istrictify f x == if x=UU then UU else f`x"  | 
| 3842 | 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 | 25  | 
consts  | 
26  | 
        ID      :: "('a::cpo) -> 'a"
 | 
|
27  | 
        cfcomp  :: "('b->'c)->(('a::cpo)->('b::cpo))->'a->('c::cpo)"
 | 
|
28  | 
||
29  | 
syntax  "@oo"   :: "('b->'c)=>('a->'b)=>'a->'c" ("_ oo _" [101,100] 100)
 | 
|
30  | 
||
31  | 
translations "f1 oo f2" == "cfcomp`f1`f2"  | 
|
32  | 
||
33  | 
defs  | 
|
34  | 
||
| 3842 | 35  | 
ID_def "ID ==(LAM x. x)"  | 
36  | 
oo_def "cfcomp == (LAM f g x. f`(g`x))"  | 
|
| 3327 | 37  | 
|
| 
243
 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 
nipkow 
parents:  
diff
changeset
 | 
38  | 
end  |