| author | nipkow | 
| Mon, 17 Mar 1997 15:37:41 +0100 | |
| changeset 2798 | f84be65745b2 | 
| parent 515 | abcc438e7c27 | 
| permissions | -rw-r--r-- | 
| 434 | 1 | (* Title: ZF/ex/Equiv.ML | 
| 0 | 2 | ID: $Id$ | 
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 434 | 4 | Copyright 1994 University of Cambridge | 
| 0 | 5 | |
| 434 | 6 | For Equiv.thy. Equivalence relations in Zermelo-Fraenkel Set Theory | 
| 0 | 7 | *) | 
| 8 | ||
| 9 | val RSLIST = curry (op MRS); | |
| 10 | ||
| 11 | open Equiv; | |
| 12 | ||
| 13 | (*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***) | |
| 14 | ||
| 15 | (** first half: equiv(A,r) ==> converse(r) O r = r **) | |
| 16 | ||
| 17 | goalw Equiv.thy [trans_def,sym_def] | |
| 18 | "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; | |
| 19 | by (fast_tac (ZF_cs addSEs [converseD,compE]) 1); | |
| 20 | val sym_trans_comp_subset = result(); | |
| 21 | ||
| 22 | goalw Equiv.thy [refl_def] | |
| 434 | 23 | "!!A r. [| refl(A,r); r <= A*A |] ==> r <= converse(r) O r"; | 
| 0 | 24 | by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 1); | 
| 25 | val refl_comp_subset = result(); | |
| 26 | ||
| 27 | goalw Equiv.thy [equiv_def] | |
| 28 | "!!A r. equiv(A,r) ==> converse(r) O r = r"; | |
| 29 | by (rtac equalityI 1); | |
| 30 | by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1 | |
| 31 | ORELSE etac conjE 1)); | |
| 32 | val equiv_comp_eq = result(); | |
| 33 | ||
| 34 | (*second half*) | |
| 35 | goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def] | |
| 36 | "!!A r. [| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"; | |
| 37 | by (etac equalityE 1); | |
| 38 | by (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r" 1); | |
| 39 | by (safe_tac ZF_cs); | |
| 40 | by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 3); | |
| 41 | by (ALLGOALS (fast_tac | |
| 42 | (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE]))); | |
| 43 | by flexflex_tac; | |
| 44 | val comp_equivI = result(); | |
| 45 | ||
| 46 | (** Equivalence classes **) | |
| 47 | ||
| 48 | (*Lemma for the next result*) | |
| 515 | 49 | goalw Equiv.thy [trans_def,sym_def] | 
| 50 |     "!!A r. [| sym(r);  trans(r);  <a,b>: r |] ==> r``{a} <= r``{b}";
 | |
| 0 | 51 | by (fast_tac ZF_cs 1); | 
| 52 | val equiv_class_subset = result(); | |
| 53 | ||
| 515 | 54 | goalw Equiv.thy [equiv_def] | 
| 55 |     "!!A r. [| equiv(A,r);  <a,b>: r |] ==> r``{a} = r``{b}";
 | |
| 56 | by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset])); | |
| 57 | by (rewrite_goals_tac [sym_def]); | |
| 0 | 58 | by (fast_tac ZF_cs 1); | 
| 59 | val equiv_class_eq = result(); | |
| 60 | ||
| 61 | val prems = goalw Equiv.thy [equiv_def,refl_def] | |
| 62 |     "[| equiv(A,r);  a: A |] ==> a: r``{a}";
 | |
| 63 | by (cut_facts_tac prems 1); | |
| 64 | by (fast_tac ZF_cs 1); | |
| 65 | val equiv_class_self = result(); | |
| 66 | ||
| 67 | (*Lemma for the next result*) | |
| 68 | goalw Equiv.thy [equiv_def,refl_def] | |
| 69 |     "!!A r. [| equiv(A,r);  r``{b} <= r``{a};  b: A |] ==> <a,b>: r";
 | |
| 70 | by (fast_tac ZF_cs 1); | |
| 71 | val subset_equiv_class = result(); | |
| 72 | ||
| 73 | val prems = goal Equiv.thy | |
| 74 |     "[| r``{a} = r``{b};  equiv(A,r);  b: A |] ==> <a,b>: r";
 | |
| 75 | by (REPEAT (resolve_tac (prems @ [equalityD2, subset_equiv_class]) 1)); | |
| 76 | val eq_equiv_class = result(); | |
| 77 | ||
| 78 | (*thus r``{a} = r``{b} as well*)
 | |
| 79 | goalw Equiv.thy [equiv_def,trans_def,sym_def] | |
| 80 |     "!!A r. [| equiv(A,r);  x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
 | |
| 81 | by (fast_tac ZF_cs 1); | |
| 82 | val equiv_class_nondisjoint = result(); | |
| 83 | ||
| 434 | 84 | goalw Equiv.thy [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A"; | 
| 85 | by (safe_tac ZF_cs); | |
| 0 | 86 | val equiv_type = result(); | 
| 87 | ||
| 88 | goal Equiv.thy | |
| 89 |     "!!A r. equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
 | |
| 90 | by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] | |
| 91 | addDs [equiv_type]) 1); | |
| 92 | val equiv_class_eq_iff = result(); | |
| 93 | ||
| 94 | goal Equiv.thy | |
| 95 |     "!!A r. [| equiv(A,r);  x: A;  y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
 | |
| 96 | by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] | |
| 97 | addDs [equiv_type]) 1); | |
| 98 | val eq_equiv_class_iff = result(); | |
| 99 | ||
| 100 | (*** Quotients ***) | |
| 101 | ||
| 102 | (** Introduction/elimination rules -- needed? **) | |
| 103 | ||
| 104 | val prems = goalw Equiv.thy [quotient_def] "x:A ==> r``{x}: A/r";
 | |
| 105 | by (rtac RepFunI 1); | |
| 106 | by (resolve_tac prems 1); | |
| 107 | val quotientI = result(); | |
| 108 | ||
| 109 | val major::prems = goalw Equiv.thy [quotient_def] | |
| 110 |     "[| X: A/r;  !!x. [| X = r``{x};  x:A |] ==> P |] 	\
 | |
| 111 | \ ==> P"; | |
| 112 | by (rtac (major RS RepFunE) 1); | |
| 113 | by (eresolve_tac prems 1); | |
| 114 | by (assume_tac 1); | |
| 115 | val quotientE = result(); | |
| 116 | ||
| 117 | goalw Equiv.thy [equiv_def,refl_def,quotient_def] | |
| 118 | "!!A r. equiv(A,r) ==> Union(A/r) = A"; | |
| 119 | by (fast_tac eq_cs 1); | |
| 120 | val Union_quotient = result(); | |
| 121 | ||
| 122 | goalw Equiv.thy [quotient_def] | |
| 123 | "!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)"; | |
| 124 | by (safe_tac (ZF_cs addSIs [equiv_class_eq])); | |
| 125 | by (assume_tac 1); | |
| 126 | by (rewrite_goals_tac [equiv_def,trans_def,sym_def]); | |
| 127 | by (fast_tac ZF_cs 1); | |
| 128 | val quotient_disj = result(); | |
| 129 | ||
| 130 | (**** Defining unary operations upon equivalence classes ****) | |
| 131 | ||
| 132 | (** These proofs really require as local premises | |
| 133 | equiv(A,r); congruent(r,b) | |
| 134 | **) | |
| 135 | ||
| 136 | (*Conversion rule*) | |
| 137 | val prems as [equivA,bcong,_] = goal Equiv.thy | |
| 138 |     "[| equiv(A,r);  congruent(r,b);  a: A |] ==> (UN x:r``{a}. b(x)) = b(a)";
 | |
| 139 | by (cut_facts_tac prems 1); | |
| 434 | 140 | by (rtac ([refl RS UN_cong, UN_constant] MRS trans) 1); | 
| 141 | by (etac equiv_class_self 2); | |
| 142 | by (assume_tac 2); | |
| 0 | 143 | by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]); | 
| 144 | by (fast_tac ZF_cs 1); | |
| 145 | val UN_equiv_class = result(); | |
| 146 | ||
| 147 | (*Resolve th against the "local" premises*) | |
| 148 | val localize = RSLIST [equivA,bcong]; | |
| 149 | ||
| 150 | (*type checking of  UN x:r``{a}. b(x) *)
 | |
| 151 | val _::_::prems = goalw Equiv.thy [quotient_def] | |
| 152 | "[| equiv(A,r); congruent(r,b); X: A/r; \ | |
| 153 | \ !!x. x : A ==> b(x) : B |] \ | |
| 154 | \ ==> (UN x:X. b(x)) : B"; | |
| 155 | by (cut_facts_tac prems 1); | |
| 156 | by (safe_tac ZF_cs); | |
| 157 | by (rtac (localize UN_equiv_class RS ssubst) 1); | |
| 158 | by (REPEAT (ares_tac prems 1)); | |
| 159 | val UN_equiv_class_type = result(); | |
| 160 | ||
| 161 | (*Sufficient conditions for injectiveness. Could weaken premises! | |
| 162 | major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B | |
| 163 | *) | |
| 164 | val _::_::prems = goalw Equiv.thy [quotient_def] | |
| 165 | "[| equiv(A,r); congruent(r,b); \ | |
| 166 | \ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ | |
| 167 | \ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |] \ | |
| 168 | \ ==> X=Y"; | |
| 169 | by (cut_facts_tac prems 1); | |
| 170 | by (safe_tac ZF_cs); | |
| 171 | by (rtac (equivA RS equiv_class_eq) 1); | |
| 172 | by (REPEAT (ares_tac prems 1)); | |
| 173 | by (etac box_equals 1); | |
| 174 | by (REPEAT (ares_tac [localize UN_equiv_class] 1)); | |
| 175 | val UN_equiv_class_inject = result(); | |
| 176 | ||
| 177 | ||
| 178 | (**** Defining binary operations upon equivalence classes ****) | |
| 179 | ||
| 180 | ||
| 181 | goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def] | |
| 182 | "!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))"; | |
| 183 | by (fast_tac ZF_cs 1); | |
| 184 | val congruent2_implies_congruent = result(); | |
| 185 | ||
| 186 | val equivA::prems = goalw Equiv.thy [congruent_def] | |
| 187 | "[| equiv(A,r); congruent2(r,b); a: A |] ==> \ | |
| 188 | \    congruent(r, %x1. UN x2:r``{a}. b(x1,x2))";
 | |
| 189 | by (cut_facts_tac (equivA::prems) 1); | |
| 190 | by (safe_tac ZF_cs); | |
| 191 | by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); | |
| 192 | by (assume_tac 1); | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 193 | by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class, | 
| 0 | 194 | congruent2_implies_congruent]) 1); | 
| 195 | by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); | |
| 196 | by (fast_tac ZF_cs 1); | |
| 197 | val congruent2_implies_congruent_UN = result(); | |
| 198 | ||
| 199 | val prems as equivA::_ = goal Equiv.thy | |
| 200 | "[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \ | |
| 201 | \    ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)";
 | |
| 202 | by (cut_facts_tac prems 1); | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 203 | by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class, | 
| 0 | 204 | congruent2_implies_congruent, | 
| 205 | congruent2_implies_congruent_UN]) 1); | |
| 206 | val UN_equiv_class2 = result(); | |
| 207 | ||
| 208 | (*type checking*) | |
| 209 | val prems = goalw Equiv.thy [quotient_def] | |
| 210 | "[| equiv(A,r); congruent2(r,b); \ | |
| 211 | \ X1: A/r; X2: A/r; \ | |
| 212 | \ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B |] \ | |
| 213 | \ ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"; | |
| 214 | by (cut_facts_tac prems 1); | |
| 215 | by (safe_tac ZF_cs); | |
| 216 | by (REPEAT (ares_tac (prems@[UN_equiv_class_type, | |
| 217 | congruent2_implies_congruent_UN, | |
| 218 | congruent2_implies_congruent, quotientI]) 1)); | |
| 219 | val UN_equiv_class_type2 = result(); | |
| 220 | ||
| 221 | ||
| 222 | (*Suggested by John Harrison -- the two subproofs may be MUCH simpler | |
| 223 | than the direct proof*) | |
| 224 | val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def] | |
| 225 | "[| equiv(A,r); \ | |
| 226 | \ !! y z w. [| w: A; <y,z> : r |] ==> b(y,w) = b(z,w); \ | |
| 227 | \ !! y z w. [| w: A; <y,z> : r |] ==> b(w,y) = b(w,z) \ | |
| 228 | \ |] ==> congruent2(r,b)"; | |
| 229 | by (cut_facts_tac prems 1); | |
| 230 | by (safe_tac ZF_cs); | |
| 231 | by (rtac trans 1); | |
| 232 | by (REPEAT (ares_tac prems 1 | |
| 233 | ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1)); | |
| 234 | val congruent2I = result(); | |
| 235 | ||
| 236 | val [equivA,commute,congt] = goal Equiv.thy | |
| 237 | "[| equiv(A,r); \ | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 238 | \ !! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y); \ | 
| 0 | 239 | \ !! y z w. [| w: A; <y,z>: r |] ==> b(w,y) = b(w,z) \ | 
| 240 | \ |] ==> congruent2(r,b)"; | |
| 241 | by (resolve_tac [equivA RS congruent2I] 1); | |
| 242 | by (rtac (commute RS trans) 1); | |
| 243 | by (rtac (commute RS trans RS sym) 3); | |
| 244 | by (rtac sym 5); | |
| 245 | by (REPEAT (ares_tac [congt] 1 | |
| 246 | ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1)); | |
| 247 | val congruent2_commuteI = result(); | |
| 248 | ||
| 249 | (***OBSOLETE VERSION | |
| 250 | (*Rules congruentI and congruentD would simplify use of rewriting below*) | |
| 251 | val [equivA,ZinA,congt,commute] = goalw Equiv.thy [quotient_def] | |
| 252 | "[| equiv(A,r); Z: A/r; \ | |
| 253 | \ !!w. [| w: A |] ==> congruent(r, %z.b(w,z)); \ | |
| 254 | \ !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) \ | |
| 255 | \ |] ==> congruent(r, %w. UN z: Z. b(w,z))"; | |
| 256 | val congt' = rewrite_rule [congruent_def] congt; | |
| 257 | by (cut_facts_tac [ZinA,congt] 1); | |
| 258 | by (rewtac congruent_def); | |
| 259 | by (safe_tac ZF_cs); | |
| 260 | by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); | |
| 261 | by (assume_tac 1); | |
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 262 | by (asm_simp_tac (ZF_ss addsimps [congt RS (equivA RS UN_equiv_class)]) 1); | 
| 0 | 263 | by (rtac (commute RS trans) 1); | 
| 264 | by (rtac (commute RS trans RS sym) 3); | |
| 265 | by (rtac sym 5); | |
| 266 | by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1)); | |
| 267 | val congruent_commuteI = result(); | |
| 268 | ***) |