src/HOL/Relation.thy
author oheimb
Wed, 12 Aug 1998 16:23:25 +0200
changeset 5305 513925de8962
parent 4746 a5dcd7e4a37d
child 5608 a82a038a3e7a
permissions -rw-r--r--
cleanup for Fun.thy: merged Update.{thy|ML} into Fun.{thy|ML} moved o_def from HOL.thy to Fun.thy added Id_def to Fun.thy moved image_compose from Set.ML to Fun.ML moved o_apply and o_assoc from simpdata.ML to Fun.ML moved fun_upd_same and fun_upd_other (from Map.ML) to Fun.ML added fun_upd_twist to Fun.ML

(*  Title:      Relation.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1996  University of Cambridge
*)

Relation = Prod +
consts
    id          :: "('a * 'a)set"               (*the identity relation*)
    O           :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
    converse    :: "('a*'b) set => ('b*'a) set"     ("(_^-1)" [1000] 999)
    "^^"        :: "[('a*'b) set,'a set] => 'b set" (infixl 90)
    Domain      :: "('a*'b) set => 'a set"
    Range       :: "('a*'b) set => 'b set"
    trans       :: "('a * 'a)set => bool"       (*transitivity predicate*)
    Univalent   :: "('a * 'b)set => bool"
defs
    id_def        "id == {p. ? x. p = (x,x)}"
    comp_def      "r O s == {(x,z). ? y. (x,y):s & (y,z):r}"
    converse_def   "r^-1 == {(y,x). (x,y):r}"
    Domain_def    "Domain(r) == {x. ? y. (x,y):r}"
    Range_def     "Range(r) == Domain(r^-1)"
    Image_def     "r ^^ s == {y. ? x:s. (x,y):r}"
    trans_def     "trans(r) == (!x y z. (x,y):r --> (y,z):r --> (x,z):r)"
    Univalent_def "Univalent r == !x y. (x,y):r --> (!z. (x,z):r --> y=z)"
end