src/HOL/Relation.thy
author nipkow
Fri, 26 Jan 1996 20:25:39 +0100
changeset 1454 d0266c81a85e
parent 1128 64b30e3cc6d4
child 1475 7f5a4cd08209
permissions -rw-r--r--
Streamlined defs in Relation and added new intro/elim rules to do with pattern matching in sets: {(x,y). ...} and UN (x,y):A. ...
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     1
(*  Title: 	Relation.thy
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     3
    Author: 	Riccardo Mattolini, Dip. Sistemi e Informatica
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     4
        and	Lawrence C Paulson, Cambridge University Computer Laboratory
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     7
*)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     8
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     9
Relation = Prod +
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    10
consts
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    11
    id	        :: "('a * 'a)set"               (*the identity relation*)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    12
    O	        :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    13
    trans       :: "('a * 'a)set => bool" 	(*transitivity predicate*)
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1128
diff changeset
    14
    converse    :: "('a * 'b)set => ('b * 'a)set"
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1128
diff changeset
    15
    "^^"        :: "[('a * 'b) set, 'a set] => 'b set" (infixl 90)
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1128
diff changeset
    16
    Domain      :: "('a * 'b) set => 'a set"
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1128
diff changeset
    17
    Range       :: "('a * 'b) set => 'b set"
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    18
defs
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    19
    id_def	"id == {p. ? x. p = (x,x)}"
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1128
diff changeset
    20
    comp_def	"r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    21
    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
    22
    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
    23
    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
    24
    Range_def     "Range(r) == Domain(converse(r))"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    25
    Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. (x,y):r)}"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    26
end