| author | paulson | 
| Wed, 25 Nov 1998 15:52:45 +0100 | |
| changeset 5968 | 06f9dbfff032 | 
| parent 3373 | b19b1456cc78 | 
| child 6812 | ac4c9707ae53 | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: Equiv.thy | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 2 | ID: $Id$ | 
| 2215 | 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 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: 
2215diff
changeset | 9 | Equiv = Relation + Finite + | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 10 | consts | 
| 1476 | 11 |     refl,equiv  ::      "['a set,('a*'a) set]=>bool"
 | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 12 |     sym         ::      "('a*'a) set=>bool"
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 13 |     "'/"        ::      "['a set,('a*'a) set]=>'a set set"  (infixl 90) 
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 14 | (*set of equiv classes*) | 
| 1476 | 15 |     congruent   ::      "[('a*'a) set,'a=>'b]=>bool"
 | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 16 |     congruent2  ::      "[('a*'a) set,['a,'a]=>'b]=>bool"
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 17 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 18 | defs | 
| 1642 | 19 | refl_def "refl A r == r <= A Times A & (ALL x: A. (x,x) : r)" | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 20 | sym_def "sym(r) == ALL x y. (x,y): r --> (y,x): r" | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 21 | equiv_def "equiv A r == refl A r & sym(r) & trans(r)" | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 22 |     quotient_def  "A/r == UN x:A. {r^^{x}}"  
 | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 23 | congruent_def "congruent r b == ALL y z. (y,z):r --> b(y)=b(z)" | 
| 1151 | 24 | congruent2_def "congruent2 r b == ALL y1 z1 y2 z2. | 
| 25 | (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 | 26 | end |