| author | paulson | 
| Wed, 10 Jan 2001 11:00:17 +0100 | |
| changeset 10840 | 28a53b68a8c0 | 
| parent 10832 | e33b47e4246d | 
| child 11136 | e34e7f6d9b57 | 
| permissions | -rw-r--r-- | 
| 10358 | 1 | (* Title: HOL/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 | |
| 10212 | 7 | Relation = Product_Type + | 
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 8 | |
| 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 9 | constdefs | 
| 10358 | 10 |   converse :: "('a * 'b) set => ('b * 'a) set"    ("(_^-1)" [1000] 999)
 | 
| 11 |   "r^-1 == {(y, x). (x, y) : r}"
 | |
| 12 | syntax (xsymbols) | |
| 13 |   converse :: "('a * 'b) set => ('b * 'a) set"    ("(_\\<inverse>)" [1000] 999)
 | |
| 7912 | 14 | |
| 10358 | 15 | constdefs | 
| 7912 | 16 |   comp  :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set"  (infixr "O" 60)
 | 
| 17 |     "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
 | |
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 18 | |
| 10832 | 19 |   Image :: "[('a*'b) set,'a set] => 'b set"                (infixl "``" 90)
 | 
| 20 |     "r `` s == {y. ? x:s. (x,y):r}"
 | |
| 7912 | 21 | |
| 22 |   Id    :: "('a * 'a)set"                            (*the identity relation*)
 | |
| 23 |     "Id == {p. ? x. p = (x,x)}"
 | |
| 24 | ||
| 25 |   diag  :: "'a set => ('a * 'a)set"          (*diagonal: identity over a set*)
 | |
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 26 |     "diag(A) == UN x:A. {(x,x)}"
 | 
| 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 27 | |
| 6806 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 28 |   Domain :: "('a*'b) set => 'a set"
 | 
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 29 |     "Domain(r) == {x. ? y. (x,y):r}"
 | 
| 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 30 | |
| 6806 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 31 |   Range  :: "('a*'b) set => 'b set"
 | 
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 32 | "Range(r) == Domain(r^-1)" | 
| 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 33 | |
| 10786 | 34 |   Field :: "('a*'a)set=>'a set"
 | 
| 35 | "Field r == Domain r Un Range r" | |
| 36 | ||
| 6806 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 37 |   refl   :: "['a set, ('a*'a) set] => bool" (*reflexivity over a set*)
 | 
| 8703 | 38 | "refl A r == r <= A <*> A & (ALL x: A. (x,x) : r)" | 
| 6806 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 39 | |
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 40 |   sym    :: "('a*'a) set=>bool"             (*symmetry predicate*)
 | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 41 | "sym(r) == ALL x y. (x,y): r --> (y,x): r" | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 42 | |
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 43 |   antisym:: "('a * 'a)set => bool"          (*antisymmetry predicate*)
 | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 44 | "antisym(r) == ALL x y. (x,y):r --> (y,x):r --> x=y" | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 45 | |
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 46 |   trans  :: "('a * 'a)set => bool"          (*transitivity predicate*)
 | 
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 47 | "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: 
5608diff
changeset | 48 | |
| 10797 | 49 |   single_valued :: "('a * 'b)set => bool"
 | 
| 50 | "single_valued r == !x y. (x,y):r --> (!z. (x,z):r --> y=z)" | |
| 5978 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 paulson parents: 
5608diff
changeset | 51 | |
| 7014 
11ee650edcd2
Added some definitions and theorems needed for the
 berghofe parents: 
6806diff
changeset | 52 |   fun_rel_comp :: "['a => 'b, ('b * 'c) set] => ('a => 'c) set"
 | 
| 
11ee650edcd2
Added some definitions and theorems needed for the
 berghofe parents: 
6806diff
changeset | 53 |     "fun_rel_comp f R == {g. !x. (f x, g x) : R}"
 | 
| 
11ee650edcd2
Added some definitions and theorems needed for the
 berghofe parents: 
6806diff
changeset | 54 | |
| 6806 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 55 | syntax | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 56 |   reflexive :: "('a * 'a)set => bool"       (*reflexivity over a type*)
 | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 57 | translations | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 58 | "reflexive" == "refl UNIV" | 
| 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 paulson parents: 
5978diff
changeset | 59 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 60 | end |