src/HOL/Relation.thy
author paulson
Wed, 14 Feb 2001 12:16:32 +0100
changeset 11115 285b31e9e026
parent 10832 e33b47e4246d
child 11136 e34e7f6d9b57
permissions -rw-r--r--
a new theorem from Bryan Ford
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10358
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
     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
f3f7bf0079fa Simplification and tidying of definitions
paulson
parents: 1695
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
f3f7bf0079fa Simplification and tidying of definitions
paulson
parents: 1695
diff changeset
     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
33fe2d701ddd *** empty log message ***
nipkow
parents: 8703
diff changeset
     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
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
    10
  converse :: "('a * 'b) set => ('b * 'a) set"    ("(_^-1)" [1000] 999)
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
    11
  "r^-1 == {(y, x). (x, y) : r}"
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
    12
syntax (xsymbols)
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
    13
  converse :: "('a * 'b) set => ('b * 'a) set"    ("(_\\<inverse>)" [1000] 999)
7912
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    14
10358
ef2a753cda2a converse: syntax \<inverse>;
wenzelm
parents: 10212
diff changeset
    15
constdefs
7912
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    16
  comp  :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set"  (infixr "O" 60)
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    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
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10797
diff changeset
    19
  Image :: "[('a*'b) set,'a set] => 'b set"                (infixl "``" 90)
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10797
diff changeset
    20
    "r `` s == {y. ? x:s. (x,y):r}"
7912
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    21
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    22
  Id    :: "('a * 'a)set"                            (*the identity relation*)
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    23
    "Id == {p. ? x. p = (x,x)}"
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    24
0e42be14f136 tidied using modern infix form
paulson
parents: 7014
diff changeset
    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
  
6806
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    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
6806
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    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
10786
04ee73606993 Field of a relation, and some Domain/Range rules
paulson
parents: 10358
diff changeset
    34
  Field :: "('a*'a)set=>'a set"
04ee73606993 Field of a relation, and some Domain/Range rules
paulson
parents: 10358
diff changeset
    35
    "Field r == Domain r Un Range r"
04ee73606993 Field of a relation, and some Domain/Range rules
paulson
parents: 10358
diff changeset
    36
6806
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    37
  refl   :: "['a set, ('a*'a) set] => bool" (*reflexivity over a set*)
8703
816d8f6513be Times -> <*>
nipkow
parents: 8268
diff changeset
    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
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    40
  sym    :: "('a*'a) set=>bool"             (*symmetry predicate*)
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
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    43
  antisym:: "('a * 'a)set => bool"          (*antisymmetry predicate*)
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
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    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
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10786
diff changeset
    49
  single_valued :: "('a * 'b)set => bool"
028d22926a41 ^^ -> ```
nipkow
parents: 10786
diff changeset
    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
6806
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    55
syntax
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    56
  reflexive :: "('a * 'a)set => bool"       (*reflexivity over a type*)
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    57
translations
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    58
  "reflexive" == "refl UNIV"
43c081a0858d new preficates refl, sym [from Integ/Equiv], antisym
paulson
parents: 5978
diff changeset
    59
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    60
end