Lemmas on Int.floor, Int.ceil and Int.fract #
This file contains basic results on the integer-valued floor and ceiling functions, as well as the fractional part operator.
TODO #
LinearOrderedRing can be relaxed to OrderedRing in many lemmas.
Tags #
rounding, floor, ceil
Extension for the positivity tactic: Int.floor is nonnegative if its input is.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Extension for the positivity tactic: Nat.ceil is positive if its input is.
Instances For
Extension for the positivity tactic: Int.ceil is positive/nonnegative if its input is.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Floor rings #
Floor #
@[simp]
@[simp]
theorem
Int.sub_one_lt_floor
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.floor_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(z : ℤ)
:
@[simp]
theorem
Int.floor_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
:
@[simp]
theorem
Int.floor_zero
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.floor_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.floor_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
theorem
Int.floor_add_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(z : ℤ)
:
@[simp]
theorem
Int.floor_add_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.floor_intCast_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(z : ℤ)
(a : R)
:
@[simp]
theorem
Int.floor_add_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
:
@[simp]
theorem
Int.floor_add_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
theorem
Int.floor_natCast_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
(a : R)
:
@[simp]
theorem
Int.floor_ofNat_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
[n.AtLeastTwo]
(a : R)
:
@[simp]
theorem
Int.floor_sub_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(z : ℤ)
:
@[simp]
theorem
Int.floor_sub_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
:
@[simp]
theorem
Int.floor_sub_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.floor_sub_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
theorem
Int.floor_eq_self_iff_mem
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.floor_div_natCast
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
(a : k)
(n : ℕ)
:
Fractional part #
@[simp]
theorem
Int.fract_add_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(m : ℤ)
:
@[simp]
theorem
Int.fract_add_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(m : ℕ)
:
@[simp]
theorem
Int.fract_add_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.fract_add_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
theorem
Int.fract_intCast_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(m : ℤ)
(a : R)
:
@[simp]
theorem
Int.fract_natCast_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
(a : R)
:
@[simp]
theorem
Int.fract_one_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.fract_ofNat_add
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
[n.AtLeastTwo]
(a : R)
:
@[simp]
theorem
Int.fract_sub_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(m : ℤ)
:
@[simp]
theorem
Int.fract_sub_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
:
@[simp]
theorem
Int.fract_sub_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.fract_sub_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
theorem
Int.fract_add_le
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a b : R)
:
theorem
Int.fract_add_fract_le
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a b : R)
:
@[simp]
theorem
Int.fract_nonneg
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.fract_pos
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
:
The fractional part of a is positive if and only if a ≠ ⌊a⌋.
theorem
Int.fract_lt_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.fract_zero
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.fract_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
theorem
Int.abs_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.abs_one_sub_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.fract_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(z : ℤ)
:
@[simp]
theorem
Int.fract_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
:
@[simp]
theorem
Int.fract_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
[n.AtLeastTwo]
:
theorem
Int.fract_floor
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.floor_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.fract_eq_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{a b : R}
:
@[simp]
theorem
Int.fract_eq_self
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{a : R}
:
@[simp]
theorem
Int.fract_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.fract_neg
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{x : R}
(hx : fract x ≠ 0)
:
@[simp]
theorem
Int.fract_neg_eq_zero
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{x : R}
:
theorem
Int.fract_mul_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(b : ℕ)
:
theorem
Int.fract_div_mul_self_mem_Ico
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
(a b : k)
(ha : 0 < a)
:
theorem
Int.sub_floor_div_mul_nonneg
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{b : k}
(a : k)
(hb : 0 < b)
:
theorem
Int.sub_floor_div_mul_lt
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{b : k}
(a : k)
(hb : 0 < b)
:
theorem
Int.fract_div_natCast_eq_div_natCast_mod
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{m n : ℕ}
:
theorem
Int.fract_div_intCast_eq_div_intCast_mod
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{m : ℤ}
{n : ℕ}
:
Ceil #
theorem
Int.ceil_nonneg_of_neg_one_lt
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
(ha : -1 < a)
:
theorem
Int.floor_neg
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
:
theorem
Int.ceil_neg
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.ceil_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(z : ℤ)
:
@[simp]
theorem
Int.ceil_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
:
@[simp]
theorem
Int.ceil_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
theorem
Int.ceil_add_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(z : ℤ)
:
@[simp]
theorem
Int.ceil_add_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
:
@[simp]
theorem
Int.ceil_add_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.ceil_add_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
@[simp]
theorem
Int.ceil_sub_intCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(z : ℤ)
:
@[simp]
theorem
Int.ceil_sub_natCast
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
:
@[simp]
theorem
Int.ceil_sub_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.ceil_sub_ofNat
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
(n : ℕ)
[n.AtLeastTwo]
:
theorem
Int.ceil_lt_add_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
@[simp]
theorem
Int.ceil_zero
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
@[simp]
theorem
Int.ceil_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
:
theorem
Int.ceil_eq_on_Ioc'
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(z : ℤ)
(a : R)
:
theorem
Int.ceil_eq_self_iff_mem
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.floor_le_ceil
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
theorem
Int.floor_lt_ceil_of_lt
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{a b : R}
(h : a < b)
:
@[deprecated Int.ceil_eq_floor_add_one_iff_notMem (since := "2025-05-23")]
theorem
Int.ceil_eq_floor_add_one_iff_not_mem
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
(a : R)
:
Alias of Int.ceil_eq_floor_add_one_iff_notMem.
theorem
Int.ceil_eq_add_one_sub_fract
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
(ha : fract a ≠ 0)
:
theorem
Int.ceil_sub_self_eq
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
{a : R}
[IsStrictOrderedRing R]
(ha : fract a ≠ 0)
:
theorem
Int.div_two_lt_floor
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{a : k}
(ha : 1 ≤ a)
:
theorem
Int.ceil_lt_two_mul
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{a : k}
(ha : 2⁻¹ < a)
:
theorem
Int.ceil_le_two_mul
{k : Type u_4}
[Field k]
[LinearOrder k]
[IsStrictOrderedRing k]
[FloorRing k]
{a : k}
(ha : 2⁻¹ ≤ a)
:
Intervals #
theorem
Int.map_floor
{F : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Ring R]
[LinearOrder R]
[Ring S]
[LinearOrder S]
[FloorRing R]
[FloorRing S]
[FunLike F R S]
[RingHomClass F R S]
(f : F)
(hf : StrictMono ⇑f)
(a : R)
:
theorem
Int.map_ceil
{F : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Ring R]
[LinearOrder R]
[Ring S]
[LinearOrder S]
[FloorRing R]
[FloorRing S]
[FunLike F R S]
[RingHomClass F R S]
(f : F)
(hf : StrictMono ⇑f)
(a : R)
:
theorem
Int.map_fract
{F : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Ring R]
[LinearOrder R]
[Ring S]
[LinearOrder S]
[FloorRing R]
[FloorRing S]
[FunLike F R S]
[RingHomClass F R S]
(f : F)
(hf : StrictMono ⇑f)
(a : R)
:
theorem
Nat.ceil_lt_add_one_of_gt_neg_one
{R : Type u_2}
[Ring R]
[LinearOrder R]
[FloorRing R]
[IsStrictOrderedRing R]
{a : R}
(ha : -1 < a)
:
a variant of Nat.ceil_lt_add_one with its condition 0 ≤ a generalized to -1 < a
A floor ring as a floor semiring #
theorem
Int.natCast_floor_eq_floor
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : 0 ≤ a)
:
theorem
Int.natCast_ceil_eq_ceil
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : 0 ≤ a)
:
theorem
Int.natCast_ceil_eq_ceil_of_neg_one_lt
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : -1 < a)
:
theorem
natCast_floor_eq_intCast_floor
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : 0 ≤ a)
:
theorem
natCast_ceil_eq_intCast_ceil
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : 0 ≤ a)
:
theorem
natCast_ceil_eq_intCast_ceil_of_neg_one_lt
{R : Type u_2}
[Ring R]
[LinearOrder R]
[IsStrictOrderedRing R]
[FloorRing R]
{a : R}
(ha : -1 < a)
:
There exists at most one FloorRing structure on a given linear ordered ring.