src/ZF/AC/first.thy
author clasohm
Tue, 24 Oct 1995 13:53:09 +0100
changeset 1291 e173be970d27
parent 1203 a39bec971684
child 1401 0c439768f45c
permissions -rw-r--r--
added generation of HTML files to thy_read.ML; removed old HTML package

(*  Title: 	ZF/AC/first.thy
    ID:         $Id$
    Author: 	Krzysztof Grabczewski

Theory helpful in proofs using first element of a well ordered set
*)

first = Order +

consts

  first                   :: "[i, i, i] => o"

defs

  first_def                "first(u, X, R) 
			    == u:X & (ALL v:X. v~=u --> <u,v> : R)"
end