author | paulson |
Sat, 29 Jun 2002 22:46:56 +0200 | |
changeset 13260 | ea36a40c004f |
parent 12114 | a8e860c86252 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
6382 | 1 |
(* Title: HOLCF/Cfun1.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 |
12030 | 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 |
|
6382 | 6 |
Definition of the type -> of continuous functions. |
243
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 |
Cfun1 = Cont + |
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 |
default cpo |
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
slotosch
parents:
2640
diff
changeset
|
13 |
|
6382 | 14 |
typedef (CFun) ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}" (CfunI) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
15 |
|
3323
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
2838
diff
changeset
|
16 |
(* to make << defineable *) |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
2838
diff
changeset
|
17 |
instance "->" :: (cpo,cpo)sq_ord |
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents:
2838
diff
changeset
|
18 |
|
5291 | 19 |
syntax |
10863 | 20 |
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999) |
1479 | 21 |
(* application *) |
10863 | 22 |
Abs_CFun :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10) |
1479 | 23 |
(* abstraction *) |
24 |
less_cfun :: "[('a -> 'b),('a -> 'b)]=>bool" |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
25 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12030
diff
changeset
|
26 |
syntax (xsymbols) |
3394 | 27 |
"->" :: [type, type] => type ("(_ \\<rightarrow>/ _)" [1,0]0) |
2394 | 28 |
"LAM " :: "[idts, 'a => 'b] => ('a -> 'b)" |
29 |
("(3\\<Lambda>_./ _)" [0, 10] 10) |
|
10863 | 30 |
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\\<cdot>_)" [999,1000] 999) |
31 |
||
32 |
syntax (HTML output) |
|
33 |
Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\\<cdot>_)" [999,1000] 999) |
|
34 |
||
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
430
diff
changeset
|
35 |
defs |
5291 | 36 |
less_cfun_def "(op <<) == (% fo1 fo2. Rep_CFun fo1 << Rep_CFun fo2 )" |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
37 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
38 |
end |