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