src/ZF/Integ/EquivClass.thy
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 9333 5cacc383157a
child 13239 f599984ec4c2
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     1
(*  Title:      ZF/EquivClass.thy
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     2
    ID:         $Id$
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     5
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     6
Equivalence relations in Zermelo-Fraenkel Set Theory 
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     7
*)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     8
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     9
EquivClass = Rel + Perm + 
9333
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    10
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    11
constdefs
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    12
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    13
  quotient    :: [i,i]=>i    (infixl "'/'/" 90)  (*set of equiv classes*)
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    14
      "A//r == {r``{x} . x:A}"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    15
9333
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    16
  congruent   :: [i,i=>i]=>o
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    17
      "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    18
9333
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    19
  congruent2  :: [i,[i,i]=>i]=>o
5cacc383157a changed the quotient syntax from / to //
paulson
parents: 5528
diff changeset
    20
      "congruent2(r,b) == ALL y1 z1 y2 z2. 
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    21
           <y1,z1>:r --> <y2,z2>:r --> b(y1,y2) = b(z1,z2)"
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    22
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    23
end