| author | berghofe | 
| Mon, 11 Jun 2001 19:21:13 +0200 | |
| changeset 11371 | 1d5d181b7e28 | 
| parent 11136 | e34e7f6d9b57 | 
| child 12487 | bbd564190c9b | 
| 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: 
5608 
diff
changeset
 | 
8  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
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  | 
| 11136 | 16  | 
  comp  :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set"  (infixr "O" 60)
 | 
| 7912 | 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: 
5608 
diff
changeset
 | 
18  | 
|
| 11136 | 19  | 
  Image :: "[('a * 'b) set, 'a set] => 'b set"                (infixl "``" 90)
 | 
| 10832 | 20  | 
    "r `` s == {y. ? x:s. (x,y):r}"
 | 
| 7912 | 21  | 
|
| 11136 | 22  | 
  Id    :: "('a * 'a) set"                            (*the identity relation*)
 | 
| 7912 | 23  | 
    "Id == {p. ? x. p = (x,x)}"
 | 
24  | 
||
| 11136 | 25  | 
  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
 | 
26  | 
    "diag(A) == UN x:A. {(x,x)}"
 | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
changeset
 | 
27  | 
|
| 11136 | 28  | 
  Domain :: "('a * 'b) set => 'a set"
 | 
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
changeset
 | 
29  | 
    "Domain(r) == {x. ? y. (x,y):r}"
 | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
changeset
 | 
30  | 
|
| 11136 | 31  | 
  Range  :: "('a * 'b) set => 'b set"
 | 
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
changeset
 | 
32  | 
"Range(r) == Domain(r^-1)"  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
changeset
 | 
33  | 
|
| 11136 | 34  | 
  Field :: "('a * 'a) set => 'a set"
 | 
| 10786 | 35  | 
"Field r == Domain r Un Range r"  | 
36  | 
||
| 11136 | 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: 
5978 
diff
changeset
 | 
39  | 
|
| 11136 | 40  | 
  sym    :: "('a * 'a) set => bool"             (*symmetry predicate*)
 | 
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
changeset
 | 
41  | 
"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
 | 
42  | 
|
| 11136 | 43  | 
  antisym:: "('a * 'a) set => bool"          (*antisymmetry predicate*)
 | 
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
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: 
5978 
diff
changeset
 | 
45  | 
|
| 11136 | 46  | 
  trans  :: "('a * 'a) set => bool"          (*transitivity predicate*)
 | 
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5608 
diff
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: 
5608 
diff
changeset
 | 
48  | 
|
| 11136 | 49  | 
  single_valued :: "('a * 'b) set => bool"
 | 
| 10797 | 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: 
5608 
diff
changeset
 | 
51  | 
|
| 
7014
 
11ee650edcd2
Added some definitions and theorems needed for the
 
berghofe 
parents: 
6806 
diff
changeset
 | 
52  | 
  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
 | 
53  | 
    "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
 | 
54  | 
|
| 11136 | 55  | 
  inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set"
 | 
56  | 
    "inv_image r f == {(x,y). (f(x), f(y)) : r}"
 | 
|
57  | 
||
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
changeset
 | 
58  | 
syntax  | 
| 11136 | 59  | 
  reflexive :: "('a * 'a) set => bool"       (*reflexivity over a type*)
 | 
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
changeset
 | 
60  | 
translations  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
changeset
 | 
61  | 
"reflexive" == "refl UNIV"  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
5978 
diff
changeset
 | 
62  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
63  | 
end  |