11479
|
1 |
(* Title: HOL/UNITY/UNITYMisc.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Sidi O Ehmety, Computer Laboratory
|
|
4 |
Copyright 2001 University of Cambridge
|
|
5 |
|
|
6 |
Some miscellaneous and add-hoc set theory concepts.
|
|
7 |
*)
|
|
8 |
|
|
9 |
|
|
10 |
|
|
11 |
UNITYMisc = Main +
|
|
12 |
constdefs
|
|
13 |
less_than :: i=>i
|
|
14 |
"less_than(A) == measure(A, %x. x)"
|
|
15 |
|
|
16 |
lessThan :: [i, i] => i
|
|
17 |
"lessThan(u,A) == {x:A. x<u}"
|
|
18 |
|
|
19 |
greaterThan :: [i,i]=> i
|
|
20 |
"greaterThan(u,A) == {x:A. u<x}"
|
|
21 |
|
|
22 |
(* Identity relation over a domain A *)
|
|
23 |
Identity :: i=>i
|
|
24 |
"Identity(A) == {p:A*A. EX x. p=<x,x>}"
|
|
25 |
end |