| author | berghofe | 
| Fri, 29 Sep 2000 18:02:24 +0200 | |
| changeset 10117 | 8e58b3045e29 | 
| parent 6382 | 8b0c9205da75 | 
| child 10834 | a7897aebbffc | 
| 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 | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 4 | Copyright 1993 Technische Universitaet Muenchen | 
| 
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: 
2640diff
changeset | 12 | default cpo | 
| 
2e908f29bc3d
changed continuous functions from pcpo to cpo (including instances)
 slotosch parents: 
2640diff
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: 
2838diff
changeset | 16 | (* to make << defineable *) | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2838diff
changeset | 17 | instance "->" :: (cpo,cpo)sq_ord | 
| 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2838diff
changeset | 18 | |
| 5291 | 19 | syntax | 
| 20 | 	Rep_CFun  :: "('a -> 'b)=>('a => 'b)" ("_`_" [999,1000] 999)
 | |
| 1479 | 21 | (* application *) | 
| 5291 | 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 | |
| 2394 | 26 | syntax (symbols) | 
| 3394 | 27 |   "->"		:: [type, type] => type	("(_ \\<rightarrow>/ _)" [1,0]0)
 | 
| 2394 | 28 |   "LAM "	:: "[idts, 'a => 'b] => ('a -> 'b)"
 | 
| 29 | 					("(3\\<Lambda>_./ _)" [0, 10] 10)
 | |
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
430diff
changeset | 30 | defs | 
| 5291 | 31 | 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 | 32 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 33 | end |