src/ZF/UNITY/UNITYMisc.thy
author paulson
Wed, 08 Aug 2001 14:33:10 +0200
changeset 11479 697dcaaf478f
child 14046 6616e6c53d48
permissions -rw-r--r--
new ZF/UNITY theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/UNITYMisc.thy
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Some miscellaneous and add-hoc set theory concepts.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
UNITYMisc = Main +
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
  less_than :: i=>i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
  "less_than(A) == measure(A, %x. x)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
  lessThan :: [i, i] => i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
  "lessThan(u,A) == {x:A. x<u}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
  greaterThan :: [i,i]=> i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
  "greaterThan(u,A) == {x:A. u<x}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
  (* Identity relation over a domain A *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
  Identity :: i=>i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
  "Identity(A) == {p:A*A. EX x. p=<x,x>}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
end