author | paulson |
Fri, 17 Jul 1998 11:13:59 +0200 | |
changeset 5157 | 6e03de8ec2b4 |
parent 5147 | 825877190618 |
child 5268 | 59ef39008514 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/QPair.ML |
0 | 2 |
ID: $Id$ |
1461 | 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 |
Many proofs are borrowed from pair.ML and sum.ML |
|
14 |
||
15 |
Do we EVER have rank(a) < rank(<a;b>) ? Perhaps if the latter rank |
|
16 |
is not a limit ordinal? |
|
17 |
*) |
|
18 |
||
19 |
open QPair; |
|
20 |
||
21 |
(**** Quine ordered pairing ****) |
|
22 |
||
23 |
(** Lemmas for showing that <a;b> uniquely determines a and b **) |
|
24 |
||
2469 | 25 |
qed_goalw "QPair_empty" thy [QPair_def] |
26 |
"<0;0> = 0" |
|
27 |
(fn _=> [Simp_tac 1]); |
|
28 |
||
1096 | 29 |
qed_goalw "QPair_iff" thy [QPair_def] |
0 | 30 |
"<a;b> = <c;d> <-> a=c & b=d" |
31 |
(fn _=> [rtac sum_equal_iff 1]); |
|
32 |
||
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
33 |
bind_thm ("QPair_inject", (QPair_iff RS iffD1 RS conjE)); |
0 | 34 |
|
2469 | 35 |
Addsimps [QPair_empty, QPair_iff]; |
36 |
AddSEs [QPair_inject]; |
|
37 |
||
1096 | 38 |
qed_goal "QPair_inject1" thy "<a;b> = <c;d> ==> a=c" |
0 | 39 |
(fn [major]=> |
40 |
[ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); |
|
41 |
||
1096 | 42 |
qed_goal "QPair_inject2" thy "<a;b> = <c;d> ==> b=d" |
0 | 43 |
(fn [major]=> |
44 |
[ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); |
|
45 |
||
46 |
||
47 |
(*** QSigma: Disjoint union of a family of sets |
|
48 |
Generalizes Cartesian product ***) |
|
49 |
||
1096 | 50 |
qed_goalw "QSigmaI" thy [QSigma_def] |
2469 | 51 |
"!!A B. [| a:A; b:B(a) |] ==> <a;b> : QSigma(A,B)" |
3016 | 52 |
(fn _ => [ Blast_tac 1 ]); |
2469 | 53 |
|
54 |
AddSIs [QSigmaI]; |
|
0 | 55 |
|
56 |
(*The general elimination rule*) |
|
1096 | 57 |
qed_goalw "QSigmaE" thy [QSigma_def] |
0 | 58 |
"[| c: QSigma(A,B); \ |
59 |
\ !!x y.[| x:A; y:B(x); c=<x;y> |] ==> P \ |
|
60 |
\ |] ==> P" |
|
61 |
(fn major::prems=> |
|
62 |
[ (cut_facts_tac [major] 1), |
|
63 |
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); |
|
64 |
||
65 |
(** Elimination rules for <a;b>:A*B -- introducing no eigenvariables **) |
|
66 |
||
67 |
val QSigmaE2 = |
|
68 |
rule_by_tactic (REPEAT_FIRST (etac QPair_inject ORELSE' bound_hyp_subst_tac) |
|
1461 | 69 |
THEN prune_params_tac) |
0 | 70 |
(read_instantiate [("c","<a;b>")] QSigmaE); |
71 |
||
1096 | 72 |
qed_goal "QSigmaD1" thy "<a;b> : QSigma(A,B) ==> a : A" |
0 | 73 |
(fn [major]=> |
74 |
[ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); |
|
75 |
||
1096 | 76 |
qed_goal "QSigmaD2" thy "<a;b> : QSigma(A,B) ==> b : B(a)" |
0 | 77 |
(fn [major]=> |
78 |
[ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); |
|
79 |
||
2469 | 80 |
AddSEs [QSigmaE2, QSigmaE]; |
744 | 81 |
|
82 |
||
1096 | 83 |
qed_goalw "QSigma_cong" thy [QSigma_def] |
0 | 84 |
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ |
85 |
\ QSigma(A,B) = QSigma(A',B')" |
|
4091 | 86 |
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]); |
0 | 87 |
|
1096 | 88 |
qed_goal "QSigma_empty1" thy "QSigma(0,B) = 0" |
3016 | 89 |
(fn _ => [ (Blast_tac 1) ]); |
1096 | 90 |
|
91 |
qed_goal "QSigma_empty2" thy "A <*> 0 = 0" |
|
3016 | 92 |
(fn _ => [ (Blast_tac 1) ]); |
2469 | 93 |
|
94 |
Addsimps [QSigma_empty1, QSigma_empty2]; |
|
0 | 95 |
|
1096 | 96 |
|
97 |
(*** Projections: qfst, qsnd ***) |
|
98 |
||
99 |
qed_goalw "qfst_conv" thy [qfst_def] "qfst(<a;b>) = a" |
|
100 |
(fn _=> |
|
4091 | 101 |
[ (blast_tac (claset() addIs [the_equality]) 1) ]); |
1096 | 102 |
|
103 |
qed_goalw "qsnd_conv" thy [qsnd_def] "qsnd(<a;b>) = b" |
|
104 |
(fn _=> |
|
4091 | 105 |
[ (blast_tac (claset() addIs [the_equality]) 1) ]); |
1096 | 106 |
|
2469 | 107 |
Addsimps [qfst_conv, qsnd_conv]; |
1096 | 108 |
|
109 |
qed_goal "qfst_type" thy |
|
110 |
"!!p. p:QSigma(A,B) ==> qfst(p) : A" |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
111 |
(fn _=> [ Auto_tac ]); |
1096 | 112 |
|
113 |
qed_goal "qsnd_type" thy |
|
114 |
"!!p. p:QSigma(A,B) ==> qsnd(p) : B(qfst(p))" |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
115 |
(fn _=> [ Auto_tac ]); |
1096 | 116 |
|
5137 | 117 |
Goal "a: QSigma(A,B) ==> <qfst(a); qsnd(a)> = a"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
118 |
by Auto_tac; |
1096 | 119 |
qed "QPair_qfst_qsnd_eq"; |
0 | 120 |
|
121 |
||
122 |
(*** Eliminator - qsplit ***) |
|
123 |
||
1096 | 124 |
(*A META-equality, so that it applies to higher types as well...*) |
125 |
qed_goalw "qsplit" thy [qsplit_def] |
|
3840 | 126 |
"qsplit(%x y. c(x,y), <a;b>) == c(a,b)" |
2469 | 127 |
(fn _ => [ (Simp_tac 1), |
1096 | 128 |
(rtac reflexive_thm 1) ]); |
0 | 129 |
|
2469 | 130 |
Addsimps [qsplit]; |
131 |
||
1096 | 132 |
qed_goal "qsplit_type" thy |
0 | 133 |
"[| p:QSigma(A,B); \ |
134 |
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x;y>) \ |
|
3840 | 135 |
\ |] ==> qsplit(%x y. c(x,y), p) : C(p)" |
0 | 136 |
(fn major::prems=> |
137 |
[ (rtac (major RS QSigmaE) 1), |
|
4091 | 138 |
(asm_simp_tac (simpset() addsimps prems) 1) ]); |
1096 | 139 |
|
5067 | 140 |
Goalw [qsplit_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
141 |
"u: A<*>B ==> R(qsplit(c,u)) <-> (ALL x:A. ALL y:B. u = <x;y> --> R(c(x,y)))"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4091
diff
changeset
|
142 |
by Auto_tac; |
1096 | 143 |
qed "expand_qsplit"; |
144 |
||
145 |
||
146 |
(*** qsplit for predicates: result type o ***) |
|
147 |
||
5137 | 148 |
Goalw [qsplit_def] "R(a,b) ==> qsplit(R, <a;b>)"; |
2469 | 149 |
by (Asm_simp_tac 1); |
1096 | 150 |
qed "qsplitI"; |
151 |
||
152 |
val major::sigma::prems = goalw thy [qsplit_def] |
|
1461 | 153 |
"[| qsplit(R,z); z:QSigma(A,B); \ |
154 |
\ !!x y. [| z = <x;y>; R(x,y) |] ==> P \ |
|
1096 | 155 |
\ |] ==> P"; |
156 |
by (rtac (sigma RS QSigmaE) 1); |
|
157 |
by (cut_facts_tac [major] 1); |
|
2469 | 158 |
by (REPEAT (ares_tac prems 1)); |
159 |
by (Asm_full_simp_tac 1); |
|
1096 | 160 |
qed "qsplitE"; |
161 |
||
5137 | 162 |
Goalw [qsplit_def] "qsplit(R,<a;b>) ==> R(a,b)"; |
2469 | 163 |
by (Asm_full_simp_tac 1); |
1096 | 164 |
qed "qsplitD"; |
0 | 165 |
|
166 |
||
167 |
(*** qconverse ***) |
|
168 |
||
1096 | 169 |
qed_goalw "qconverseI" thy [qconverse_def] |
0 | 170 |
"!!a b r. <a;b>:r ==> <b;a>:qconverse(r)" |
3016 | 171 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 172 |
|
1096 | 173 |
qed_goalw "qconverseD" thy [qconverse_def] |
0 | 174 |
"!!a b r. <a;b> : qconverse(r) ==> <b;a> : r" |
3016 | 175 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 176 |
|
1096 | 177 |
qed_goalw "qconverseE" thy [qconverse_def] |
0 | 178 |
"[| yx : qconverse(r); \ |
179 |
\ !!x y. [| yx=<y;x>; <x;y>:r |] ==> P \ |
|
180 |
\ |] ==> P" |
|
181 |
(fn [major,minor]=> |
|
182 |
[ (rtac (major RS ReplaceE) 1), |
|
183 |
(REPEAT (eresolve_tac [exE, conjE, minor] 1)), |
|
184 |
(hyp_subst_tac 1), |
|
185 |
(assume_tac 1) ]); |
|
186 |
||
2469 | 187 |
AddSIs [qconverseI]; |
188 |
AddSEs [qconverseD, qconverseE]; |
|
0 | 189 |
|
1096 | 190 |
qed_goal "qconverse_qconverse" thy |
0 | 191 |
"!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" |
3016 | 192 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 193 |
|
1096 | 194 |
qed_goal "qconverse_type" thy |
0 | 195 |
"!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A" |
3016 | 196 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 197 |
|
1096 | 198 |
qed_goal "qconverse_prod" thy "qconverse(A <*> B) = B <*> A" |
3016 | 199 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 200 |
|
1096 | 201 |
qed_goal "qconverse_empty" thy "qconverse(0) = 0" |
3016 | 202 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 203 |
|
204 |
||
205 |
(**** The Quine-inspired notion of disjoint sum ****) |
|
206 |
||
207 |
val qsum_defs = [qsum_def,QInl_def,QInr_def,qcase_def]; |
|
208 |
||
209 |
(** Introduction rules for the injections **) |
|
210 |
||
5137 | 211 |
Goalw qsum_defs "a : A ==> QInl(a) : A <+> B"; |
3016 | 212 |
by (Blast_tac 1); |
760 | 213 |
qed "QInlI"; |
0 | 214 |
|
5137 | 215 |
Goalw qsum_defs "b : B ==> QInr(b) : A <+> B"; |
3016 | 216 |
by (Blast_tac 1); |
760 | 217 |
qed "QInrI"; |
0 | 218 |
|
219 |
(** Elimination rules **) |
|
220 |
||
1096 | 221 |
val major::prems = goalw thy qsum_defs |
0 | 222 |
"[| u: A <+> B; \ |
223 |
\ !!x. [| x:A; u=QInl(x) |] ==> P; \ |
|
224 |
\ !!y. [| y:B; u=QInr(y) |] ==> P \ |
|
225 |
\ |] ==> P"; |
|
226 |
by (rtac (major RS UnE) 1); |
|
227 |
by (REPEAT (rtac refl 1 |
|
228 |
ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1)); |
|
760 | 229 |
qed "qsumE"; |
0 | 230 |
|
2469 | 231 |
AddSIs [QInlI, QInrI]; |
232 |
||
0 | 233 |
(** Injection and freeness equivalences, for rewriting **) |
234 |
||
5067 | 235 |
Goalw qsum_defs "QInl(a)=QInl(b) <-> a=b"; |
2469 | 236 |
by (Simp_tac 1); |
760 | 237 |
qed "QInl_iff"; |
0 | 238 |
|
5067 | 239 |
Goalw qsum_defs "QInr(a)=QInr(b) <-> a=b"; |
2469 | 240 |
by (Simp_tac 1); |
760 | 241 |
qed "QInr_iff"; |
0 | 242 |
|
5067 | 243 |
Goalw qsum_defs "QInl(a)=QInr(b) <-> False"; |
2469 | 244 |
by (Simp_tac 1); |
760 | 245 |
qed "QInl_QInr_iff"; |
0 | 246 |
|
5067 | 247 |
Goalw qsum_defs "QInr(b)=QInl(a) <-> False"; |
2469 | 248 |
by (Simp_tac 1); |
760 | 249 |
qed "QInr_QInl_iff"; |
0 | 250 |
|
5067 | 251 |
Goalw qsum_defs "0<+>0 = 0"; |
2469 | 252 |
by (Simp_tac 1); |
253 |
qed "qsum_empty"; |
|
254 |
||
55 | 255 |
(*Injection and freeness rules*) |
256 |
||
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
257 |
bind_thm ("QInl_inject", (QInl_iff RS iffD1)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
258 |
bind_thm ("QInr_inject", (QInr_iff RS iffD1)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
259 |
bind_thm ("QInl_neq_QInr", (QInl_QInr_iff RS iffD1 RS FalseE)); |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
260 |
bind_thm ("QInr_neq_QInl", (QInr_QInl_iff RS iffD1 RS FalseE)); |
55 | 261 |
|
2469 | 262 |
AddSEs [qsumE, QInl_neq_QInr, QInr_neq_QInl]; |
263 |
AddSDs [QInl_inject, QInr_inject]; |
|
264 |
Addsimps [QInl_iff, QInr_iff, QInl_QInr_iff, QInr_QInl_iff, qsum_empty]; |
|
0 | 265 |
|
5137 | 266 |
Goal "QInl(a): A<+>B ==> a: A"; |
3016 | 267 |
by (Blast_tac 1); |
760 | 268 |
qed "QInlD"; |
55 | 269 |
|
5137 | 270 |
Goal "QInr(b): A<+>B ==> b: B"; |
3016 | 271 |
by (Blast_tac 1); |
760 | 272 |
qed "QInrD"; |
55 | 273 |
|
0 | 274 |
(** <+> is itself injective... who cares?? **) |
275 |
||
5067 | 276 |
Goal |
0 | 277 |
"u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"; |
3016 | 278 |
by (Blast_tac 1); |
760 | 279 |
qed "qsum_iff"; |
0 | 280 |
|
5067 | 281 |
Goal "A <+> B <= C <+> D <-> A<=C & B<=D"; |
3016 | 282 |
by (Blast_tac 1); |
760 | 283 |
qed "qsum_subset_iff"; |
0 | 284 |
|
5067 | 285 |
Goal "A <+> B = C <+> D <-> A=C & B=D"; |
4091 | 286 |
by (simp_tac (simpset() addsimps [extension,qsum_subset_iff]) 1); |
3016 | 287 |
by (Blast_tac 1); |
760 | 288 |
qed "qsum_equal_iff"; |
0 | 289 |
|
290 |
(*** Eliminator -- qcase ***) |
|
291 |
||
5067 | 292 |
Goalw qsum_defs "qcase(c, d, QInl(a)) = c(a)"; |
2469 | 293 |
by (Simp_tac 1); |
760 | 294 |
qed "qcase_QInl"; |
0 | 295 |
|
5067 | 296 |
Goalw qsum_defs "qcase(c, d, QInr(b)) = d(b)"; |
2469 | 297 |
by (Simp_tac 1); |
760 | 298 |
qed "qcase_QInr"; |
0 | 299 |
|
2469 | 300 |
Addsimps [qcase_QInl, qcase_QInr]; |
301 |
||
1096 | 302 |
val major::prems = goal thy |
0 | 303 |
"[| u: A <+> B; \ |
304 |
\ !!x. x: A ==> c(x): C(QInl(x)); \ |
|
305 |
\ !!y. y: B ==> d(y): C(QInr(y)) \ |
|
306 |
\ |] ==> qcase(c,d,u) : C(u)"; |
|
307 |
by (rtac (major RS qsumE) 1); |
|
308 |
by (ALLGOALS (etac ssubst)); |
|
4091 | 309 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); |
760 | 310 |
qed "qcase_type"; |
0 | 311 |
|
312 |
(** Rules for the Part primitive **) |
|
313 |
||
5067 | 314 |
Goal "Part(A <+> B,QInl) = {QInl(x). x: A}"; |
3016 | 315 |
by (Blast_tac 1); |
760 | 316 |
qed "Part_QInl"; |
0 | 317 |
|
5067 | 318 |
Goal "Part(A <+> B,QInr) = {QInr(y). y: B}"; |
3016 | 319 |
by (Blast_tac 1); |
760 | 320 |
qed "Part_QInr"; |
0 | 321 |
|
5067 | 322 |
Goal "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y: Part(B,h)}"; |
3016 | 323 |
by (Blast_tac 1); |
760 | 324 |
qed "Part_QInr2"; |
0 | 325 |
|
5137 | 326 |
Goal "C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; |
3016 | 327 |
by (Blast_tac 1); |
760 | 328 |
qed "Part_qsum_equality"; |