| author | wenzelm | 
| Fri, 21 May 1999 16:26:06 +0200 | |
| changeset 6699 | 1471f2bd74a0 | 
| parent 6112 | 5e4871c5136b | 
| child 8551 | 5c22595bc599 | 
| 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: 
760diff
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 _=> | |
| 5505 | 101 | [ (Blast_tac 1) ]); | 
| 1096 | 102 | |
| 103 | qed_goalw "qsnd_conv" thy [qsnd_def] "qsnd(<a;b>) = b" | |
| 104 | (fn _=> | |
| 5505 | 105 | [ (Blast_tac 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: 
4091diff
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: 
4091diff
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: 
4091diff
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...*) | 
| 6112 | 125 | Goalw [qsplit_def] "qsplit(%x y. c(x,y), <a;b>) == c(a,b)"; | 
| 126 | by (Simp_tac 1); | |
| 127 | qed "qsplit"; | |
| 2469 | 128 | Addsimps [qsplit]; | 
| 129 | ||
| 1096 | 130 | qed_goal "qsplit_type" thy | 
| 0 | 131 | "[| p:QSigma(A,B); \ | 
| 132 | \ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x;y>) \ | |
| 3840 | 133 | \ |] ==> qsplit(%x y. c(x,y), p) : C(p)" | 
| 0 | 134 | (fn major::prems=> | 
| 135 | [ (rtac (major RS QSigmaE) 1), | |
| 4091 | 136 | (asm_simp_tac (simpset() addsimps prems) 1) ]); | 
| 1096 | 137 | |
| 5067 | 138 | Goalw [qsplit_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 139 | "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: 
4091diff
changeset | 140 | by Auto_tac; | 
| 1096 | 141 | qed "expand_qsplit"; | 
| 142 | ||
| 143 | ||
| 144 | (*** qsplit for predicates: result type o ***) | |
| 145 | ||
| 5137 | 146 | Goalw [qsplit_def] "R(a,b) ==> qsplit(R, <a;b>)"; | 
| 2469 | 147 | by (Asm_simp_tac 1); | 
| 1096 | 148 | qed "qsplitI"; | 
| 149 | ||
| 150 | val major::sigma::prems = goalw thy [qsplit_def] | |
| 1461 | 151 | "[| qsplit(R,z); z:QSigma(A,B); \ | 
| 152 | \ !!x y. [| z = <x;y>; R(x,y) |] ==> P \ | |
| 1096 | 153 | \ |] ==> P"; | 
| 154 | by (rtac (sigma RS QSigmaE) 1); | |
| 155 | by (cut_facts_tac [major] 1); | |
| 2469 | 156 | by (REPEAT (ares_tac prems 1)); | 
| 157 | by (Asm_full_simp_tac 1); | |
| 1096 | 158 | qed "qsplitE"; | 
| 159 | ||
| 5137 | 160 | Goalw [qsplit_def] "qsplit(R,<a;b>) ==> R(a,b)"; | 
| 2469 | 161 | by (Asm_full_simp_tac 1); | 
| 1096 | 162 | qed "qsplitD"; | 
| 0 | 163 | |
| 164 | ||
| 165 | (*** qconverse ***) | |
| 166 | ||
| 1096 | 167 | qed_goalw "qconverseI" thy [qconverse_def] | 
| 0 | 168 | "!!a b r. <a;b>:r ==> <b;a>:qconverse(r)" | 
| 3016 | 169 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 170 | |
| 1096 | 171 | qed_goalw "qconverseD" thy [qconverse_def] | 
| 0 | 172 | "!!a b r. <a;b> : qconverse(r) ==> <b;a> : r" | 
| 3016 | 173 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 174 | |
| 1096 | 175 | qed_goalw "qconverseE" thy [qconverse_def] | 
| 0 | 176 | "[| yx : qconverse(r); \ | 
| 177 | \ !!x y. [| yx=<y;x>; <x;y>:r |] ==> P \ | |
| 178 | \ |] ==> P" | |
| 179 | (fn [major,minor]=> | |
| 180 | [ (rtac (major RS ReplaceE) 1), | |
| 181 | (REPEAT (eresolve_tac [exE, conjE, minor] 1)), | |
| 182 | (hyp_subst_tac 1), | |
| 183 | (assume_tac 1) ]); | |
| 184 | ||
| 2469 | 185 | AddSIs [qconverseI]; | 
| 186 | AddSEs [qconverseD, qconverseE]; | |
| 0 | 187 | |
| 1096 | 188 | qed_goal "qconverse_qconverse" thy | 
| 0 | 189 | "!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" | 
| 3016 | 190 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 191 | |
| 1096 | 192 | qed_goal "qconverse_type" thy | 
| 0 | 193 | "!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A" | 
| 3016 | 194 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 195 | |
| 1096 | 196 | qed_goal "qconverse_prod" thy "qconverse(A <*> B) = B <*> A" | 
| 3016 | 197 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 198 | |
| 1096 | 199 | qed_goal "qconverse_empty" thy "qconverse(0) = 0" | 
| 3016 | 200 | (fn _ => [ (Blast_tac 1) ]); | 
| 0 | 201 | |
| 202 | ||
| 203 | (**** The Quine-inspired notion of disjoint sum ****) | |
| 204 | ||
| 205 | val qsum_defs = [qsum_def,QInl_def,QInr_def,qcase_def]; | |
| 206 | ||
| 207 | (** Introduction rules for the injections **) | |
| 208 | ||
| 5137 | 209 | Goalw qsum_defs "a : A ==> QInl(a) : A <+> B"; | 
| 3016 | 210 | by (Blast_tac 1); | 
| 760 | 211 | qed "QInlI"; | 
| 0 | 212 | |
| 5137 | 213 | Goalw qsum_defs "b : B ==> QInr(b) : A <+> B"; | 
| 3016 | 214 | by (Blast_tac 1); | 
| 760 | 215 | qed "QInrI"; | 
| 0 | 216 | |
| 217 | (** Elimination rules **) | |
| 218 | ||
| 1096 | 219 | val major::prems = goalw thy qsum_defs | 
| 0 | 220 | "[| u: A <+> B; \ | 
| 221 | \ !!x. [| x:A; u=QInl(x) |] ==> P; \ | |
| 222 | \ !!y. [| y:B; u=QInr(y) |] ==> P \ | |
| 223 | \ |] ==> P"; | |
| 224 | by (rtac (major RS UnE) 1); | |
| 225 | by (REPEAT (rtac refl 1 | |
| 226 | ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1)); | |
| 760 | 227 | qed "qsumE"; | 
| 0 | 228 | |
| 2469 | 229 | AddSIs [QInlI, QInrI]; | 
| 230 | ||
| 0 | 231 | (** Injection and freeness equivalences, for rewriting **) | 
| 232 | ||
| 5067 | 233 | Goalw qsum_defs "QInl(a)=QInl(b) <-> a=b"; | 
| 2469 | 234 | by (Simp_tac 1); | 
| 760 | 235 | qed "QInl_iff"; | 
| 0 | 236 | |
| 5067 | 237 | Goalw qsum_defs "QInr(a)=QInr(b) <-> a=b"; | 
| 2469 | 238 | by (Simp_tac 1); | 
| 760 | 239 | qed "QInr_iff"; | 
| 0 | 240 | |
| 5067 | 241 | Goalw qsum_defs "QInl(a)=QInr(b) <-> False"; | 
| 2469 | 242 | by (Simp_tac 1); | 
| 760 | 243 | qed "QInl_QInr_iff"; | 
| 0 | 244 | |
| 5067 | 245 | Goalw qsum_defs "QInr(b)=QInl(a) <-> False"; | 
| 2469 | 246 | by (Simp_tac 1); | 
| 760 | 247 | qed "QInr_QInl_iff"; | 
| 0 | 248 | |
| 5067 | 249 | Goalw qsum_defs "0<+>0 = 0"; | 
| 2469 | 250 | by (Simp_tac 1); | 
| 251 | qed "qsum_empty"; | |
| 252 | ||
| 55 | 253 | (*Injection and freeness rules*) | 
| 254 | ||
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 255 | bind_thm ("QInl_inject", (QInl_iff RS iffD1));
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 256 | bind_thm ("QInr_inject", (QInr_iff RS iffD1));
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 257 | bind_thm ("QInl_neq_QInr", (QInl_QInr_iff RS iffD1 RS FalseE));
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 258 | bind_thm ("QInr_neq_QInl", (QInr_QInl_iff RS iffD1 RS FalseE));
 | 
| 55 | 259 | |
| 2469 | 260 | AddSEs [qsumE, QInl_neq_QInr, QInr_neq_QInl]; | 
| 261 | AddSDs [QInl_inject, QInr_inject]; | |
| 262 | Addsimps [QInl_iff, QInr_iff, QInl_QInr_iff, QInr_QInl_iff, qsum_empty]; | |
| 0 | 263 | |
| 5137 | 264 | Goal "QInl(a): A<+>B ==> a: A"; | 
| 3016 | 265 | by (Blast_tac 1); | 
| 760 | 266 | qed "QInlD"; | 
| 55 | 267 | |
| 5137 | 268 | Goal "QInr(b): A<+>B ==> b: B"; | 
| 3016 | 269 | by (Blast_tac 1); | 
| 760 | 270 | qed "QInrD"; | 
| 55 | 271 | |
| 0 | 272 | (** <+> is itself injective... who cares?? **) | 
| 273 | ||
| 5268 | 274 | Goal "u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"; | 
| 3016 | 275 | by (Blast_tac 1); | 
| 760 | 276 | qed "qsum_iff"; | 
| 0 | 277 | |
| 5067 | 278 | Goal "A <+> B <= C <+> D <-> A<=C & B<=D"; | 
| 3016 | 279 | by (Blast_tac 1); | 
| 760 | 280 | qed "qsum_subset_iff"; | 
| 0 | 281 | |
| 5067 | 282 | Goal "A <+> B = C <+> D <-> A=C & B=D"; | 
| 4091 | 283 | by (simp_tac (simpset() addsimps [extension,qsum_subset_iff]) 1); | 
| 3016 | 284 | by (Blast_tac 1); | 
| 760 | 285 | qed "qsum_equal_iff"; | 
| 0 | 286 | |
| 287 | (*** Eliminator -- qcase ***) | |
| 288 | ||
| 5067 | 289 | Goalw qsum_defs "qcase(c, d, QInl(a)) = c(a)"; | 
| 2469 | 290 | by (Simp_tac 1); | 
| 760 | 291 | qed "qcase_QInl"; | 
| 0 | 292 | |
| 5067 | 293 | Goalw qsum_defs "qcase(c, d, QInr(b)) = d(b)"; | 
| 2469 | 294 | by (Simp_tac 1); | 
| 760 | 295 | qed "qcase_QInr"; | 
| 0 | 296 | |
| 2469 | 297 | Addsimps [qcase_QInl, qcase_QInr]; | 
| 298 | ||
| 1096 | 299 | val major::prems = goal thy | 
| 0 | 300 | "[| u: A <+> B; \ | 
| 301 | \ !!x. x: A ==> c(x): C(QInl(x)); \ | |
| 302 | \ !!y. y: B ==> d(y): C(QInr(y)) \ | |
| 303 | \ |] ==> qcase(c,d,u) : C(u)"; | |
| 304 | by (rtac (major RS qsumE) 1); | |
| 305 | by (ALLGOALS (etac ssubst)); | |
| 4091 | 306 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 307 | qed "qcase_type"; | 
| 0 | 308 | |
| 309 | (** Rules for the Part primitive **) | |
| 310 | ||
| 5067 | 311 | Goal "Part(A <+> B,QInl) = {QInl(x). x: A}";
 | 
| 3016 | 312 | by (Blast_tac 1); | 
| 760 | 313 | qed "Part_QInl"; | 
| 0 | 314 | |
| 5067 | 315 | Goal "Part(A <+> B,QInr) = {QInr(y). y: B}";
 | 
| 3016 | 316 | by (Blast_tac 1); | 
| 760 | 317 | qed "Part_QInr"; | 
| 0 | 318 | |
| 5067 | 319 | Goal "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y: Part(B,h)}";
 | 
| 3016 | 320 | by (Blast_tac 1); | 
| 760 | 321 | qed "Part_QInr2"; | 
| 0 | 322 | |
| 5137 | 323 | Goal "C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; | 
| 3016 | 324 | by (Blast_tac 1); | 
| 760 | 325 | qed "Part_qsum_equality"; |