author | wenzelm |
Fri, 01 Dec 2000 19:39:54 +0100 | |
changeset 10562 | fcd29e58c40c |
parent 9964 | 7966a2902266 |
child 12114 | a8e860c86252 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/OrderType.thy |
435 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
850 | 6 |
Order types and ordinal arithmetic. |
435 | 7 |
|
8 |
The order type of a well-ordering is the least ordinal isomorphic to it. |
|
9 |
*) |
|
10 |
||
2469 | 11 |
OrderType = OrderArith + OrdQuant + |
435 | 12 |
consts |
1401 | 13 |
ordermap :: [i,i]=>i |
14 |
ordertype :: [i,i]=>i |
|
435 | 15 |
|
1401 | 16 |
Ord_alt :: i => o |
850 | 17 |
|
1401 | 18 |
"**" :: [i,i]=>i (infixl 70) |
19 |
"++" :: [i,i]=>i (infixl 65) |
|
20 |
"--" :: [i,i]=>i (infixl 65) |
|
850 | 21 |
|
22 |
||
753 | 23 |
defs |
435 | 24 |
ordermap_def |
25 |
"ordermap(A,r) == lam x:A. wfrec[A](r, x, %x f. f `` pred(A,x,r))" |
|
26 |
||
27 |
ordertype_def "ordertype(A,r) == ordermap(A,r)``A" |
|
28 |
||
850 | 29 |
Ord_alt_def (*alternative definition of ordinal numbers*) |
30 |
"Ord_alt(X) == well_ord(X, Memrel(X)) & (ALL u:X. u=pred(X, u, Memrel(X)))" |
|
31 |
||
32 |
(*ordinal multiplication*) |
|
33 |
omult_def "i ** j == ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))" |
|
34 |
||
35 |
(*ordinal addition*) |
|
36 |
oadd_def "i ++ j == ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))" |
|
37 |
||
1033 | 38 |
(*ordinal subtraction*) |
39 |
odiff_def "i -- j == ordertype(i-j, Memrel(i))" |
|
40 |
||
9964 | 41 |
syntax (symbols) |
42 |
"op **" :: [i,i] => i (infixl "\\<times>\\<times>" 70) |
|
43 |
||
44 |
syntax (HTML output) |
|
9683 | 45 |
"op **" :: [i,i] => i (infixl "\\<times>\\<times>" 70) |
9654
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents:
2469
diff
changeset
|
46 |
|
435 | 47 |
end |