| author | wenzelm | 
| Thu, 08 Nov 2001 23:52:56 +0100 | |
| changeset 12106 | 4a8558dbb6a0 | 
| parent 9211 | 6236c5285bd8 | 
| child 12199 | 8213fd95acb5 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/pair | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 6 | Ordered pairs in Zermelo-Fraenkel Set Theory | |
| 7 | *) | |
| 8 | ||
| 9 | (** Lemmas for showing that <a,b> uniquely determines a and b **) | |
| 10 | ||
| 9180 | 11 | Goal "{a} = {b} <-> a=b";
 | 
| 12 | by (resolve_tac [extension RS iff_trans] 1); | |
| 13 | by (Blast_tac 1) ; | |
| 14 | qed "singleton_eq_iff"; | |
| 822 
8d5748202c53
Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
 lcp parents: 
782diff
changeset | 15 | |
| 9180 | 16 | Goal "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)";
 | 
| 17 | by (resolve_tac [extension RS iff_trans] 1); | |
| 18 | by (Blast_tac 1) ; | |
| 19 | qed "doubleton_eq_iff"; | |
| 0 | 20 | |
| 9211 | 21 | Goalw [Pair_def] "<a,b> = <c,d> <-> a=c & b=d"; | 
| 22 | by (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1); | |
| 23 | by (Blast_tac 1) ; | |
| 24 | qed "Pair_iff"; | |
| 0 | 25 | |
| 2469 | 26 | Addsimps [Pair_iff]; | 
| 27 | ||
| 28 | bind_thm ("Pair_inject", Pair_iff RS iffD1 RS conjE);
 | |
| 0 | 29 | |
| 2469 | 30 | AddSEs [Pair_inject]; | 
| 31 | ||
| 32 | bind_thm ("Pair_inject1", Pair_iff RS iffD1 RS conjunct1);
 | |
| 33 | bind_thm ("Pair_inject2", Pair_iff RS iffD1 RS conjunct2);
 | |
| 0 | 34 | |
| 9211 | 35 | Goalw [Pair_def] "<a,b> ~= 0"; | 
| 36 | by (blast_tac (claset() addEs [equalityE]) 1) ; | |
| 37 | qed "Pair_not_0"; | |
| 0 | 38 | |
| 2469 | 39 | bind_thm ("Pair_neq_0", Pair_not_0 RS notE);
 | 
| 0 | 40 | |
| 2469 | 41 | AddSEs [Pair_neq_0, sym RS Pair_neq_0]; | 
| 42 | ||
| 9211 | 43 | Goalw [Pair_def] "<a,b>=a ==> P"; | 
| 44 | by (rtac (consI1 RS mem_asym RS FalseE) 1); | |
| 45 | by (etac subst 1); | |
| 46 | by (rtac consI1 1) ; | |
| 47 | qed "Pair_neq_fst"; | |
| 0 | 48 | |
| 9211 | 49 | Goalw [Pair_def] "<a,b>=b ==> P"; | 
| 50 | by (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1); | |
| 51 | by (etac subst 1); | |
| 52 | by (rtac (consI1 RS consI2) 1) ; | |
| 53 | qed "Pair_neq_snd"; | |
| 0 | 54 | |
| 55 | ||
| 56 | (*** Sigma: Disjoint union of a family of sets | |
| 57 | Generalizes Cartesian product ***) | |
| 58 | ||
| 9211 | 59 | Goalw [Sigma_def] "<a,b>: Sigma(A,B) <-> a:A & b:B(a)"; | 
| 60 | by (Blast_tac 1) ; | |
| 61 | qed "Sigma_iff"; | |
| 2469 | 62 | |
| 63 | Addsimps [Sigma_iff]; | |
| 64 | ||
| 9180 | 65 | Goal "[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"; | 
| 66 | by (Asm_simp_tac 1); | |
| 67 | qed "SigmaI"; | |
| 68 | ||
| 6153 | 69 | AddTCs [SigmaI]; | 
| 2469 | 70 | |
| 71 | bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1);
 | |
| 72 | bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2);
 | |
| 0 | 73 | |
| 74 | (*The general elimination rule*) | |
| 9211 | 75 | val major::prems= Goalw [Sigma_def] | 
| 0 | 76 | "[| c: Sigma(A,B); \ | 
| 77 | \ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \ | |
| 9211 | 78 | \ |] ==> P"; | 
| 79 | by (cut_facts_tac [major] 1); | |
| 80 | by (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ; | |
| 81 | qed "SigmaE"; | |
| 0 | 82 | |
| 9180 | 83 | val [major,minor]= Goal | 
| 0 | 84 | "[| <a,b> : Sigma(A,B); \ | 
| 85 | \ [| a:A; b:B(a) |] ==> P \ | |
| 9180 | 86 | \ |] ==> P"; | 
| 87 | by (rtac minor 1); | |
| 88 | by (rtac (major RS SigmaD1) 1); | |
| 89 | by (rtac (major RS SigmaD2) 1) ; | |
| 90 | qed "SigmaE2"; | |
| 0 | 91 | |
| 9211 | 92 | val prems= Goalw [Sigma_def] | 
| 0 | 93 | "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ | 
| 9211 | 94 | \ Sigma(A,B) = Sigma(A',B')"; | 
| 95 | by (simp_tac (simpset() addsimps prems) 1) ; | |
| 96 | qed "Sigma_cong"; | |
| 2469 | 97 | |
| 0 | 98 | |
| 2469 | 99 | (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause | 
| 100 | flex-flex pairs and the "Check your prover" error. Most | |
| 101 | Sigmas and Pis are abbreviated as * or -> *) | |
| 0 | 102 | |
| 2469 | 103 | AddSIs [SigmaI]; | 
| 104 | AddSEs [SigmaE2, SigmaE]; | |
| 0 | 105 | |
| 9180 | 106 | Goal "Sigma(0,B) = 0"; | 
| 107 | by (Blast_tac 1) ; | |
| 108 | qed "Sigma_empty1"; | |
| 2469 | 109 | |
| 9180 | 110 | Goal "A*0 = 0"; | 
| 111 | by (Blast_tac 1) ; | |
| 112 | qed "Sigma_empty2"; | |
| 2469 | 113 | |
| 114 | Addsimps [Sigma_empty1, Sigma_empty2]; | |
| 1105 | 115 | |
| 5264 | 116 | Goal "A*B=0 <-> A=0 | B=0"; | 
| 117 | by (Blast_tac 1); | |
| 118 | qed "Sigma_empty_iff"; | |
| 119 | ||
| 1105 | 120 | |
| 121 | (*** Projections: fst, snd ***) | |
| 122 | ||
| 9211 | 123 | Goalw [fst_def] "fst(<a,b>) = a"; | 
| 124 | by (Blast_tac 1) ; | |
| 125 | qed "fst_conv"; | |
| 1105 | 126 | |
| 9211 | 127 | Goalw [snd_def] "snd(<a,b>) = b"; | 
| 128 | by (Blast_tac 1) ; | |
| 129 | qed "snd_conv"; | |
| 1105 | 130 | |
| 2469 | 131 | Addsimps [fst_conv,snd_conv]; | 
| 1105 | 132 | |
| 9180 | 133 | Goal "p:Sigma(A,B) ==> fst(p) : A"; | 
| 134 | by (Auto_tac) ; | |
| 135 | qed "fst_type"; | |
| 1105 | 136 | |
| 9180 | 137 | Goal "p:Sigma(A,B) ==> snd(p) : B(fst(p))"; | 
| 138 | by (Auto_tac) ; | |
| 139 | qed "snd_type"; | |
| 1105 | 140 | |
| 9180 | 141 | Goal "a: Sigma(A,B) ==> <fst(a),snd(a)> = a"; | 
| 142 | by (Auto_tac) ; | |
| 143 | qed "Pair_fst_snd_eq"; | |
| 1105 | 144 | |
| 0 | 145 | |
| 146 | (*** Eliminator - split ***) | |
| 147 | ||
| 1105 | 148 | (*A META-equality, so that it applies to higher types as well...*) | 
| 6112 | 149 | Goalw [split_def] "split(%x y. c(x,y), <a,b>) == c(a,b)"; | 
| 150 | by (Simp_tac 1); | |
| 151 | qed "split"; | |
| 2469 | 152 | Addsimps [split]; | 
| 153 | ||
| 9180 | 154 | val major::prems= Goal | 
| 0 | 155 | "[| p:Sigma(A,B); \ | 
| 156 | \ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \ | |
| 9180 | 157 | \ |] ==> split(%x y. c(x,y), p) : C(p)"; | 
| 158 | by (rtac (major RS SigmaE) 1); | |
| 159 | by (asm_simp_tac (simpset() addsimps prems) 1); | |
| 160 | qed "split_type"; | |
| 6153 | 161 | AddTCs [split_type]; | 
| 0 | 162 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5264diff
changeset | 163 | Goalw [split_def] | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5264diff
changeset | 164 | "u: A*B ==> \ | 
| 435 | 165 | \ R(split(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 | 166 | by Auto_tac; | 
| 760 | 167 | qed "expand_split"; | 
| 435 | 168 | |
| 169 | ||
| 0 | 170 | (*** split for predicates: result type o ***) | 
| 171 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5264diff
changeset | 172 | Goalw [split_def] "R(a,b) ==> split(R, <a,b>)"; | 
| 2469 | 173 | by (Asm_simp_tac 1); | 
| 1105 | 174 | qed "splitI"; | 
| 0 | 175 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5264diff
changeset | 176 | val major::sigma::prems = Goalw [split_def] | 
| 1461 | 177 | "[| split(R,z); z:Sigma(A,B); \ | 
| 178 | \ !!x y. [| z = <x,y>; R(x,y) |] ==> P \ | |
| 1105 | 179 | \ |] ==> P"; | 
| 180 | by (rtac (sigma RS SigmaE) 1); | |
| 0 | 181 | by (cut_facts_tac [major] 1); | 
| 2469 | 182 | by (REPEAT (ares_tac prems 1)); | 
| 183 | by (Asm_full_simp_tac 1); | |
| 1105 | 184 | qed "splitE"; | 
| 0 | 185 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5264diff
changeset | 186 | Goalw [split_def] "split(R,<a,b>) ==> R(a,b)"; | 
| 2469 | 187 | by (Full_simp_tac 1); | 
| 1105 | 188 | qed "splitD"; | 
| 0 | 189 | |
| 533 
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
 lcp parents: 
437diff
changeset | 190 | |
| 
7357160bc56a
ZF/pair.ML: moved some definitions here from simpdata.ML
 lcp parents: 
437diff
changeset | 191 |