0
|
1 |
(* Title: ZF/zf.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory
|
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
|
|
6 |
Zermelo-Fraenkel Set Theory
|
|
7 |
*)
|
|
8 |
|
|
9 |
ZF = FOL +
|
|
10 |
|
|
11 |
types
|
49
|
12 |
i, is 0
|
0
|
13 |
|
|
14 |
arities
|
|
15 |
i :: term
|
|
16 |
|
|
17 |
|
|
18 |
consts
|
|
19 |
|
80
|
20 |
"0" :: "i" ("0") (*the empty set*)
|
|
21 |
Pow :: "i => i" (*power sets*)
|
|
22 |
Inf :: "i" (*infinite set*)
|
0
|
23 |
|
|
24 |
(* Bounded Quantifiers *)
|
|
25 |
|
80
|
26 |
"@Ball" :: "[idt, i, o] => o" ("(3ALL _:_./ _)" 10)
|
|
27 |
"@Bex" :: "[idt, i, o] => o" ("(3EX _:_./ _)" 10)
|
|
28 |
Ball :: "[i, i => o] => o"
|
|
29 |
Bex :: "[i, i => o] => o"
|
0
|
30 |
|
|
31 |
(* General Union and Intersection *)
|
|
32 |
|
80
|
33 |
"@INTER" :: "[idt, i, i] => i" ("(3INT _:_./ _)" 10)
|
|
34 |
"@UNION" :: "[idt, i, i] => i" ("(3UN _:_./ _)" 10)
|
|
35 |
Union, Inter :: "i => i"
|
0
|
36 |
|
|
37 |
(* Variations on Replacement *)
|
|
38 |
|
80
|
39 |
"@Replace" :: "[idt, idt, i, o] => i" ("(1{_ ./ _: _, _})")
|
|
40 |
"@RepFun" :: "[i, idt, i] => i" ("(1{_ ./ _: _})")
|
|
41 |
"@Collect" :: "[idt, i, o] => i" ("(1{_: _ ./ _})")
|
|
42 |
PrimReplace :: "[i, [i, i] => o] => i"
|
|
43 |
Replace :: "[i, [i, i] => o] => i"
|
|
44 |
RepFun :: "[i, i => i] => i"
|
|
45 |
Collect :: "[i, i => o] => i"
|
0
|
46 |
|
|
47 |
(* Descriptions *)
|
|
48 |
|
80
|
49 |
The :: "(i => o) => i" (binder "THE " 10)
|
|
50 |
if :: "[o, i, i] => i"
|
0
|
51 |
|
|
52 |
(* Enumerations of type i *)
|
|
53 |
|
80
|
54 |
"" :: "i => is" ("_")
|
|
55 |
"@Enum" :: "[i, is] => is" ("_,/ _")
|
0
|
56 |
|
|
57 |
(* Finite Sets *)
|
|
58 |
|
80
|
59 |
"@Finset" :: "is => i" ("{(_)}")
|
|
60 |
Upair, cons :: "[i, i] => i"
|
|
61 |
succ :: "i => i"
|
0
|
62 |
|
|
63 |
(* Ordered Pairing and n-Tuples *)
|
|
64 |
|
80
|
65 |
"@Tuple" :: "[i, is] => i" ("<(_,/ _)>")
|
|
66 |
Pair :: "[i, i] => i"
|
|
67 |
fst, snd :: "i => i"
|
|
68 |
split :: "[[i, i] => i, i] => i"
|
|
69 |
fsplit :: "[[i, i] => o, i] => o"
|
0
|
70 |
|
|
71 |
(* Sigma and Pi Operators *)
|
|
72 |
|
80
|
73 |
"@PROD" :: "[idt, i, i] => i" ("(3PROD _:_./ _)" 10)
|
|
74 |
"@SUM" :: "[idt, i, i] => i" ("(3SUM _:_./ _)" 10)
|
|
75 |
"@lam" :: "[idt, i, i] => i" ("(3lam _:_./ _)" 10)
|
|
76 |
Pi, Sigma :: "[i, i => i] => i"
|
0
|
77 |
|
|
78 |
(* Relations and Functions *)
|
|
79 |
|
80
|
80 |
domain :: "i => i"
|
|
81 |
range :: "i => i"
|
|
82 |
field :: "i => i"
|
|
83 |
converse :: "i => i"
|
|
84 |
Lambda :: "[i, i => i] => i"
|
|
85 |
restrict :: "[i, i] => i"
|
0
|
86 |
|
|
87 |
(* Infixes in order of decreasing precedence *)
|
|
88 |
|
80
|
89 |
"``" :: "[i, i] => i" (infixl 90) (*image*)
|
|
90 |
"-``" :: "[i, i] => i" (infixl 90) (*inverse image*)
|
|
91 |
"`" :: "[i, i] => i" (infixl 90) (*function application*)
|
0
|
92 |
|
80
|
93 |
(*Except for their translations, * and -> are right and ~: left associative infixes*)
|
|
94 |
" *" :: "[i, i] => i" ("(_ */ _)" [81, 80] 80) (*Cartesian product*)
|
|
95 |
"Int" :: "[i, i] => i" (infixl 70) (*binary intersection*)
|
|
96 |
"Un" :: "[i, i] => i" (infixl 65) (*binary union*)
|
|
97 |
"-" :: "[i, i] => i" (infixl 65) (*set difference*)
|
|
98 |
" ->" :: "[i, i] => i" ("(_ ->/ _)" [61, 60] 60) (*function space*)
|
|
99 |
"<=" :: "[i, i] => o" (infixl 50) (*subset relation*)
|
|
100 |
":" :: "[i, i] => o" (infixl 50) (*membership relation*)
|
|
101 |
"~:" :: "[i, i] => o" ("(_ ~:/ _)" [50, 51] 50) (*negated membership relation*)
|
0
|
102 |
|
|
103 |
|
|
104 |
translations
|
|
105 |
"{x, xs}" == "cons(x, {xs})"
|
|
106 |
"{x}" == "cons(x, 0)"
|
49
|
107 |
"<x, y, z>" == "<x, <y, z>>"
|
|
108 |
"<x, y>" == "Pair(x, y)"
|
0
|
109 |
"{x:A. P}" == "Collect(A, %x. P)"
|
|
110 |
"{y. x:A, Q}" == "Replace(A, %x y. Q)"
|
|
111 |
"{f. x:A}" == "RepFun(A, %x. f)"
|
|
112 |
"INT x:A. B" == "Inter({B. x:A})"
|
|
113 |
"UN x:A. B" == "Union({B. x:A})"
|
|
114 |
"PROD x:A. B" => "Pi(A, %x. B)"
|
|
115 |
"SUM x:A. B" => "Sigma(A, %x. B)"
|
49
|
116 |
"A -> B" => "Pi(A, _K(B))"
|
|
117 |
"A * B" => "Sigma(A, _K(B))"
|
0
|
118 |
"lam x:A. f" == "Lambda(A, %x. f)"
|
|
119 |
"ALL x:A. P" == "Ball(A, %x. P)"
|
|
120 |
"EX x:A. P" == "Bex(A, %x. P)"
|
80
|
121 |
"x ~: y" == "~ (x : y)"
|
37
|
122 |
|
0
|
123 |
|
|
124 |
rules
|
|
125 |
|
|
126 |
(* Bounded Quantifiers *)
|
|
127 |
Ball_def "Ball(A,P) == ALL x. x:A --> P(x)"
|
|
128 |
Bex_def "Bex(A,P) == EX x. x:A & P(x)"
|
|
129 |
subset_def "A <= B == ALL x:A. x:B"
|
|
130 |
|
|
131 |
(* ZF axioms -- see Suppes p.238
|
|
132 |
Axioms for Union, Pow and Replace state existence only,
|
|
133 |
uniqueness is derivable using extensionality. *)
|
|
134 |
|
|
135 |
extension "A = B <-> A <= B & B <= A"
|
|
136 |
union_iff "A : Union(C) <-> (EX B:C. A:B)"
|
|
137 |
power_set "A : Pow(B) <-> A <= B"
|
|
138 |
succ_def "succ(i) == cons(i,i)"
|
|
139 |
|
|
140 |
(*We may name this set, though it is not uniquely defined. *)
|
|
141 |
infinity "0:Inf & (ALL y:Inf. succ(y): Inf)"
|
|
142 |
|
|
143 |
(*This formulation facilitates case analysis on A. *)
|
37
|
144 |
foundation "A=0 | (EX x:A. ALL y:x. y~:A)"
|
0
|
145 |
|
|
146 |
(* Schema axiom since predicate P is a higher-order variable *)
|
|
147 |
replacement "(ALL x:A. ALL y z. P(x,y) & P(x,z) --> y=z) ==> \
|
|
148 |
\ b : PrimReplace(A,P) <-> (EX x:A. P(x,b))"
|
|
149 |
|
|
150 |
(* Derived form of replacement, restricting P to its functional part.
|
|
151 |
The resulting set (for functional P) is the same as with
|
|
152 |
PrimReplace, but the rules are simpler. *)
|
|
153 |
Replace_def "Replace(A,P) == PrimReplace(A, %x y. (EX!z.P(x,z)) & P(x,y))"
|
|
154 |
|
|
155 |
(* Functional form of replacement -- analgous to ML's map functional *)
|
|
156 |
RepFun_def "RepFun(A,f) == {y . x:A, y=f(x)}"
|
|
157 |
|
|
158 |
(* Separation and Pairing can be derived from the Replacement
|
|
159 |
and Powerset Axioms using the following definitions. *)
|
|
160 |
|
|
161 |
Collect_def "Collect(A,P) == {y . x:A, x=y & P(x)}"
|
|
162 |
|
|
163 |
(*Unordered pairs (Upair) express binary union/intersection and cons;
|
|
164 |
set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...) *)
|
|
165 |
Upair_def "Upair(a,b) == {y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
|
|
166 |
cons_def "cons(a,A) == Upair(a,a) Un A"
|
|
167 |
|
|
168 |
(* Difference, general intersection, binary union and small intersection *)
|
|
169 |
|
|
170 |
Diff_def "A - B == { x:A . ~(x:B) }"
|
|
171 |
Inter_def "Inter(A) == { x:Union(A) . ALL y:A. x:y}"
|
|
172 |
Un_def "A Un B == Union(Upair(A,B))"
|
|
173 |
Int_def "A Int B == Inter(Upair(A,B))"
|
|
174 |
|
|
175 |
(* Definite descriptions -- via Replace over the set "1" *)
|
|
176 |
|
|
177 |
the_def "The(P) == Union({y . x:{0}, P(y)})"
|
|
178 |
if_def "if(P,a,b) == THE z. P & z=a | ~P & z=b"
|
|
179 |
|
|
180 |
(* Ordered pairs and disjoint union of a family of sets *)
|
|
181 |
|
|
182 |
(* this "symmetric" definition works better than {{a}, {a,b}} *)
|
|
183 |
Pair_def "<a,b> == {{a,a}, {a,b}}"
|
|
184 |
fst_def "fst == split(%x y.x)"
|
|
185 |
snd_def "snd == split(%x y.y)"
|
|
186 |
split_def "split(c,p) == THE y. EX a b. p=<a,b> & y=c(a,b)"
|
|
187 |
fsplit_def "fsplit(R,z) == EX x y. z=<x,y> & R(x,y)"
|
|
188 |
Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {<x,y>}"
|
|
189 |
|
|
190 |
(* Operations on relations *)
|
|
191 |
|
|
192 |
(*converse of relation r, inverse of function*)
|
|
193 |
converse_def "converse(r) == {z. w:r, EX x y. w=<x,y> & z=<y,x>}"
|
|
194 |
|
|
195 |
domain_def "domain(r) == {x. w:r, EX y. w=<x,y>}"
|
|
196 |
range_def "range(r) == domain(converse(r))"
|
|
197 |
field_def "field(r) == domain(r) Un range(r)"
|
|
198 |
image_def "r `` A == {y : range(r) . EX x:A. <x,y> : r}"
|
|
199 |
vimage_def "r -`` A == converse(r)``A"
|
|
200 |
|
|
201 |
(* Abstraction, application and Cartesian product of a family of sets *)
|
|
202 |
|
|
203 |
lam_def "Lambda(A,b) == {<x,b(x)> . x:A}"
|
|
204 |
apply_def "f`a == THE y. <a,y> : f"
|
|
205 |
Pi_def "Pi(A,B) == {f: Pow(Sigma(A,B)). ALL x:A. EX! y. <x,y>: f}"
|
|
206 |
|
|
207 |
(* Restrict the function f to the domain A *)
|
|
208 |
restrict_def "restrict(f,A) == lam x:A.f`x"
|
|
209 |
|
|
210 |
end
|
|
211 |
|
|
212 |
|
|
213 |
ML
|
|
214 |
|
|
215 |
(* 'Dependent' type operators *)
|
|
216 |
|
|
217 |
val print_translation =
|
|
218 |
[("Pi", dependent_tr' ("@PROD", " ->")),
|
|
219 |
("Sigma", dependent_tr' ("@SUM", " *"))];
|
49
|
220 |
|