src/HOL/Integ/Relation.thy
author clasohm
Fri, 03 Mar 1995 12:04:45 +0100
changeset 925 15539deb6863
child 972 e61b058d58d2
permissions -rw-r--r--
new version of HOL/Integ with curried function application
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	Relation.thy
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Riccardo Mattolini, Dip. Sistemi e Informatica
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     4
        and	Lawrence C Paulson, Cambridge University Computer Laboratory
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     7
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     8
Functions represented as relations in Higher-Order Set Theory 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     9
*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    10
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    11
Relation = Trancl +
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    12
consts
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    13
    converse    ::      "('a*'a) set => ('a*'a) set"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    14
    "^^"        ::      "[('a*'a) set,'a set] => 'a set" (infixl 90)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    15
    Domain      ::      "('a*'a) set => 'a set"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    16
    Range       ::      "('a*'a) set => 'a set"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    17
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    18
defs
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    19
    converse_def  "converse(r) == {z. (? w:r. ? x y. w=<x,y> & z=<y,x>)}"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    20
    Domain_def    "Domain(r) == {z. ! x. (z=x --> (? y. <x,y>:r))}"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    21
    Range_def     "Range(r) == Domain(converse(r))"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    22
    Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. <x,y>:r)}"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    23
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    24
end