author | lcp |
Fri, 28 Apr 1995 11:24:32 +0200 | |
changeset 1074 | d60f203eeddf |
parent 929 | 137035296ad6 |
child 1097 | 01379c46ad2d |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: ZF/qpair.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
Quine-inspired ordered pairs and disjoint sums, for non-well-founded data |
|
7 |
structures in ZF. Does not precisely follow Quine's construction. Thanks |
|
8 |
to Thomas Forster for suggesting this approach! |
|
9 |
||
10 |
W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, |
|
11 |
1966. |
|
12 |
*) |
|
13 |
||
124 | 14 |
QPair = Sum + "simpdata" + |
0 | 15 |
consts |
16 |
QPair :: "[i, i] => i" ("<(_;/ _)>") |
|
17 |
qsplit :: "[[i,i] => i, i] => i" |
|
18 |
qfsplit :: "[[i,i] => o, i] => o" |
|
19 |
qconverse :: "i => i" |
|
20 |
QSigma :: "[i, i => i] => i" |
|
21 |
||
22 |
"<+>" :: "[i,i]=>i" (infixr 65) |
|
23 |
QInl,QInr :: "i=>i" |
|
24 |
qcase :: "[i=>i, i=>i, i]=>i" |
|
25 |
||
929
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
26 |
syntax |
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
27 |
"@QSUM" :: "[idt, i, i] => i" ("(3QSUM _:_./ _)" 10) |
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
28 |
"<*>" :: "[i, i] => i" (infixr 80) |
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
29 |
|
0 | 30 |
translations |
31 |
"QSUM x:A. B" => "QSigma(A, %x. B)" |
|
44 | 32 |
"A <*> B" => "QSigma(A, _K(B))" |
0 | 33 |
|
753 | 34 |
defs |
0 | 35 |
QPair_def "<a;b> == a+b" |
120 | 36 |
qsplit_def "qsplit(c,p) == THE y. EX a b. p=<a;b> & y=c(a,b)" |
0 | 37 |
qfsplit_def "qfsplit(R,z) == EX x y. z=<x;y> & R(x,y)" |
38 |
qconverse_def "qconverse(r) == {z. w:r, EX x y. w=<x;y> & z=<y;x>}" |
|
39 |
QSigma_def "QSigma(A,B) == UN x:A. UN y:B(x). {<x;y>}" |
|
40 |
||
120 | 41 |
qsum_def "A <+> B == ({0} <*> A) Un ({1} <*> B)" |
0 | 42 |
QInl_def "QInl(a) == <0;a>" |
43 |
QInr_def "QInr(b) == <1;b>" |
|
44 |
qcase_def "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))" |
|
45 |
end |
|
46 |
||
47 |
ML |
|
48 |
||
49 |
val print_translation = |
|
929
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
50 |
[("QSigma", dependent_tr' ("@QSUM", "op <*>"))]; |