| Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > ILE Home > Th. List > df-fl | GIF version | ||
| Description: Define the floor
(greatest integer less than or equal to) function. See
flval 10379 for its value, flqlelt 10383 for its basic property, and flqcl 10380 for
its closure. For example, (⌊‘(3 / 2)) =
1 while
(⌊‘-(3 / 2)) = -2 (ex-fl 15455).
Although we define this on real numbers so that notations are similar to the Metamath Proof Explorer, in the absence of excluded middle few theorems will be possible for all real numbers. Imagine a real number which is around 2.99995 or 3.00001 . In order to determine whether its floor is 2 or 3, it would be necessary to compute the number to arbitrary precision. The term "floor" was coined by Ken Iverson. He also invented a mathematical notation for floor, consisting of an L-shaped left bracket and its reflection as a right bracket. In APL, the left-bracket alone is used, and we borrow this idea. (Thanks to Paul Chapman for this information.) (Contributed by NM, 14-Nov-2004.) |
| Ref | Expression |
|---|---|
| df-fl | ⊢ ⌊ = (𝑥 ∈ ℝ ↦ (℩𝑦 ∈ ℤ (𝑦 ≤ 𝑥 ∧ 𝑥 < (𝑦 + 1)))) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | cfl 10375 | . 2 class ⌊ | |
| 2 | vx | . . 3 setvar 𝑥 | |
| 3 | cr 7895 | . . 3 class ℝ | |
| 4 | vy | . . . . . . 7 setvar 𝑦 | |
| 5 | 4 | cv 1363 | . . . . . 6 class 𝑦 |
| 6 | 2 | cv 1363 | . . . . . 6 class 𝑥 |
| 7 | cle 8079 | . . . . . 6 class ≤ | |
| 8 | 5, 6, 7 | wbr 4034 | . . . . 5 wff 𝑦 ≤ 𝑥 |
| 9 | c1 7897 | . . . . . . 7 class 1 | |
| 10 | caddc 7899 | . . . . . . 7 class + | |
| 11 | 5, 9, 10 | co 5925 | . . . . . 6 class (𝑦 + 1) |
| 12 | clt 8078 | . . . . . 6 class < | |
| 13 | 6, 11, 12 | wbr 4034 | . . . . 5 wff 𝑥 < (𝑦 + 1) |
| 14 | 8, 13 | wa 104 | . . . 4 wff (𝑦 ≤ 𝑥 ∧ 𝑥 < (𝑦 + 1)) |
| 15 | cz 9343 | . . . 4 class ℤ | |
| 16 | 14, 4, 15 | crio 5879 | . . 3 class (℩𝑦 ∈ ℤ (𝑦 ≤ 𝑥 ∧ 𝑥 < (𝑦 + 1))) |
| 17 | 2, 3, 16 | cmpt 4095 | . 2 class (𝑥 ∈ ℝ ↦ (℩𝑦 ∈ ℤ (𝑦 ≤ 𝑥 ∧ 𝑥 < (𝑦 + 1)))) |
| 18 | 1, 17 | wceq 1364 | 1 wff ⌊ = (𝑥 ∈ ℝ ↦ (℩𝑦 ∈ ℤ (𝑦 ≤ 𝑥 ∧ 𝑥 < (𝑦 + 1)))) |
| Colors of variables: wff set class |
| This definition is referenced by: flval 10379 |
| Copyright terms: Public domain | W3C validator |