author | wenzelm |
Fri, 16 Oct 1998 18:52:17 +0200 | |
changeset 5660 | f2c5354cd32f |
parent 4721 | c8a8482a8124 |
child 11346 | 0d28bc664955 |
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 |
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 |
|
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 |
*) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
9 |
|
297 | 10 |
Porder = Porder0 + |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
11 |
|
1479 | 12 |
consts |
13 |
"<|" :: "['a set,'a::po] => bool" (infixl 55) |
|
14 |
"<<|" :: "['a set,'a::po] => bool" (infixl 55) |
|
15 |
lub :: "'a set => 'a::po" |
|
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
3842
diff
changeset
|
16 |
tord :: "'a::po set => bool" |
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
3842
diff
changeset
|
17 |
chain :: "(nat=>'a::po) => bool" |
1479 | 18 |
max_in_chain :: "[nat,nat=>'a::po]=>bool" |
19 |
finite_chain :: "(nat=>'a::po)=>bool" |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
20 |
|
2394 | 21 |
syntax |
22 |
||
23 |
"@LUB" :: "(('b::term) => 'a) => 'a" (binder "LUB " 10) |
|
24 |
||
25 |
translations |
|
26 |
||
3842 | 27 |
"LUB x. t" == "lub(range(%x. t))" |
2394 | 28 |
|
29 |
syntax (symbols) |
|
30 |
||
31 |
"LUB " :: "[idts, 'a] => 'a" ("(3\\<Squnion>_./ _)"[0,10] 10) |
|
32 |
||
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
297
diff
changeset
|
33 |
defs |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
34 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
35 |
(* class definitions *) |
3842 | 36 |
is_ub "S <| x == ! y. y:S --> y<<x" |
1479 | 37 |
is_lub "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
|
38 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
39 |
(* Arbitrary chains are total orders *) |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
3842
diff
changeset
|
40 |
tord "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
|
41 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
42 |
(* Here we use countable chains and I prefer to code them as functions! *) |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
3842
diff
changeset
|
43 |
chain "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
|
44 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
45 |
(* finite chains, needed for monotony of continouous functions *) |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
297
diff
changeset
|
46 |
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:
3842
diff
changeset
|
47 |
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
|
48 |
|
2640 | 49 |
lub_def "lub S == (@x. S <<| x)" |
1274 | 50 |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
51 |
end |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
297
diff
changeset
|
52 |
|
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
297
diff
changeset
|
53 |