author | berghofe |
Thu, 27 Jan 2005 12:37:02 +0100 | |
changeset 15475 | fdf9434b04ea |
parent 14981 | e73f8140af78 |
child 15563 | 9e125b675253 |
permissions | -rw-r--r-- |
2640 | 1 |
(* Title: HOLCF/Pcpo.thy |
2 |
ID: $Id$ |
|
3 |
Author: Franz Regensburger |
|
4 |
||
5 |
introduction of the classes cpo and pcpo |
|
6 |
*) |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
7 |
Pcpo = Porder + |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
8 |
|
2640 | 9 |
(* The class cpo of chain complete partial orders *) |
10 |
(* ********************************************** *) |
|
11 |
axclass cpo < po |
|
12 |
(* class axiom: *) |
|
5438 | 13 |
cpo "chain S ==> ? x. range S <<| x" |
2394 | 14 |
|
2640 | 15 |
(* The class pcpo of pointed cpos *) |
16 |
(* ****************************** *) |
|
17 |
axclass pcpo < cpo |
|
18 |
||
3842 | 19 |
least "? x.!y. x<<y" |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
20 |
|
2394 | 21 |
consts |
2640 | 22 |
UU :: "'a::pcpo" |
2394 | 23 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12030
diff
changeset
|
24 |
syntax (xsymbols) |
2640 | 25 |
UU :: "'a::pcpo" ("\\<bottom>") |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
26 |
|
2640 | 27 |
defs |
3842 | 28 |
UU_def "UU == @x.!y. x<<y" |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
29 |
|
3326 | 30 |
(* further useful classes for HOLCF domains *) |
31 |
||
32 |
axclass chfin<cpo |
|
33 |
||
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
3842
diff
changeset
|
34 |
chfin "!Y. chain Y-->(? n. max_in_chain n Y)" |
3326 | 35 |
|
36 |
axclass flat<pcpo |
|
37 |
||
3842 | 38 |
ax_flat "! x y. x << y --> (x = UU) | (x=y)" |
3326 | 39 |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
40 |
end |