author | paulson |
Wed, 13 Nov 1996 10:47:08 +0100 | |
changeset 2183 | 8d42a7bccf0b |
parent 1983 | f3f7bf0079fa |
child 3439 | 54785105178c |
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 + |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
8 |
consts |
1475 | 9 |
id :: "('a * 'a)set" (*the identity relation*) |
10 |
O :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60) |
|
11 |
trans :: "('a * 'a)set => bool" (*transitivity predicate*) |
|
1695
0f9b9eda2a2c
Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents:
1475
diff
changeset
|
12 |
converse :: "('a*'b) set => ('b*'a) set" |
0f9b9eda2a2c
Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents:
1475
diff
changeset
|
13 |
"^^" :: "[('a*'b) set,'a set] => 'b set" (infixl 90) |
0f9b9eda2a2c
Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents:
1475
diff
changeset
|
14 |
Domain :: "('a*'b) set => 'a set" |
0f9b9eda2a2c
Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents:
1475
diff
changeset
|
15 |
Range :: "('a*'b) set => 'b set" |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
16 |
defs |
1983 | 17 |
id_def "id == {p. ? x. p = (x,x)}" |
18 |
comp_def "r O s == {(x,z). ? y. (x,y):s & (y,z):r}" |
|
1475 | 19 |
trans_def "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)" |
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1128
diff
changeset
|
20 |
converse_def "converse(r) == {(y,x). (x,y):r}" |
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1128
diff
changeset
|
21 |
Domain_def "Domain(r) == {x. ? y. (x,y):r}" |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
22 |
Range_def "Range(r) == Domain(converse(r))" |
1983 | 23 |
Image_def "r ^^ s == {y. ? x:s. (x,y):r}" |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
24 |
end |