author | wenzelm |
Fri, 21 May 1999 11:40:15 +0200 | |
changeset 6685 | e33ae2af0d36 |
parent 5978 | fa2c2dd74f8c |
child 6806 | 43c081a0858d |
permissions | -rw-r--r-- |
1475 | 1 |
(* Title: Relation.thy |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
2 |
ID: $Id$ |
1983 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
4 |
Copyright 1996 University of Cambridge |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
5 |
*) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
6 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
7 |
Relation = Prod + |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
8 |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
9 |
consts |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
10 |
O :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60) |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
11 |
converse :: "('a*'b) set => ('b*'a) set" ("(_^-1)" [1000] 999) |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
12 |
"^^" :: "[('a*'b) set,'a set] => 'b set" (infixl 90) |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
13 |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
14 |
defs |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
15 |
comp_def "r O s == {(x,z). ? y. (x,y):s & (y,z):r}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
16 |
converse_def "r^-1 == {(y,x). (x,y):r}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
17 |
Image_def "r ^^ s == {y. ? x:s. (x,y):r}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
18 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
19 |
constdefs |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
20 |
Id :: "('a * 'a)set" (*the identity relation*) |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
21 |
"Id == {p. ? x. p = (x,x)}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
22 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
23 |
diag :: "'a set => ('a * 'a)set" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
24 |
"diag(A) == UN x:A. {(x,x)}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
25 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
26 |
Domain :: "('a*'b) set => 'a set" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
27 |
"Domain(r) == {x. ? y. (x,y):r}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
28 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
29 |
Range :: "('a*'b) set => 'b set" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
30 |
"Range(r) == Domain(r^-1)" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
31 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
32 |
trans :: "('a * 'a)set => bool" (*transitivity predicate*) |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
33 |
"trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
34 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
35 |
Univalent :: "('a * 'b)set => bool" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
36 |
"Univalent r == !x y. (x,y):r --> (!z. (x,z):r --> y=z)" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
37 |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
38 |
end |