| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > df-ceil | Structured version Visualization version GIF version | ||
| Description: The ceiling (least
integer greater than or equal to) function. Defined in
ISO 80000-2:2009(E) operation 2-9.18 and the "NIST Digital Library of
Mathematical Functions" , front introduction, "Common Notations
and
Definitions" section at http://dlmf.nist.gov/front/introduction#Sx4.
See ceilval 13878 for its value, ceilge 13885 and ceilm1lt 13888 for its basic
properties, and ceilcl 13882 for its closure. For example,
(⌈‘(3 / 2)) = 2 while (⌈‘-(3 / 2)) = -1
(ex-ceil 30467).
The symbol ⌈ is inspired by the gamma shaped left bracket of the usual notation. (Contributed by David A. Wheeler, 19-May-2015.) |
| Ref | Expression |
|---|---|
| df-ceil | ⊢ ⌈ = (𝑥 ∈ ℝ ↦ -(⌊‘-𝑥)) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | cceil 13831 | . 2 class ⌈ | |
| 2 | vx | . . 3 setvar 𝑥 | |
| 3 | cr 11154 | . . 3 class ℝ | |
| 4 | 2 | cv 1539 | . . . . . 6 class 𝑥 |
| 5 | 4 | cneg 11493 | . . . . 5 class -𝑥 |
| 6 | cfl 13830 | . . . . 5 class ⌊ | |
| 7 | 5, 6 | cfv 6561 | . . . 4 class (⌊‘-𝑥) |
| 8 | 7 | cneg 11493 | . . 3 class -(⌊‘-𝑥) |
| 9 | 2, 3, 8 | cmpt 5225 | . 2 class (𝑥 ∈ ℝ ↦ -(⌊‘-𝑥)) |
| 10 | 1, 9 | wceq 1540 | 1 wff ⌈ = (𝑥 ∈ ℝ ↦ -(⌊‘-𝑥)) |
| Colors of variables: wff setvar class |
| This definition is referenced by: ceilval 13878 dfceil2 13879 |
| Copyright terms: Public domain | W3C validator |