src/HOL/Relation.thy
author paulson
Thu, 12 Sep 1996 10:40:05 +0200
changeset 1985 84cf16192e03
parent 1983 f3f7bf0079fa
child 3439 54785105178c
permissions -rw-r--r--
Tidied many proofs, using AddIffs to let equivalences take the place of separate Intr and Elim rules. Also deleted most named clasets.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1454
diff changeset
     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
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
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     7
Relation = Prod +
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     8
consts
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1454
diff changeset
     9
    id          :: "('a * 'a)set"               (*the identity relation*)
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1454
diff changeset
    10
    O           :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1454
diff changeset
    11
    trans       :: "('a * 'a)set => bool"       (*transitivity predicate*)
1695
0f9b9eda2a2c Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents: 1475
diff changeset
    12
    converse    :: "('a*'b) set => ('b*'a) set"
0f9b9eda2a2c Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents: 1475
diff changeset
    13
    "^^"        :: "[('a*'b) set,'a set] => 'b set" (infixl 90)
0f9b9eda2a2c Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents: 1475
diff changeset
    14
    Domain      :: "('a*'b) set => 'a set"
0f9b9eda2a2c Generalized types of some of the operators (thanks to Norbert Voelker)
nipkow
parents: 1475
diff changeset
    15
    Range       :: "('a*'b) set => 'b set"
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    16
defs
1983
f3f7bf0079fa Simplification and tidying of definitions
paulson
parents: 1695
diff changeset
    17
    id_def        "id == {p. ? x. p = (x,x)}"
f3f7bf0079fa Simplification and tidying of definitions
paulson
parents: 1695
diff changeset
    18
    comp_def      "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1454
diff changeset
    19
    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
    20
    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
    21
    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
    22
    Range_def     "Range(r) == Domain(converse(r))"
1983
f3f7bf0079fa Simplification and tidying of definitions
paulson
parents: 1695
diff changeset
    23
    Image_def     "r ^^ s == {y. ? x:s. (x,y):r}"
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    24
end