diff -r 38928c4a8eb2 -r 4896b4e4077b src/ZF/Integ/EquivClass.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Integ/EquivClass.thy Tue Sep 22 13:49:22 1998 +0200 @@ -0,0 +1,23 @@ +(* Title: ZF/EquivClass.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Equivalence relations in Zermelo-Fraenkel Set Theory +*) + +EquivClass = Rel + Perm + +consts + "'/" :: [i,i]=>i (infixl 90) (*set of equiv classes*) + congruent :: [i,i=>i]=>o + congruent2 :: [i,[i,i]=>i]=>o + +defs + quotient_def "A/r == {r``{x} . x:A}" + congruent_def "congruent(r,b) == ALL y z. :r --> b(y)=b(z)" + + congruent2_def + "congruent2(r,b) == ALL y1 z1 y2 z2. + :r --> :r --> b(y1,y2) = b(z1,z2)" + +end