| 15300 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      3 |     Copyright   1996  University of Cambridge
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Equivalence Relations in Higher-Order Set Theory *}
 | 
|  |      7 | 
 | 
|  |      8 | theory Equiv_Relations
 | 
|  |      9 | imports Relation Finite_Set
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | subsection {* Equivalence relations *}
 | 
|  |     13 | 
 | 
|  |     14 | locale equiv =
 | 
|  |     15 |   fixes A and r
 | 
|  |     16 |   assumes refl: "refl A r"
 | 
|  |     17 |     and sym: "sym r"
 | 
|  |     18 |     and trans: "trans r"
 | 
|  |     19 | 
 | 
|  |     20 | text {*
 | 
|  |     21 |   Suppes, Theorem 70: @{text r} is an equiv relation iff @{text "r\<inverse> O
 | 
|  |     22 |   r = r"}.
 | 
|  |     23 | 
 | 
|  |     24 |   First half: @{text "equiv A r ==> r\<inverse> O r = r"}.
 | 
|  |     25 | *}
 | 
|  |     26 | 
 | 
|  |     27 | lemma sym_trans_comp_subset:
 | 
|  |     28 |     "sym r ==> trans r ==> r\<inverse> O r \<subseteq> r"
 | 
|  |     29 |   by (unfold trans_def sym_def converse_def) blast
 | 
|  |     30 | 
 | 
|  |     31 | lemma refl_comp_subset: "refl A r ==> r \<subseteq> r\<inverse> O r"
 | 
|  |     32 |   by (unfold refl_def) blast
 | 
|  |     33 | 
 | 
|  |     34 | lemma equiv_comp_eq: "equiv A r ==> r\<inverse> O r = r"
 | 
|  |     35 |   apply (unfold equiv_def)
 | 
|  |     36 |   apply clarify
 | 
|  |     37 |   apply (rule equalityI)
 | 
| 17589 |     38 |    apply (iprover intro: sym_trans_comp_subset refl_comp_subset)+
 | 
| 15300 |     39 |   done
 | 
|  |     40 | 
 | 
|  |     41 | text {* Second half. *}
 | 
|  |     42 | 
 | 
|  |     43 | lemma comp_equivI:
 | 
|  |     44 |     "r\<inverse> O r = r ==> Domain r = A ==> equiv A r"
 | 
|  |     45 |   apply (unfold equiv_def refl_def sym_def trans_def)
 | 
|  |     46 |   apply (erule equalityE)
 | 
|  |     47 |   apply (subgoal_tac "\<forall>x y. (x, y) \<in> r --> (y, x) \<in> r")
 | 
|  |     48 |    apply fast
 | 
|  |     49 |   apply fast
 | 
|  |     50 |   done
 | 
|  |     51 | 
 | 
|  |     52 | 
 | 
|  |     53 | subsection {* Equivalence classes *}
 | 
|  |     54 | 
 | 
|  |     55 | lemma equiv_class_subset:
 | 
|  |     56 |   "equiv A r ==> (a, b) \<in> r ==> r``{a} \<subseteq> r``{b}"
 | 
|  |     57 |   -- {* lemma for the next result *}
 | 
|  |     58 |   by (unfold equiv_def trans_def sym_def) blast
 | 
|  |     59 | 
 | 
|  |     60 | theorem equiv_class_eq: "equiv A r ==> (a, b) \<in> r ==> r``{a} = r``{b}"
 | 
|  |     61 |   apply (assumption | rule equalityI equiv_class_subset)+
 | 
|  |     62 |   apply (unfold equiv_def sym_def)
 | 
|  |     63 |   apply blast
 | 
|  |     64 |   done
 | 
|  |     65 | 
 | 
|  |     66 | lemma equiv_class_self: "equiv A r ==> a \<in> A ==> a \<in> r``{a}"
 | 
|  |     67 |   by (unfold equiv_def refl_def) blast
 | 
|  |     68 | 
 | 
|  |     69 | lemma subset_equiv_class:
 | 
|  |     70 |     "equiv A r ==> r``{b} \<subseteq> r``{a} ==> b \<in> A ==> (a,b) \<in> r"
 | 
|  |     71 |   -- {* lemma for the next result *}
 | 
|  |     72 |   by (unfold equiv_def refl_def) blast
 | 
|  |     73 | 
 | 
|  |     74 | lemma eq_equiv_class:
 | 
|  |     75 |     "r``{a} = r``{b} ==> equiv A r ==> b \<in> A ==> (a, b) \<in> r"
 | 
| 17589 |     76 |   by (iprover intro: equalityD2 subset_equiv_class)
 | 
| 15300 |     77 | 
 | 
|  |     78 | lemma equiv_class_nondisjoint:
 | 
|  |     79 |     "equiv A r ==> x \<in> (r``{a} \<inter> r``{b}) ==> (a, b) \<in> r"
 | 
|  |     80 |   by (unfold equiv_def trans_def sym_def) blast
 | 
|  |     81 | 
 | 
|  |     82 | lemma equiv_type: "equiv A r ==> r \<subseteq> A \<times> A"
 | 
|  |     83 |   by (unfold equiv_def refl_def) blast
 | 
|  |     84 | 
 | 
|  |     85 | theorem equiv_class_eq_iff:
 | 
|  |     86 |   "equiv A r ==> ((x, y) \<in> r) = (r``{x} = r``{y} & x \<in> A & y \<in> A)"
 | 
|  |     87 |   by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type)
 | 
|  |     88 | 
 | 
|  |     89 | theorem eq_equiv_class_iff:
 | 
|  |     90 |   "equiv A r ==> x \<in> A ==> y \<in> A ==> (r``{x} = r``{y}) = ((x, y) \<in> r)"
 | 
|  |     91 |   by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type)
 | 
|  |     92 | 
 | 
|  |     93 | 
 | 
|  |     94 | subsection {* Quotients *}
 | 
|  |     95 | 
 | 
|  |     96 | constdefs
 | 
|  |     97 |   quotient :: "['a set, ('a*'a) set] => 'a set set"  (infixl "'/'/" 90)
 | 
|  |     98 |   "A//r == \<Union>x \<in> A. {r``{x}}"  -- {* set of equiv classes *}
 | 
|  |     99 | 
 | 
|  |    100 | lemma quotientI: "x \<in> A ==> r``{x} \<in> A//r"
 | 
|  |    101 |   by (unfold quotient_def) blast
 | 
|  |    102 | 
 | 
|  |    103 | lemma quotientE:
 | 
|  |    104 |   "X \<in> A//r ==> (!!x. X = r``{x} ==> x \<in> A ==> P) ==> P"
 | 
|  |    105 |   by (unfold quotient_def) blast
 | 
|  |    106 | 
 | 
|  |    107 | lemma Union_quotient: "equiv A r ==> Union (A//r) = A"
 | 
|  |    108 |   by (unfold equiv_def refl_def quotient_def) blast
 | 
|  |    109 | 
 | 
|  |    110 | lemma quotient_disj:
 | 
|  |    111 |   "equiv A r ==> X \<in> A//r ==> Y \<in> A//r ==> X = Y | (X \<inter> Y = {})"
 | 
|  |    112 |   apply (unfold quotient_def)
 | 
|  |    113 |   apply clarify
 | 
|  |    114 |   apply (rule equiv_class_eq)
 | 
|  |    115 |    apply assumption
 | 
|  |    116 |   apply (unfold equiv_def trans_def sym_def)
 | 
|  |    117 |   apply blast
 | 
|  |    118 |   done
 | 
|  |    119 | 
 | 
|  |    120 | lemma quotient_eqI:
 | 
|  |    121 |   "[|equiv A r; X \<in> A//r; Y \<in> A//r; x \<in> X; y \<in> Y; (x,y) \<in> r|] ==> X = Y" 
 | 
|  |    122 |   apply (clarify elim!: quotientE)
 | 
|  |    123 |   apply (rule equiv_class_eq, assumption)
 | 
|  |    124 |   apply (unfold equiv_def sym_def trans_def, blast)
 | 
|  |    125 |   done
 | 
|  |    126 | 
 | 
|  |    127 | lemma quotient_eq_iff:
 | 
|  |    128 |   "[|equiv A r; X \<in> A//r; Y \<in> A//r; x \<in> X; y \<in> Y|] ==> (X = Y) = ((x,y) \<in> r)" 
 | 
|  |    129 |   apply (rule iffI)  
 | 
|  |    130 |    prefer 2 apply (blast del: equalityI intro: quotient_eqI) 
 | 
|  |    131 |   apply (clarify elim!: quotientE)
 | 
|  |    132 |   apply (unfold equiv_def sym_def trans_def, blast)
 | 
|  |    133 |   done
 | 
|  |    134 | 
 | 
|  |    135 | 
 | 
|  |    136 | lemma quotient_empty [simp]: "{}//r = {}"
 | 
|  |    137 | by(simp add: quotient_def)
 | 
|  |    138 | 
 | 
|  |    139 | lemma quotient_is_empty [iff]: "(A//r = {}) = (A = {})"
 | 
|  |    140 | by(simp add: quotient_def)
 | 
|  |    141 | 
 | 
|  |    142 | lemma quotient_is_empty2 [iff]: "({} = A//r) = (A = {})"
 | 
|  |    143 | by(simp add: quotient_def)
 | 
|  |    144 | 
 | 
|  |    145 | 
 | 
| 15302 |    146 | lemma singleton_quotient: "{x}//r = {r `` {x}}"
 | 
|  |    147 | by(simp add:quotient_def)
 | 
|  |    148 | 
 | 
|  |    149 | lemma quotient_diff1:
 | 
|  |    150 |   "\<lbrakk> inj_on (%a. {a}//r) A; a \<in> A \<rbrakk> \<Longrightarrow> (A - {a})//r = A//r - {a}//r"
 | 
|  |    151 | apply(simp add:quotient_def inj_on_def)
 | 
|  |    152 | apply blast
 | 
|  |    153 | done
 | 
|  |    154 | 
 | 
| 15300 |    155 | subsection {* Defining unary operations upon equivalence classes *}
 | 
|  |    156 | 
 | 
|  |    157 | text{*A congruence-preserving function*}
 | 
|  |    158 | locale congruent =
 | 
|  |    159 |   fixes r and f
 | 
|  |    160 |   assumes congruent: "(y,z) \<in> r ==> f y = f z"
 | 
|  |    161 | 
 | 
|  |    162 | syntax
 | 
|  |    163 |   RESPECTS ::"['a => 'b, ('a * 'a) set] => bool"  (infixr "respects" 80)
 | 
|  |    164 | 
 | 
|  |    165 | translations
 | 
|  |    166 |   "f respects r" == "congruent r f"
 | 
|  |    167 | 
 | 
|  |    168 | 
 | 
|  |    169 | lemma UN_constant_eq: "a \<in> A ==> \<forall>y \<in> A. f y = c ==> (\<Union>y \<in> A. f(y))=c"
 | 
|  |    170 |   -- {* lemma required to prove @{text UN_equiv_class} *}
 | 
|  |    171 |   by auto
 | 
|  |    172 | 
 | 
|  |    173 | lemma UN_equiv_class:
 | 
|  |    174 |   "equiv A r ==> f respects r ==> a \<in> A
 | 
|  |    175 |     ==> (\<Union>x \<in> r``{a}. f x) = f a"
 | 
|  |    176 |   -- {* Conversion rule *}
 | 
|  |    177 |   apply (rule equiv_class_self [THEN UN_constant_eq], assumption+)
 | 
|  |    178 |   apply (unfold equiv_def congruent_def sym_def)
 | 
|  |    179 |   apply (blast del: equalityI)
 | 
|  |    180 |   done
 | 
|  |    181 | 
 | 
|  |    182 | lemma UN_equiv_class_type:
 | 
|  |    183 |   "equiv A r ==> f respects r ==> X \<in> A//r ==>
 | 
|  |    184 |     (!!x. x \<in> A ==> f x \<in> B) ==> (\<Union>x \<in> X. f x) \<in> B"
 | 
|  |    185 |   apply (unfold quotient_def)
 | 
|  |    186 |   apply clarify
 | 
|  |    187 |   apply (subst UN_equiv_class)
 | 
|  |    188 |      apply auto
 | 
|  |    189 |   done
 | 
|  |    190 | 
 | 
|  |    191 | text {*
 | 
|  |    192 |   Sufficient conditions for injectiveness.  Could weaken premises!
 | 
|  |    193 |   major premise could be an inclusion; bcong could be @{text "!!y. y \<in>
 | 
|  |    194 |   A ==> f y \<in> B"}.
 | 
|  |    195 | *}
 | 
|  |    196 | 
 | 
|  |    197 | lemma UN_equiv_class_inject:
 | 
|  |    198 |   "equiv A r ==> f respects r ==>
 | 
|  |    199 |     (\<Union>x \<in> X. f x) = (\<Union>y \<in> Y. f y) ==> X \<in> A//r ==> Y \<in> A//r
 | 
|  |    200 |     ==> (!!x y. x \<in> A ==> y \<in> A ==> f x = f y ==> (x, y) \<in> r)
 | 
|  |    201 |     ==> X = Y"
 | 
|  |    202 |   apply (unfold quotient_def)
 | 
|  |    203 |   apply clarify
 | 
|  |    204 |   apply (rule equiv_class_eq)
 | 
|  |    205 |    apply assumption
 | 
|  |    206 |   apply (subgoal_tac "f x = f xa")
 | 
|  |    207 |    apply blast
 | 
|  |    208 |   apply (erule box_equals)
 | 
|  |    209 |    apply (assumption | rule UN_equiv_class)+
 | 
|  |    210 |   done
 | 
|  |    211 | 
 | 
|  |    212 | 
 | 
|  |    213 | subsection {* Defining binary operations upon equivalence classes *}
 | 
|  |    214 | 
 | 
|  |    215 | text{*A congruence-preserving function of two arguments*}
 | 
|  |    216 | locale congruent2 =
 | 
|  |    217 |   fixes r1 and r2 and f
 | 
|  |    218 |   assumes congruent2:
 | 
|  |    219 |     "(y1,z1) \<in> r1 ==> (y2,z2) \<in> r2 ==> f y1 y2 = f z1 z2"
 | 
|  |    220 | 
 | 
|  |    221 | text{*Abbreviation for the common case where the relations are identical*}
 | 
|  |    222 | syntax
 | 
|  |    223 |   RESPECTS2 ::"['a => 'b, ('a * 'a) set] => bool"  (infixr "respects2 " 80)
 | 
|  |    224 | 
 | 
|  |    225 | translations
 | 
|  |    226 |   "f respects2 r" => "congruent2 r r f"
 | 
|  |    227 | 
 | 
|  |    228 | lemma congruent2_implies_congruent:
 | 
|  |    229 |     "equiv A r1 ==> congruent2 r1 r2 f ==> a \<in> A ==> congruent r2 (f a)"
 | 
|  |    230 |   by (unfold congruent_def congruent2_def equiv_def refl_def) blast
 | 
|  |    231 | 
 | 
|  |    232 | lemma congruent2_implies_congruent_UN:
 | 
|  |    233 |   "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a \<in> A2 ==>
 | 
|  |    234 |     congruent r1 (\<lambda>x1. \<Union>x2 \<in> r2``{a}. f x1 x2)"
 | 
|  |    235 |   apply (unfold congruent_def)
 | 
|  |    236 |   apply clarify
 | 
|  |    237 |   apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+)
 | 
|  |    238 |   apply (simp add: UN_equiv_class congruent2_implies_congruent)
 | 
|  |    239 |   apply (unfold congruent2_def equiv_def refl_def)
 | 
|  |    240 |   apply (blast del: equalityI)
 | 
|  |    241 |   done
 | 
|  |    242 | 
 | 
|  |    243 | lemma UN_equiv_class2:
 | 
|  |    244 |   "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a1 \<in> A1 ==> a2 \<in> A2
 | 
|  |    245 |     ==> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. f x1 x2) = f a1 a2"
 | 
|  |    246 |   by (simp add: UN_equiv_class congruent2_implies_congruent
 | 
|  |    247 |     congruent2_implies_congruent_UN)
 | 
|  |    248 | 
 | 
|  |    249 | lemma UN_equiv_class_type2:
 | 
|  |    250 |   "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f
 | 
|  |    251 |     ==> X1 \<in> A1//r1 ==> X2 \<in> A2//r2
 | 
|  |    252 |     ==> (!!x1 x2. x1 \<in> A1 ==> x2 \<in> A2 ==> f x1 x2 \<in> B)
 | 
|  |    253 |     ==> (\<Union>x1 \<in> X1. \<Union>x2 \<in> X2. f x1 x2) \<in> B"
 | 
|  |    254 |   apply (unfold quotient_def)
 | 
|  |    255 |   apply clarify
 | 
|  |    256 |   apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
 | 
|  |    257 |     congruent2_implies_congruent quotientI)
 | 
|  |    258 |   done
 | 
|  |    259 | 
 | 
|  |    260 | lemma UN_UN_split_split_eq:
 | 
|  |    261 |   "(\<Union>(x1, x2) \<in> X. \<Union>(y1, y2) \<in> Y. A x1 x2 y1 y2) =
 | 
|  |    262 |     (\<Union>x \<in> X. \<Union>y \<in> Y. (\<lambda>(x1, x2). (\<lambda>(y1, y2). A x1 x2 y1 y2) y) x)"
 | 
|  |    263 |   -- {* Allows a natural expression of binary operators, *}
 | 
|  |    264 |   -- {* without explicit calls to @{text split} *}
 | 
|  |    265 |   by auto
 | 
|  |    266 | 
 | 
|  |    267 | lemma congruent2I:
 | 
|  |    268 |   "equiv A1 r1 ==> equiv A2 r2
 | 
|  |    269 |     ==> (!!y z w. w \<in> A2 ==> (y,z) \<in> r1 ==> f y w = f z w)
 | 
|  |    270 |     ==> (!!y z w. w \<in> A1 ==> (y,z) \<in> r2 ==> f w y = f w z)
 | 
|  |    271 |     ==> congruent2 r1 r2 f"
 | 
|  |    272 |   -- {* Suggested by John Harrison -- the two subproofs may be *}
 | 
|  |    273 |   -- {* \emph{much} simpler than the direct proof. *}
 | 
|  |    274 |   apply (unfold congruent2_def equiv_def refl_def)
 | 
|  |    275 |   apply clarify
 | 
|  |    276 |   apply (blast intro: trans)
 | 
|  |    277 |   done
 | 
|  |    278 | 
 | 
|  |    279 | lemma congruent2_commuteI:
 | 
|  |    280 |   assumes equivA: "equiv A r"
 | 
|  |    281 |     and commute: "!!y z. y \<in> A ==> z \<in> A ==> f y z = f z y"
 | 
|  |    282 |     and congt: "!!y z w. w \<in> A ==> (y,z) \<in> r ==> f w y = f w z"
 | 
|  |    283 |   shows "f respects2 r"
 | 
|  |    284 |   apply (rule congruent2I [OF equivA equivA])
 | 
|  |    285 |    apply (rule commute [THEN trans])
 | 
|  |    286 |      apply (rule_tac [3] commute [THEN trans, symmetric])
 | 
|  |    287 |        apply (rule_tac [5] sym)
 | 
|  |    288 |        apply (assumption | rule congt |
 | 
|  |    289 |          erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+
 | 
|  |    290 |   done
 | 
|  |    291 | 
 | 
|  |    292 | 
 | 
|  |    293 | subsection {* Cardinality results *}
 | 
|  |    294 | 
 | 
|  |    295 | text {*Suggested by Florian Kammüller*}
 | 
|  |    296 | 
 | 
|  |    297 | lemma finite_quotient: "finite A ==> r \<subseteq> A \<times> A ==> finite (A//r)"
 | 
|  |    298 |   -- {* recall @{thm equiv_type} *}
 | 
|  |    299 |   apply (rule finite_subset)
 | 
|  |    300 |    apply (erule_tac [2] finite_Pow_iff [THEN iffD2])
 | 
|  |    301 |   apply (unfold quotient_def)
 | 
|  |    302 |   apply blast
 | 
|  |    303 |   done
 | 
|  |    304 | 
 | 
|  |    305 | lemma finite_equiv_class:
 | 
|  |    306 |   "finite A ==> r \<subseteq> A \<times> A ==> X \<in> A//r ==> finite X"
 | 
|  |    307 |   apply (unfold quotient_def)
 | 
|  |    308 |   apply (rule finite_subset)
 | 
|  |    309 |    prefer 2 apply assumption
 | 
|  |    310 |   apply blast
 | 
|  |    311 |   done
 | 
|  |    312 | 
 | 
|  |    313 | lemma equiv_imp_dvd_card:
 | 
|  |    314 |   "finite A ==> equiv A r ==> \<forall>X \<in> A//r. k dvd card X
 | 
|  |    315 |     ==> k dvd card A"
 | 
|  |    316 |   apply (rule Union_quotient [THEN subst])
 | 
|  |    317 |    apply assumption
 | 
|  |    318 |   apply (rule dvd_partition)
 | 
| 15392 |    319 |      prefer 3 apply (blast dest: quotient_disj)
 | 
|  |    320 |     apply (simp_all add: Union_quotient equiv_type)
 | 
| 15300 |    321 |   done
 | 
|  |    322 | 
 | 
| 15303 |    323 | lemma card_quotient_disjoint:
 | 
|  |    324 |  "\<lbrakk> finite A; inj_on (\<lambda>x. {x} // r) A \<rbrakk> \<Longrightarrow> card(A//r) = card A"
 | 
|  |    325 | apply(simp add:quotient_def)
 | 
|  |    326 | apply(subst card_UN_disjoint)
 | 
|  |    327 |    apply assumption
 | 
|  |    328 |   apply simp
 | 
|  |    329 |  apply(fastsimp simp add:inj_on_def)
 | 
| 15539 |    330 | apply (simp add:setsum_constant)
 | 
| 15303 |    331 | done
 | 
|  |    332 | 
 | 
| 15300 |    333 | ML
 | 
|  |    334 | {*
 | 
|  |    335 | val UN_UN_split_split_eq = thm "UN_UN_split_split_eq";
 | 
|  |    336 | val UN_constant_eq = thm "UN_constant_eq";
 | 
|  |    337 | val UN_equiv_class = thm "UN_equiv_class";
 | 
|  |    338 | val UN_equiv_class2 = thm "UN_equiv_class2";
 | 
|  |    339 | val UN_equiv_class_inject = thm "UN_equiv_class_inject";
 | 
|  |    340 | val UN_equiv_class_type = thm "UN_equiv_class_type";
 | 
|  |    341 | val UN_equiv_class_type2 = thm "UN_equiv_class_type2";
 | 
|  |    342 | val Union_quotient = thm "Union_quotient";
 | 
|  |    343 | val comp_equivI = thm "comp_equivI";
 | 
|  |    344 | val congruent2I = thm "congruent2I";
 | 
|  |    345 | val congruent2_commuteI = thm "congruent2_commuteI";
 | 
|  |    346 | val congruent2_def = thm "congruent2_def";
 | 
|  |    347 | val congruent2_implies_congruent = thm "congruent2_implies_congruent";
 | 
|  |    348 | val congruent2_implies_congruent_UN = thm "congruent2_implies_congruent_UN";
 | 
|  |    349 | val congruent_def = thm "congruent_def";
 | 
|  |    350 | val eq_equiv_class = thm "eq_equiv_class";
 | 
|  |    351 | val eq_equiv_class_iff = thm "eq_equiv_class_iff";
 | 
|  |    352 | val equiv_class_eq = thm "equiv_class_eq";
 | 
|  |    353 | val equiv_class_eq_iff = thm "equiv_class_eq_iff";
 | 
|  |    354 | val equiv_class_nondisjoint = thm "equiv_class_nondisjoint";
 | 
|  |    355 | val equiv_class_self = thm "equiv_class_self";
 | 
|  |    356 | val equiv_comp_eq = thm "equiv_comp_eq";
 | 
|  |    357 | val equiv_def = thm "equiv_def";
 | 
|  |    358 | val equiv_imp_dvd_card = thm "equiv_imp_dvd_card";
 | 
|  |    359 | val equiv_type = thm "equiv_type";
 | 
|  |    360 | val finite_equiv_class = thm "finite_equiv_class";
 | 
|  |    361 | val finite_quotient = thm "finite_quotient";
 | 
|  |    362 | val quotientE = thm "quotientE";
 | 
|  |    363 | val quotientI = thm "quotientI";
 | 
|  |    364 | val quotient_def = thm "quotient_def";
 | 
|  |    365 | val quotient_disj = thm "quotient_disj";
 | 
|  |    366 | val refl_comp_subset = thm "refl_comp_subset";
 | 
|  |    367 | val subset_equiv_class = thm "subset_equiv_class";
 | 
|  |    368 | val sym_trans_comp_subset = thm "sym_trans_comp_subset";
 | 
|  |    369 | *}
 | 
|  |    370 | 
 | 
|  |    371 | end
 |