Documentation

Mathlib.Data.Rat.Floor

Floor Function for Rational Numbers #

Summary #

We define the FloorRing instance on . Some technical lemmas relating floor to integer division and modulo arithmetic are derived as well as some simple inequalities.

Tags #

rat, rationals, ℚ, floor

theorem Rat.floor_def' (a : ) :
a.floor = a.num / a.den
theorem Rat.le_floor {z : } {r : } :
z r.floor z r
theorem Rat.floor_def {q : } :
q = q.num / q.den
theorem Rat.ceil_def (q : ) :
q = -(-q.num / q.den)
theorem Rat.floor_intCast_div_natCast (n : ) (d : ) :
n / d = n / d
theorem Rat.ceil_intCast_div_natCast (n : ) (d : ) :
n / d = -(-n / d)
theorem Rat.floor_natCast_div_natCast (n d : ) :
n / d = n / d
theorem Rat.ceil_natCast_div_natCast (n d : ) :
n / d = -(-n / d)
@[simp]
theorem Rat.floor_cast {α : Type u_1} [LinearOrderedField α] [FloorRing α] (x : ) :
@[simp]
theorem Rat.ceil_cast {α : Type u_1} [LinearOrderedField α] [FloorRing α] (x : ) :
@[simp]
theorem Rat.round_cast {α : Type u_1} [LinearOrderedField α] [FloorRing α] (x : ) :
round x = round x
@[simp]
theorem Rat.cast_fract {α : Type u_1} [LinearOrderedField α] [FloorRing α] (x : ) :
(Int.fract x) = Int.fract x
theorem Int.mod_nat_eq_sub_mul_floor_rat_div {n : } {d : } :
n % d = n - d * n / d
theorem Nat.coprime_sub_mul_floor_rat_div_of_coprime {n d : } (n_coprime_d : n.Coprime d) :
(n - d * n / d).natAbs.Coprime d