diff -r 0f57375aafce -r 697dcaaf478f src/ZF/UNITY/UNITYMisc.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/UNITY/UNITYMisc.thy Wed Aug 08 14:33:10 2001 +0200 @@ -0,0 +1,25 @@ +(* Title: HOL/UNITY/UNITYMisc.thy + ID: $Id$ + Author: Sidi O Ehmety, Computer Laboratory + Copyright 2001 University of Cambridge + +Some miscellaneous and add-hoc set theory concepts. +*) + + + +UNITYMisc = Main + +constdefs + less_than :: i=>i + "less_than(A) == measure(A, %x. x)" + + lessThan :: [i, i] => i + "lessThan(u,A) == {x:A. x i + "greaterThan(u,A) == {x:A. ui + "Identity(A) == {p:A*A. EX x. p=}" +end \ No newline at end of file