src/ZF/Rel.thy
author wenzelm
Thu, 08 Nov 2001 23:52:56 +0100
changeset 12106 4a8558dbb6a0
parent 2469 b50b8c0eec01
child 13168 afcbca3498b0
permissions -rw-r--r--
* Isar/Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.) now consider the syntactic context of assumptions, giving a better chance to get type-inference of the arguments right (this is especially important for locales); * system: refrain from any attempt at filtering input streams; no longer support ``8bit'' encoding of old isabelle font, instead proper iso-latin characters may now be used;

(*  Title:      ZF/Rel.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1994  University of Cambridge

Relations in Zermelo-Fraenkel Set Theory 
*)

Rel = domrange +
consts
    refl,irrefl,equiv      :: [i,i]=>o
    sym,asym,antisym,trans :: i=>o
    trans_on               :: [i,i]=>o  ("trans[_]'(_')")

defs
  refl_def     "refl(A,r) == (ALL x: A. <x,x> : r)"

  irrefl_def   "irrefl(A,r) == ALL x: A. <x,x> ~: r"

  sym_def      "sym(r) == ALL x y. <x,y>: r --> <y,x>: r"

  asym_def     "asym(r) == ALL x y. <x,y>:r --> ~ <y,x>:r"

  antisym_def  "antisym(r) == ALL x y.<x,y>:r --> <y,x>:r --> x=y"

  trans_def    "trans(r) == ALL x y z. <x,y>: r --> <y,z>: r --> <x,z>: r"

  trans_on_def "trans[A](r) == ALL x:A. ALL y:A. ALL z:A.       
                          <x,y>: r --> <y,z>: r --> <x,z>: r"

  equiv_def    "equiv(A,r) == r <= A*A & refl(A,r) & sym(r) & trans(r)"

end