| author | nipkow | 
| Tue, 26 Feb 2002 13:47:19 +0100 | |
| changeset 12950 | 3aadb133632d | 
| parent 12338 | de0f4a63baa5 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 1479 | 1 | (* Title: HOLCF/porder.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 | |
| 297 | 6 | Conservative extension of theory Porder0 by constant definitions | 
| 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 | |
| 297 | 9 | Porder = Porder0 + | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 10 | |
| 1479 | 11 | consts | 
| 12 | "<|" :: "['a set,'a::po] => bool" (infixl 55) | |
| 13 | "<<|" :: "['a set,'a::po] => bool" (infixl 55) | |
| 14 | lub :: "'a set => 'a::po" | |
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
3842diff
changeset | 15 | tord :: "'a::po set => bool" | 
| 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
3842diff
changeset | 16 | chain :: "(nat=>'a::po) => bool" | 
| 1479 | 17 | max_in_chain :: "[nat,nat=>'a::po]=>bool" | 
| 18 | finite_chain :: "(nat=>'a::po)=>bool" | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 19 | |
| 2394 | 20 | syntax | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12114diff
changeset | 21 |   "@LUB"	:: "('b => 'a) => 'a"	(binder "LUB " 10)
 | 
| 2394 | 22 | |
| 23 | translations | |
| 24 | ||
| 3842 | 25 | "LUB x. t" == "lub(range(%x. t))" | 
| 2394 | 26 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12030diff
changeset | 27 | syntax (xsymbols) | 
| 2394 | 28 | |
| 29 |   "LUB "	:: "[idts, 'a] => 'a"		("(3\\<Squnion>_./ _)"[0,10] 10)
 | |
| 30 | ||
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
297diff
changeset | 31 | defs | 
| 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 | (* class definitions *) | 
| 11346 | 34 | is_ub_def "S <| x == ! y. y:S --> y<<x" | 
| 35 | is_lub_def "S <<| x == S <| x & (!u. S <| u --> x << u)" | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 36 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 37 | (* Arbitrary chains are total orders *) | 
| 11346 | 38 | tord_def "tord S == !x y. x:S & y:S --> (x<<y | y<<x)" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 39 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 40 | (* Here we use countable chains and I prefer to code them as functions! *) | 
| 11346 | 41 | chain_def "chain F == !i. F i << F (Suc i)" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 42 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 43 | (* finite chains, needed for monotony of continouous functions *) | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
297diff
changeset | 44 | max_in_chain_def "max_in_chain i C == ! j. i <= j --> C(i) = C(j)" | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
3842diff
changeset | 45 | finite_chain_def "finite_chain C == chain(C) & (? i. max_in_chain i C)" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 46 | |
| 2640 | 47 | lub_def "lub S == (@x. S <<| x)" | 
| 1274 | 48 | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 49 | end | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
297diff
changeset | 50 | |
| 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
297diff
changeset | 51 |