| author | wenzelm | 
| Thu, 11 Sep 1997 16:16:03 +0200 | |
| changeset 3669 | 3384c6f1f095 | 
| parent 2469 | b50b8c0eec01 | 
| child 3923 | c257b82a1200 | 
| permissions | -rw-r--r-- | 
| 1478 | 1  | 
(* Title: ZF/qpair.thy  | 
| 0 | 2  | 
ID: $Id$  | 
| 1478 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 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  | 
||
| 2469 | 14  | 
QPair = Sum +  | 
| 0 | 15  | 
consts  | 
| 1478 | 16  | 
  QPair     :: [i, i] => i                      ("<(_;/ _)>")
 | 
| 1401 | 17  | 
qfst,qsnd :: i => i  | 
18  | 
qsplit :: [[i, i] => 'a, i] => 'a::logic (*for pattern-matching*)  | 
|
19  | 
qconverse :: i => i  | 
|
20  | 
QSigma :: [i, i => i] => i  | 
|
| 0 | 21  | 
|
| 1478 | 22  | 
"<+>" :: [i,i]=>i (infixr 65)  | 
| 1401 | 23  | 
QInl,QInr :: i=>i  | 
24  | 
qcase :: [i=>i, i=>i, i]=>i  | 
|
| 0 | 25  | 
|
| 
929
 
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
 
lcp 
parents: 
753 
diff
changeset
 | 
26  | 
syntax  | 
| 1401 | 27  | 
  "@QSUM"   :: [idt, i, i] => i               ("(3QSUM _:_./ _)" 10)
 | 
| 1478 | 28  | 
"<*>" :: [i, i] => i (infixr 80)  | 
| 
929
 
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"  | 
| 
1097
 
01379c46ad2d
Changed definitions so that qsplit is now defined in terms of
 
lcp 
parents: 
929 
diff
changeset
 | 
36  | 
qfst_def "qfst(p) == THE a. EX b. p=<a;b>"  | 
| 
 
01379c46ad2d
Changed definitions so that qsplit is now defined in terms of
 
lcp 
parents: 
929 
diff
changeset
 | 
37  | 
qsnd_def "qsnd(p) == THE b. EX a. p=<a;b>"  | 
| 
 
01379c46ad2d
Changed definitions so that qsplit is now defined in terms of
 
lcp 
parents: 
929 
diff
changeset
 | 
38  | 
qsplit_def "qsplit(c,p) == c(qfst(p), qsnd(p))"  | 
| 
 
01379c46ad2d
Changed definitions so that qsplit is now defined in terms of
 
lcp 
parents: 
929 
diff
changeset
 | 
39  | 
|
| 0 | 40  | 
  qconverse_def   "qconverse(r) == {z. w:r, EX x y. w=<x;y> & z=<y;x>}"
 | 
41  | 
  QSigma_def      "QSigma(A,B)  ==  UN x:A. UN y:B(x). {<x;y>}"
 | 
|
42  | 
||
| 120 | 43  | 
  qsum_def        "A <+> B      == ({0} <*> A) Un ({1} <*> B)"
 | 
| 0 | 44  | 
QInl_def "QInl(a) == <0;a>"  | 
45  | 
QInr_def "QInr(b) == <1;b>"  | 
|
46  | 
qcase_def "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))"  | 
|
47  | 
end  | 
|
48  | 
||
49  | 
ML  | 
|
50  | 
||
51  | 
val print_translation =  | 
|
| 
929
 
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
 
lcp 
parents: 
753 
diff
changeset
 | 
52  | 
  [("QSigma", dependent_tr' ("@QSUM", "op <*>"))];
 |