author | berghofe |
Fri, 10 Dec 2004 16:48:01 +0100 | |
changeset 15394 | a2c34e6ca4f8 |
parent 14981 | e73f8140af78 |
child 15566 | eb3b1a5c304e |
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 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
5 |
Class instance of -> for class pcpo |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
6 |
|
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 |
Cfun3 = Cfun2 + |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
10 |
|
2838
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
slotosch
parents:
2640
diff
changeset
|
11 |
instance "->" :: (cpo,cpo)cpo (cpo_cfun) |
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
slotosch
parents:
2640
diff
changeset
|
12 |
instance "->" :: (cpo,pcpo)pcpo (least_cfun) |
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
slotosch
parents:
2640
diff
changeset
|
13 |
|
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
slotosch
parents:
2640
diff
changeset
|
14 |
default pcpo |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
15 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
16 |
consts |
1479 | 17 |
Istrictify :: "('a->'b)=>'a=>'b" |
18 |
strictify :: "('a->'b)->'a->'b" |
|
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
1150
diff
changeset
|
19 |
defs |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
20 |
|
10834 | 21 |
Istrictify_def "Istrictify f x == if x=UU then UU else f$x" |
3842 | 22 |
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
|
23 |
|
3327 | 24 |
consts |
25 |
ID :: "('a::cpo) -> 'a" |
|
26 |
cfcomp :: "('b->'c)->(('a::cpo)->('b::cpo))->'a->('c::cpo)" |
|
27 |
||
28 |
syntax "@oo" :: "('b->'c)=>('a->'b)=>'a->'c" ("_ oo _" [101,100] 100) |
|
29 |
||
10834 | 30 |
translations "f1 oo f2" == "cfcomp$f1$f2" |
3327 | 31 |
|
32 |
defs |
|
33 |
||
3842 | 34 |
ID_def "ID ==(LAM x. x)" |
10834 | 35 |
oo_def "cfcomp == (LAM f g x. f$(g$x))" |
3327 | 36 |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
37 |
end |