src/HOL/Integ/Equiv.thy
author nipkow
Fri, 05 Jan 2001 18:48:18 +0100
changeset 10797 028d22926a41
parent 9392 c8e6529cc082
child 10834 a7897aebbffc
permissions -rw-r--r--
^^ -> ``` Univalent -> single_valued
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1151
diff changeset
     1
(*  Title:      Equiv.thy
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 1642
diff changeset
     3
    Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 1642
diff changeset
     4
    Copyright   1996  University of Cambridge
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     5
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     6
Equivalence relations in Higher-Order Set Theory 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     7
*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     8
3373
b19b1456cc78 Addition of Finite as parent allows cardinality theorems
paulson
parents: 2215
diff changeset
     9
Equiv = Relation + Finite + 
9392
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    10
constdefs
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    11
  equiv    :: "['a set, ('a*'a) set] => bool"
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    12
    "equiv A r == refl A r & sym(r) & trans(r)"
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    13
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    14
  quotient :: "['a set, ('a*'a) set] => 'a set set"  (infixl "'/'/" 90) 
10797
028d22926a41 ^^ -> ```
nipkow
parents: 9392
diff changeset
    15
    "A//r == UN x:A. {r```{x}}"      (*set of equiv classes*)
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    16
9392
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    17
  congruent  :: "[('a*'a) set, 'a=>'b] => bool"
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    18
    "congruent r b  == ALL y z. (y,z):r --> b(y)=b(z)"
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    19
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    20
  congruent2 :: "[('a*'a) set, ['a,'a]=>'b] => bool"
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    21
    "congruent2 r b == ALL y1 z1 y2 z2.
c8e6529cc082 changed / to // for quotienting
paulson
parents: 6812
diff changeset
    22
                         (y1,z1):r --> (y2,z2):r --> b y1 y2 = b z1 z2"
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    23
end