src/HOL/Integ/Equiv.thy
author wenzelm
Mon, 29 Nov 1999 15:52:49 +0100
changeset 8039 a901bafe4578
parent 6812 ac4c9707ae53
child 9392 c8e6529cc082
permissions -rw-r--r--
Goal: tuned pris;

(*  Title:      Equiv.thy
    ID:         $Id$
    Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1996  University of Cambridge

Equivalence relations in Higher-Order Set Theory 
*)

Equiv = Relation + Finite + 
consts
    equiv      :: "['a set, ('a*'a) set] => bool"
    quotient   :: "['a set, ('a*'a) set] => 'a set set"  (infixl "'/" 90) 
                         (*set of equiv classes*)
    congruent  :: "[('a*'a) set, 'a=>'b] => bool"
    congruent2 :: "[('a*'a) set, ['a,'a]=>'b] => bool"

defs
    equiv_def     "equiv A r == refl A r & sym(r) & trans(r)"
    quotient_def  "A/r == UN x:A. {r^^{x}}"  
    congruent_def   "congruent r b  == ALL y z. (y,z):r --> b(y)=b(z)"
    congruent2_def  "congruent2 r b == ALL y1 z1 y2 z2. 
           (y1,z1):r --> (y2,z2):r --> b y1 y2 = b z1 z2"
end