1 (* Title: Equiv.ML |
|
2 ID: $Id$ |
|
3 Authors: Riccardo Mattolini, Dip. Sistemi e Informatica |
|
4 Lawrence C Paulson, Cambridge University Computer Laboratory |
|
5 Copyright 1994 Universita' di Firenze |
|
6 Copyright 1993 University of Cambridge |
|
7 |
|
8 Equivalence relations in HOL Set Theory |
|
9 *) |
|
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,converse_def] |
|
18 "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; |
|
19 by (fast_tac (comp_cs addSEs [converseD]) 1); |
|
20 qed "sym_trans_comp_subset"; |
|
21 |
|
22 goalw Equiv.thy [refl_def] |
|
23 "!!A r. refl(A,r) ==> r <= converse(r) O r"; |
|
24 by (fast_tac (rel_cs addIs [compI]) 1); |
|
25 qed "refl_comp_subset"; |
|
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 qed "equiv_comp_eq"; |
|
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 set_cs); |
|
40 by (fast_tac (set_cs addSIs [converseI] addIs [compI]) 3); |
|
41 by (ALLGOALS (fast_tac (rel_cs addIs [compI] addSEs [compE]))); |
|
42 qed "comp_equivI"; |
|
43 |
|
44 (** Equivalence classes **) |
|
45 |
|
46 (*Lemma for the next result*) |
|
47 goalw Equiv.thy [equiv_def,trans_def,sym_def] |
|
48 "!!A r. [| equiv(A,r); <a,b>: r |] ==> r^^{a} <= r^^{b}"; |
|
49 by (safe_tac rel_cs); |
|
50 by (rtac ImageI 1); |
|
51 by (fast_tac rel_cs 2); |
|
52 by (fast_tac rel_cs 1); |
|
53 qed "equiv_class_subset"; |
|
54 |
|
55 goal Equiv.thy "!!A r. [| equiv(A,r); <a,b>: r |] ==> r^^{a} = r^^{b}"; |
|
56 by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1)); |
|
57 by (rewrite_goals_tac [equiv_def,sym_def]); |
|
58 by (fast_tac rel_cs 1); |
|
59 qed "equiv_class_eq"; |
|
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 rel_cs 1); |
|
65 qed "equiv_class_self"; |
|
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 rel_cs 1); |
|
71 qed "subset_equiv_class"; |
|
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 qed "eq_equiv_class"; |
|
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 rel_cs 1); |
|
82 qed "equiv_class_nondisjoint"; |
|
83 |
|
84 val [major] = goalw Equiv.thy [equiv_def,refl_def] |
|
85 "equiv(A,r) ==> r <= Sigma(A,%x.A)"; |
|
86 by (rtac (major RS conjunct1 RS conjunct1) 1); |
|
87 qed "equiv_type"; |
|
88 |
|
89 goal Equiv.thy |
|
90 "!!A r. equiv(A,r) ==> (<x,y>: r) = (r^^{x} = r^^{y} & x:A & y:A)"; |
|
91 by (safe_tac rel_cs); |
|
92 by ((rtac equiv_class_eq 1) THEN (assume_tac 1) THEN (assume_tac 1)); |
|
93 by ((rtac eq_equiv_class 3) THEN |
|
94 (assume_tac 4) THEN (assume_tac 4) THEN (assume_tac 3)); |
|
95 by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN |
|
96 (assume_tac 1) THEN (dtac SigmaD1 1) THEN (assume_tac 1)); |
|
97 by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN |
|
98 (assume_tac 1) THEN (dtac SigmaD2 1) THEN (assume_tac 1)); |
|
99 qed "equiv_class_eq_iff"; |
|
100 |
|
101 goal Equiv.thy |
|
102 "!!A r. [| equiv(A,r); x: A; y: A |] ==> (r^^{x} = r^^{y}) = (<x,y>: r)"; |
|
103 by (safe_tac rel_cs); |
|
104 by ((rtac eq_equiv_class 1) THEN |
|
105 (assume_tac 1) THEN (assume_tac 1) THEN (assume_tac 1)); |
|
106 by ((rtac equiv_class_eq 1) THEN |
|
107 (assume_tac 1) THEN (assume_tac 1)); |
|
108 qed "eq_equiv_class_iff"; |
|
109 |
|
110 (*** Quotients ***) |
|
111 |
|
112 (** Introduction/elimination rules -- needed? **) |
|
113 |
|
114 val prems = goalw Equiv.thy [quotient_def] "x:A ==> r^^{x}: A/r"; |
|
115 by (rtac UN_I 1); |
|
116 by (resolve_tac prems 1); |
|
117 by (rtac singletonI 1); |
|
118 qed "quotientI"; |
|
119 |
|
120 val [major,minor] = goalw Equiv.thy [quotient_def] |
|
121 "[| X:(A/r); !!x. [| X = r^^{x}; x:A |] ==> P |] \ |
|
122 \ ==> P"; |
|
123 by (resolve_tac [major RS UN_E] 1); |
|
124 by (rtac minor 1); |
|
125 by (assume_tac 2); |
|
126 by (fast_tac rel_cs 1); |
|
127 qed "quotientE"; |
|
128 |
|
129 (** Not needed by Theory Integ --> bypassed **) |
|
130 (**goalw Equiv.thy [equiv_def,refl_def,quotient_def] |
|
131 "!!A r. equiv(A,r) ==> Union(A/r) = A"; |
|
132 by (fast_tac eq_cs 1); |
|
133 qed "Union_quotient"; |
|
134 **) |
|
135 |
|
136 (** Not needed by Theory Integ --> bypassed **) |
|
137 (*goalw Equiv.thy [quotient_def] |
|
138 "!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)"; |
|
139 by (safe_tac (ZF_cs addSIs [equiv_class_eq])); |
|
140 by (assume_tac 1); |
|
141 by (rewrite_goals_tac [equiv_def,trans_def,sym_def]); |
|
142 by (fast_tac ZF_cs 1); |
|
143 qed "quotient_disj"; |
|
144 **) |
|
145 |
|
146 (**** Defining unary operations upon equivalence classes ****) |
|
147 |
|
148 (* theorem needed to prove UN_equiv_class *) |
|
149 goal Set.thy "!!A. [| a:A; ! y:A. b(y)=b(a) |] ==> (UN y:A. b(y))=b(a)"; |
|
150 by (fast_tac (eq_cs addSEs [equalityE]) 1); |
|
151 qed "UN_singleton_lemma"; |
|
152 val UN_singleton = ballI RSN (2,UN_singleton_lemma); |
|
153 |
|
154 |
|
155 (** These proofs really require as local premises |
|
156 equiv(A,r); congruent(r,b) |
|
157 **) |
|
158 |
|
159 (*Conversion rule*) |
|
160 val prems as [equivA,bcong,_] = goal Equiv.thy |
|
161 "[| equiv(A,r); congruent(r,b); a: A |] ==> (UN x:r^^{a}. b(x)) = b(a)"; |
|
162 by (cut_facts_tac prems 1); |
|
163 by (rtac UN_singleton 1); |
|
164 by (rtac equiv_class_self 1); |
|
165 by (assume_tac 1); |
|
166 by (assume_tac 1); |
|
167 by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]); |
|
168 by (fast_tac rel_cs 1); |
|
169 qed "UN_equiv_class"; |
|
170 |
|
171 (*Resolve th against the "local" premises*) |
|
172 val localize = RSLIST [equivA,bcong]; |
|
173 |
|
174 (*type checking of UN x:r``{a}. b(x) *) |
|
175 val _::_::prems = goalw Equiv.thy [quotient_def] |
|
176 "[| equiv(A,r); congruent(r,b); X: A/r; \ |
|
177 \ !!x. x : A ==> b(x) : B |] \ |
|
178 \ ==> (UN x:X. b(x)) : B"; |
|
179 by (cut_facts_tac prems 1); |
|
180 by (safe_tac rel_cs); |
|
181 by (rtac (localize UN_equiv_class RS ssubst) 1); |
|
182 by (REPEAT (ares_tac prems 1)); |
|
183 qed "UN_equiv_class_type"; |
|
184 |
|
185 (*Sufficient conditions for injectiveness. Could weaken premises! |
|
186 major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B |
|
187 *) |
|
188 val _::_::prems = goalw Equiv.thy [quotient_def] |
|
189 "[| equiv(A,r); congruent(r,b); \ |
|
190 \ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ |
|
191 \ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |] \ |
|
192 \ ==> X=Y"; |
|
193 by (cut_facts_tac prems 1); |
|
194 by (safe_tac rel_cs); |
|
195 by (rtac (equivA RS equiv_class_eq) 1); |
|
196 by (REPEAT (ares_tac prems 1)); |
|
197 by (etac box_equals 1); |
|
198 by (REPEAT (ares_tac [localize UN_equiv_class] 1)); |
|
199 qed "UN_equiv_class_inject"; |
|
200 |
|
201 |
|
202 (**** Defining binary operations upon equivalence classes ****) |
|
203 |
|
204 |
|
205 goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def] |
|
206 "!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))"; |
|
207 by (fast_tac rel_cs 1); |
|
208 qed "congruent2_implies_congruent"; |
|
209 |
|
210 val equivA::prems = goalw Equiv.thy [congruent_def] |
|
211 "[| equiv(A,r); congruent2(r,b); a: A |] ==> \ |
|
212 \ congruent(r, %x1. UN x2:r^^{a}. b(x1,x2))"; |
|
213 by (cut_facts_tac (equivA::prems) 1); |
|
214 by (safe_tac rel_cs); |
|
215 by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); |
|
216 by (assume_tac 1); |
|
217 by (asm_simp_tac (prod_ss addsimps [equivA RS UN_equiv_class, |
|
218 congruent2_implies_congruent]) 1); |
|
219 by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); |
|
220 by (fast_tac rel_cs 1); |
|
221 qed "congruent2_implies_congruent_UN"; |
|
222 |
|
223 val prems as equivA::_ = goal Equiv.thy |
|
224 "[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \ |
|
225 \ ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b(x1,x2)) = b(a1,a2)"; |
|
226 by (cut_facts_tac prems 1); |
|
227 by (asm_simp_tac (prod_ss addsimps [equivA RS UN_equiv_class, |
|
228 congruent2_implies_congruent, |
|
229 congruent2_implies_congruent_UN]) 1); |
|
230 qed "UN_equiv_class2"; |
|
231 |
|
232 (*type checking*) |
|
233 val prems = goalw Equiv.thy [quotient_def] |
|
234 "[| equiv(A,r); congruent2(r,b); \ |
|
235 \ X1: A/r; X2: A/r; \ |
|
236 \ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B |] \ |
|
237 \ ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"; |
|
238 by (cut_facts_tac prems 1); |
|
239 by (safe_tac rel_cs); |
|
240 by (REPEAT (ares_tac (prems@[UN_equiv_class_type, |
|
241 congruent2_implies_congruent_UN, |
|
242 congruent2_implies_congruent, quotientI]) 1)); |
|
243 qed "UN_equiv_class_type2"; |
|
244 |
|
245 |
|
246 (*Suggested by John Harrison -- the two subproofs may be MUCH simpler |
|
247 than the direct proof*) |
|
248 val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def] |
|
249 "[| equiv(A,r); \ |
|
250 \ !! y z w. [| w: A; <y,z> : r |] ==> b(y,w) = b(z,w); \ |
|
251 \ !! y z w. [| w: A; <y,z> : r |] ==> b(w,y) = b(w,z) \ |
|
252 \ |] ==> congruent2(r,b)"; |
|
253 by (cut_facts_tac prems 1); |
|
254 by (safe_tac rel_cs); |
|
255 by (rtac trans 1); |
|
256 by (REPEAT (ares_tac prems 1 |
|
257 ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1)); |
|
258 qed "congruent2I"; |
|
259 |
|
260 val [equivA,commute,congt] = goal Equiv.thy |
|
261 "[| equiv(A,r); \ |
|
262 \ !! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y); \ |
|
263 \ !! y z w. [| w: A; <y,z>: r |] ==> b(w,y) = b(w,z) \ |
|
264 \ |] ==> congruent2(r,b)"; |
|
265 by (resolve_tac [equivA RS congruent2I] 1); |
|
266 by (rtac (commute RS trans) 1); |
|
267 by (rtac (commute RS trans RS sym) 3); |
|
268 by (rtac sym 5); |
|
269 by (REPEAT (ares_tac [congt] 1 |
|
270 ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1)); |
|
271 qed "congruent2_commuteI"; |
|
272 |
|