src/HOL/Hyperreal/IntFloor.thy
author oheimb
Fri, 11 Jul 2003 14:12:06 +0200
changeset 14100 804be4c4b642
parent 13958 c1c67582c9b5
child 14425 0a76d4633bb6
permissions -rw-r--r--
added map_image, restrict_map, some thms

(*  Title:       IntFloor.thy
    Author:      Jacques D. Fleuriot
    Copyright:   2001,2002  University of Edinburgh
    Description: Floor and ceiling operations over reals
*)

IntFloor = Integration + 

constdefs
    
    floor :: real => int
   "floor r == (LEAST n. r < real (n + (1::int)))"

    ceiling :: real => int
    "ceiling r == - floor (- r)"
  
end