434
|
1 |
(* Title: ZF/ex/Equiv.thy
|
0
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
434
|
4 |
Copyright 1994 University of Cambridge
|
0
|
5 |
|
|
6 |
Equivalence relations in Zermelo-Fraenkel Set Theory
|
|
7 |
*)
|
|
8 |
|
434
|
9 |
Equiv = Rel + Perm +
|
0
|
10 |
consts
|
|
11 |
"'/" :: "[i,i]=>i" (infixl 90) (*set of equiv classes*)
|
|
12 |
congruent :: "[i,i=>i]=>o"
|
|
13 |
congruent2 :: "[i,[i,i]=>i]=>o"
|
|
14 |
|
|
15 |
rules
|
|
16 |
quotient_def "A/r == {r``{x} . x:A}"
|
|
17 |
congruent_def "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)"
|
|
18 |
|
|
19 |
congruent2_def
|
|
20 |
"congruent2(r,b) == ALL y1 z1 y2 z2. \
|
|
21 |
\ <y1,z1>:r --> <y2,z2>:r --> b(y1,y2) = b(z1,z2)"
|
|
22 |
|
|
23 |
end
|