author | nipkow |
Tue, 23 May 2000 07:32:24 +0200 | |
changeset 8924 | c434283b4cfa |
child 8957 | 26b6e8f43305 |
permissions | -rw-r--r-- |
(* Title: HOL/SetInterval.thy ID: $Id$ Author: Tobias Nipkow Copyright 1998 TU Muenchen lessThan, greaterThan, atLeast, atMost *) SetInterval = equalities + constdefs lessThan :: "('a::ord) => 'a set" ("(1{.._'(})") "{..u(} == {x. x<u}" atMost :: "('a::ord) => 'a set" ("(1{.._})") "{..u} == {x. x<=u}" greaterThan :: "('a::ord) => 'a set" ("(1{')_..})") "{)l..} == {x. l<x}" atLeast :: "('a::ord) => 'a set" ("(1{_..})") "{l..} == {x. l<=x}" end