| author | wenzelm | 
| Fri, 17 Jan 2025 12:19:11 +0100 | |
| changeset 81847 | c163ad6d18a5 | 
| parent 76217 | 8655344f1cf6 | 
| permissions | -rw-r--r-- | 
| 23146 | 1  | 
(* Title: ZF/EquivClass.thy  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 1994 University of Cambridge  | 
|
4  | 
*)  | 
|
5  | 
||
| 60770 | 6  | 
section\<open>Equivalence Relations\<close>  | 
| 23146 | 7  | 
|
8  | 
theory EquivClass imports Trancl Perm begin  | 
|
9  | 
||
| 24893 | 10  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
11  | 
quotient :: "[i,i]\<Rightarrow>i" (infixl \<open>'/'/\<close> 90) (*set of equiv classes*) where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
12  | 
      "A//r \<equiv> {r``{x} . x \<in> A}"
 | 
| 23146 | 13  | 
|
| 24893 | 14  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
15  | 
congruent :: "[i,i\<Rightarrow>i]\<Rightarrow>o" where  | 
| 
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
16  | 
"congruent(r,b) \<equiv> \<forall>y z. \<langle>y,z\<rangle>:r \<longrightarrow> b(y)=b(z)"  | 
| 23146 | 17  | 
|
| 24893 | 18  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
19  | 
congruent2 :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>o" where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
20  | 
"congruent2(r1,r2,b) \<equiv> \<forall>y1 z1 y2 z2.  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
21  | 
\<langle>y1,z1\<rangle>:r1 \<longrightarrow> \<langle>y2,z2\<rangle>:r2 \<longrightarrow> b(y1,y2) = b(z1,z2)"  | 
| 23146 | 22  | 
|
| 24892 | 23  | 
abbreviation  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
24  | 
RESPECTS ::"[i\<Rightarrow>i, i] \<Rightarrow> o" (infixr \<open>respects\<close> 80) where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
25  | 
"f respects r \<equiv> congruent(r,f)"  | 
| 24892 | 26  | 
|
27  | 
abbreviation  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
28  | 
RESPECTS2 ::"[i\<Rightarrow>i\<Rightarrow>i, i] \<Rightarrow> o" (infixr \<open>respects2 \<close> 80) where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
29  | 
"f respects2 r \<equiv> congruent2(r,r,f)"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
61798 
diff
changeset
 | 
30  | 
\<comment> \<open>Abbreviation for the common case where the relations are identical\<close>  | 
| 23146 | 31  | 
|
32  | 
||
| 60770 | 33  | 
subsection\<open>Suppes, Theorem 70:  | 
| 69593 | 34  | 
\<^term>\<open>r\<close> is an equiv relation iff \<^term>\<open>converse(r) O r = r\<close>\<close>  | 
| 23146 | 35  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
36  | 
(** first half: equiv(A,r) \<Longrightarrow> converse(r) O r = r **)  | 
| 23146 | 37  | 
|
38  | 
lemma sym_trans_comp_subset:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
39  | 
"\<lbrakk>sym(r); trans(r)\<rbrakk> \<Longrightarrow> converse(r) O r \<subseteq> r"  | 
| 23146 | 40  | 
by (unfold trans_def sym_def, blast)  | 
41  | 
||
42  | 
lemma refl_comp_subset:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
43  | 
"\<lbrakk>refl(A,r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> r \<subseteq> converse(r) O r"  | 
| 23146 | 44  | 
by (unfold refl_def, blast)  | 
45  | 
||
46  | 
lemma equiv_comp_eq:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
47  | 
"equiv(A,r) \<Longrightarrow> converse(r) O r = r"  | 
| 
76216
 
9fc34f76b4e8
getting rid of apply (unfold ...)
 
paulson <lp15@cam.ac.uk> 
parents: 
76215 
diff
changeset
 | 
48  | 
unfolding equiv_def  | 
| 23146 | 49  | 
apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset)  | 
50  | 
done  | 
|
51  | 
||
52  | 
(*second half*)  | 
|
53  | 
lemma comp_equivI:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
54  | 
"\<lbrakk>converse(r) O r = r; domain(r) = A\<rbrakk> \<Longrightarrow> equiv(A,r)"  | 
| 76217 | 55  | 
unfolding equiv_def refl_def sym_def trans_def  | 
| 23146 | 56  | 
apply (erule equalityE)  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
57  | 
apply (subgoal_tac "\<forall>x y. \<langle>x,y\<rangle> \<in> r \<longrightarrow> \<langle>y,x\<rangle> \<in> r", blast+)  | 
| 23146 | 58  | 
done  | 
59  | 
||
60  | 
(** Equivalence classes **)  | 
|
61  | 
||
62  | 
(*Lemma for the next result*)  | 
|
63  | 
lemma equiv_class_subset:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
64  | 
    "\<lbrakk>sym(r);  trans(r);  \<langle>a,b\<rangle>: r\<rbrakk> \<Longrightarrow> r``{a} \<subseteq> r``{b}"
 | 
| 23146 | 65  | 
by (unfold trans_def sym_def, blast)  | 
66  | 
||
67  | 
lemma equiv_class_eq:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
68  | 
    "\<lbrakk>equiv(A,r);  \<langle>a,b\<rangle>: r\<rbrakk> \<Longrightarrow> r``{a} = r``{b}"
 | 
| 
76216
 
9fc34f76b4e8
getting rid of apply (unfold ...)
 
paulson <lp15@cam.ac.uk> 
parents: 
76215 
diff
changeset
 | 
69  | 
unfolding equiv_def  | 
| 23146 | 70  | 
apply (safe del: subsetI intro!: equalityI equiv_class_subset)  | 
71  | 
apply (unfold sym_def, blast)  | 
|
72  | 
done  | 
|
73  | 
||
74  | 
lemma equiv_class_self:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
75  | 
    "\<lbrakk>equiv(A,r);  a \<in> A\<rbrakk> \<Longrightarrow> a \<in> r``{a}"
 | 
| 23146 | 76  | 
by (unfold equiv_def refl_def, blast)  | 
77  | 
||
78  | 
(*Lemma for the next result*)  | 
|
79  | 
lemma subset_equiv_class:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
80  | 
    "\<lbrakk>equiv(A,r);  r``{b} \<subseteq> r``{a};  b \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 | 
| 23146 | 81  | 
by (unfold equiv_def refl_def, blast)  | 
82  | 
||
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
83  | 
lemma eq_equiv_class: "\<lbrakk>r``{a} = r``{b};  equiv(A,r);  b \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 | 
| 23146 | 84  | 
by (assumption | rule equalityD2 subset_equiv_class)+  | 
85  | 
||
86  | 
(*thus r``{a} = r``{b} as well*)
 | 
|
87  | 
lemma equiv_class_nondisjoint:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
88  | 
    "\<lbrakk>equiv(A,r);  x: (r``{a} \<inter> r``{b})\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 | 
| 23146 | 89  | 
by (unfold equiv_def trans_def sym_def, blast)  | 
90  | 
||
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
91  | 
lemma equiv_type: "equiv(A,r) \<Longrightarrow> r \<subseteq> A*A"  | 
| 23146 | 92  | 
by (unfold equiv_def, blast)  | 
93  | 
||
94  | 
lemma equiv_class_eq_iff:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
95  | 
     "equiv(A,r) \<Longrightarrow> \<langle>x,y\<rangle>: r \<longleftrightarrow> r``{x} = r``{y} \<and> x \<in> A \<and> y \<in> A"
 | 
| 23146 | 96  | 
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)  | 
97  | 
||
98  | 
lemma eq_equiv_class_iff:  | 
|
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
99  | 
     "\<lbrakk>equiv(A,r);  x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> r``{x} = r``{y} \<longleftrightarrow> \<langle>x,y\<rangle>: r"
 | 
| 23146 | 100  | 
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)  | 
101  | 
||
102  | 
(*** Quotients ***)  | 
|
103  | 
||
104  | 
(** Introduction/elimination rules -- needed? **)  | 
|
105  | 
||
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
106  | 
lemma quotientI [TC]: "x \<in> A \<Longrightarrow> r``{x}: A//r"
 | 
| 
76216
 
9fc34f76b4e8
getting rid of apply (unfold ...)
 
paulson <lp15@cam.ac.uk> 
parents: 
76215 
diff
changeset
 | 
107  | 
unfolding quotient_def  | 
| 23146 | 108  | 
apply (erule RepFunI)  | 
109  | 
done  | 
|
110  | 
||
111  | 
lemma quotientE:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
112  | 
    "\<lbrakk>X \<in> A//r;  \<And>x. \<lbrakk>X = r``{x};  x \<in> A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 | 
| 23146 | 113  | 
by (unfold quotient_def, blast)  | 
114  | 
||
115  | 
lemma Union_quotient:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
116  | 
"equiv(A,r) \<Longrightarrow> \<Union>(A//r) = A"  | 
| 23146 | 117  | 
by (unfold equiv_def refl_def quotient_def, blast)  | 
118  | 
||
119  | 
lemma quotient_disj:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
120  | 
"\<lbrakk>equiv(A,r); X \<in> A//r; Y \<in> A//r\<rbrakk> \<Longrightarrow> X=Y | (X \<inter> Y \<subseteq> 0)"  | 
| 
76216
 
9fc34f76b4e8
getting rid of apply (unfold ...)
 
paulson <lp15@cam.ac.uk> 
parents: 
76215 
diff
changeset
 | 
121  | 
unfolding quotient_def  | 
| 23146 | 122  | 
apply (safe intro!: equiv_class_eq, assumption)  | 
123  | 
apply (unfold equiv_def trans_def sym_def, blast)  | 
|
124  | 
done  | 
|
125  | 
||
| 60770 | 126  | 
subsection\<open>Defining Unary Operations upon Equivalence Classes\<close>  | 
| 23146 | 127  | 
|
128  | 
(** Could have a locale with the premises equiv(A,r) and congruent(r,b)  | 
|
129  | 
**)  | 
|
130  | 
||
131  | 
(*Conversion rule*)  | 
|
132  | 
lemma UN_equiv_class:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
133  | 
    "\<lbrakk>equiv(A,r);  b respects r;  a \<in> A\<rbrakk> \<Longrightarrow> (\<Union>x\<in>r``{a}. b(x)) = b(a)"
 | 
| 46820 | 134  | 
apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)")
 | 
| 23146 | 135  | 
apply simp  | 
| 46820 | 136  | 
apply (blast intro: equiv_class_self)  | 
| 23146 | 137  | 
apply (unfold equiv_def sym_def congruent_def, blast)  | 
138  | 
done  | 
|
139  | 
||
| 46820 | 140  | 
(*type checking of  @{term"\<Union>x\<in>r``{a}. b(x)"} *)
 | 
| 23146 | 141  | 
lemma UN_equiv_class_type:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
142  | 
"\<lbrakk>equiv(A,r); b respects r; X \<in> A//r; \<And>x. x \<in> A \<Longrightarrow> b(x) \<in> B\<rbrakk>  | 
| 
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
143  | 
\<Longrightarrow> (\<Union>x\<in>X. b(x)) \<in> B"  | 
| 23146 | 144  | 
apply (unfold quotient_def, safe)  | 
145  | 
apply (simp (no_asm_simp) add: UN_equiv_class)  | 
|
146  | 
done  | 
|
147  | 
||
148  | 
(*Sufficient conditions for injectiveness. Could weaken premises!  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
149  | 
major premise could be an inclusion; bcong could be \<And>y. y \<in> A \<Longrightarrow> b(y):B  | 
| 23146 | 150  | 
*)  | 
151  | 
lemma UN_equiv_class_inject:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
152  | 
"\<lbrakk>equiv(A,r); b respects r;  | 
| 46953 | 153  | 
(\<Union>x\<in>X. b(x))=(\<Union>y\<in>Y. b(y)); X \<in> A//r; Y \<in> A//r;  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
154  | 
\<And>x y. \<lbrakk>x \<in> A; y \<in> A; b(x)=b(y)\<rbrakk> \<Longrightarrow> \<langle>x,y\<rangle>:r\<rbrakk>  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
155  | 
\<Longrightarrow> X=Y"  | 
| 23146 | 156  | 
apply (unfold quotient_def, safe)  | 
157  | 
apply (rule equiv_class_eq, assumption)  | 
|
| 46820 | 158  | 
apply (simp add: UN_equiv_class [of A r b])  | 
| 23146 | 159  | 
done  | 
160  | 
||
161  | 
||
| 60770 | 162  | 
subsection\<open>Defining Binary Operations upon Equivalence Classes\<close>  | 
| 23146 | 163  | 
|
164  | 
lemma congruent2_implies_congruent:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
165  | 
"\<lbrakk>equiv(A,r1); congruent2(r1,r2,b); a \<in> A\<rbrakk> \<Longrightarrow> congruent(r2,b(a))"  | 
| 23146 | 166  | 
by (unfold congruent_def congruent2_def equiv_def refl_def, blast)  | 
167  | 
||
168  | 
lemma congruent2_implies_congruent_UN:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
169  | 
"\<lbrakk>equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \<in> A2\<rbrakk> \<Longrightarrow>  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
170  | 
     congruent(r1, \<lambda>x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
 | 
| 23146 | 171  | 
apply (unfold congruent_def, safe)  | 
172  | 
apply (frule equiv_type [THEN subsetD], assumption)  | 
|
| 46820 | 173  | 
apply clarify  | 
| 23146 | 174  | 
apply (simp add: UN_equiv_class congruent2_implies_congruent)  | 
175  | 
apply (unfold congruent2_def equiv_def refl_def, blast)  | 
|
176  | 
done  | 
|
177  | 
||
178  | 
lemma UN_equiv_class2:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
179  | 
"\<lbrakk>equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2\<rbrakk>  | 
| 
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
180  | 
     \<Longrightarrow> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)"
 | 
| 23146 | 181  | 
by (simp add: UN_equiv_class congruent2_implies_congruent  | 
182  | 
congruent2_implies_congruent_UN)  | 
|
183  | 
||
184  | 
(*type checking*)  | 
|
185  | 
lemma UN_equiv_class_type2:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
186  | 
"\<lbrakk>equiv(A,r); b respects2 r;  | 
| 23146 | 187  | 
X1: A//r; X2: A//r;  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
188  | 
\<And>x1 x2. \<lbrakk>x1: A; x2: A\<rbrakk> \<Longrightarrow> b(x1,x2) \<in> B  | 
| 
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
189  | 
\<rbrakk> \<Longrightarrow> (\<Union>x1\<in>X1. \<Union>x2\<in>X2. b(x1,x2)) \<in> B"  | 
| 23146 | 190  | 
apply (unfold quotient_def, safe)  | 
| 46820 | 191  | 
apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN  | 
| 23146 | 192  | 
congruent2_implies_congruent quotientI)  | 
193  | 
done  | 
|
194  | 
||
195  | 
||
196  | 
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler  | 
|
197  | 
than the direct proof*)  | 
|
198  | 
lemma congruent2I:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
199  | 
"\<lbrakk>equiv(A1,r1); equiv(A2,r2);  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
200  | 
\<And>y z w. \<lbrakk>w \<in> A2; \<langle>y,z\<rangle> \<in> r1\<rbrakk> \<Longrightarrow> b(y,w) = b(z,w);  | 
| 
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
201  | 
\<And>y z w. \<lbrakk>w \<in> A1; \<langle>y,z\<rangle> \<in> r2\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
202  | 
\<rbrakk> \<Longrightarrow> congruent2(r1,r2,b)"  | 
| 23146 | 203  | 
apply (unfold congruent2_def equiv_def refl_def, safe)  | 
| 46820 | 204  | 
apply (blast intro: trans)  | 
| 23146 | 205  | 
done  | 
206  | 
||
207  | 
lemma congruent2_commuteI:  | 
|
208  | 
assumes equivA: "equiv(A,r)"  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
209  | 
and commute: "\<And>y z. \<lbrakk>y \<in> A; z \<in> A\<rbrakk> \<Longrightarrow> b(y,z) = b(z,y)"  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
210  | 
and congt: "\<And>y z w. \<lbrakk>w \<in> A; \<langle>y,z\<rangle>: r\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)"  | 
| 23146 | 211  | 
shows "b respects2 r"  | 
| 46820 | 212  | 
apply (insert equivA [THEN equiv_type, THEN subsetD])  | 
| 23146 | 213  | 
apply (rule congruent2I [OF equivA equivA])  | 
214  | 
apply (rule commute [THEN trans])  | 
|
215  | 
apply (rule_tac [3] commute [THEN trans, symmetric])  | 
|
| 46820 | 216  | 
apply (rule_tac [5] sym)  | 
| 23146 | 217  | 
apply (blast intro: congt)+  | 
218  | 
done  | 
|
219  | 
||
220  | 
(*Obsolete?*)  | 
|
221  | 
lemma congruent_commuteI:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
222  | 
"\<lbrakk>equiv(A,r); Z \<in> A//r;  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
223  | 
\<And>w. \<lbrakk>w \<in> A\<rbrakk> \<Longrightarrow> congruent(r, \<lambda>z. b(w,z));  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
224  | 
\<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> b(y,x) = b(x,y)  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
225  | 
\<rbrakk> \<Longrightarrow> congruent(r, \<lambda>w. \<Union>z\<in>Z. b(w,z))"  | 
| 23146 | 226  | 
apply (simp (no_asm) add: congruent_def)  | 
227  | 
apply (safe elim!: quotientE)  | 
|
228  | 
apply (frule equiv_type [THEN subsetD], assumption)  | 
|
| 46820 | 229  | 
apply (simp add: UN_equiv_class [of A r])  | 
230  | 
apply (simp add: congruent_def)  | 
|
| 23146 | 231  | 
done  | 
232  | 
||
233  | 
end  |