| author | paulson | 
| Tue, 22 Jul 1997 11:15:14 +0200 | |
| changeset 3539 | d4443afc8d28 | 
| parent 3326 | 930c9bed5a09 | 
| child 3842 | b55686a7b22c | 
| 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: *) | |
| 14 | cpo "is_chain S ==> ? x. range(S) <<| (x::'a::po)" | |
| 2394 | 15 | |
| 2640 | 16 | (* The class pcpo of pointed cpos *) | 
| 17 | (* ****************************** *) | |
| 18 | axclass pcpo < cpo | |
| 19 | ||
| 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 | 
| 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 | ||
| 35 | chfin "!Y.is_chain Y-->(? n.max_in_chain n Y)" | |
| 36 | ||
| 37 | axclass flat<pcpo | |
| 38 | ||
| 39 | ax_flat "! x y.x << y --> (x = UU) | (x=y)" | |
| 40 | ||
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 41 | end |