author | paulson |
Fri, 05 Nov 1999 12:45:37 +0100 | |
changeset 7999 | 7acf6eb8eec1 |
parent 7912 | 0e42be14f136 |
child 8268 | 722074b93cdd |
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 |
|
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
9 |
constdefs |
7912 | 10 |
converse :: "('a*'b) set => ('b*'a) set" ("(_^-1)" [1000] 999) |
11 |
"r^-1 == {(y,x). (x,y):r}" |
|
12 |
||
13 |
comp :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr "O" 60) |
|
14 |
"r O s == {(x,z). ? y. (x,y):s & (y,z):r}" |
|
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
15 |
|
7912 | 16 |
Image :: "[('a*'b) set,'a set] => 'b set" (infixl "^^" 90) |
17 |
"r ^^ s == {y. ? x:s. (x,y):r}" |
|
18 |
||
19 |
Id :: "('a * 'a)set" (*the identity relation*) |
|
20 |
"Id == {p. ? x. p = (x,x)}" |
|
21 |
||
22 |
diag :: "'a set => ('a * 'a)set" (*diagonal: identity over a set*) |
|
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
23 |
"diag(A) == UN x:A. {(x,x)}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
24 |
|
6806
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
25 |
Domain :: "('a*'b) set => 'a set" |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
26 |
"Domain(r) == {x. ? y. (x,y):r}" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
27 |
|
6806
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
28 |
Range :: "('a*'b) set => 'b set" |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
29 |
"Range(r) == Domain(r^-1)" |
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
30 |
|
6806
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
31 |
refl :: "['a set, ('a*'a) set] => bool" (*reflexivity over a set*) |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
32 |
"refl A r == r <= A Times A & (ALL x: A. (x,x) : r)" |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
33 |
|
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
34 |
sym :: "('a*'a) set=>bool" (*symmetry predicate*) |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
35 |
"sym(r) == ALL x y. (x,y): r --> (y,x): r" |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
36 |
|
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
37 |
antisym:: "('a * 'a)set => bool" (*antisymmetry predicate*) |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
38 |
"antisym(r) == ALL x y. (x,y):r --> (y,x):r --> x=y" |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
39 |
|
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
40 |
trans :: "('a * 'a)set => bool" (*transitivity predicate*) |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
41 |
"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
|
42 |
|
6806
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
43 |
Univalent :: "('a * 'b)set => bool" |
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5608
diff
changeset
|
44 |
"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
|
45 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
6806
diff
changeset
|
46 |
fun_rel_comp :: "['a => 'b, ('b * 'c) set] => ('a => 'c) set" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
6806
diff
changeset
|
47 |
"fun_rel_comp f R == {g. !x. (f x, g x) : R}" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
6806
diff
changeset
|
48 |
|
6806
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
49 |
syntax |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
50 |
reflexive :: "('a * 'a)set => bool" (*reflexivity over a type*) |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
51 |
|
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
52 |
translations |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
53 |
"reflexive" == "refl UNIV" |
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents:
5978
diff
changeset
|
54 |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
55 |
end |