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