Integ/Relation.thy
author clasohm
Tue, 24 Oct 1995 14:59:17 +0100
changeset 251 f04b33ce250f
parent 217 b6c0407f203e
permissions -rw-r--r--
added calls of init_html and make_chart
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
217
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     1
(*  Title: 	Relation.thy
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     2
    ID:         $Id$
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     3
    Author: 	Riccardo Mattolini, Dip. Sistemi e Informatica
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     4
        and	Lawrence C Paulson, Cambridge University Computer Laboratory
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     7
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     8
Functions represented as relations in Higher-Order Set Theory 
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     9
*)
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    10
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    11
Relation = Trancl +
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    12
consts
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    13
    converse    ::      "('a*'a) set => ('a*'a) set"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    14
    "^^"        ::      "[('a*'a) set,'a set] => 'a set" (infixl 90)
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    15
    Domain      ::      "('a*'a) set => 'a set"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    16
    Range       ::      "('a*'a) set => 'a set"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    17
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    18
defs
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    19
    converse_def  "converse(r) == {z. (? w:r. ? x y. w=<x,y> & z=<y,x>)}"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    20
    Domain_def    "Domain(r) == {z. ! x. (z=x --> (? y. <x,y>:r))}"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    21
    Range_def     "Range(r) == Domain(converse(r))"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    22
    Image_def     "r ^^ s == {y. y:Range(r) &  (? x:s. <x,y>:r)}"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    23
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    24
end