Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ca3351f

Browse files
committed
feat(rat/{basic,floor}) add floor lemmas (#5148)
1 parent 2b074be commit ca3351f

File tree

2 files changed

+111
-1
lines changed

2 files changed

+111
-1
lines changed

src/data/rat/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -586,6 +586,14 @@ begin
586586
simp [division_def, coe_int_eq_mk, mul_def one_ne_zero d0]
587587
end
588588

589+
lemma exists_eq_mul_div_num_and_eq_mul_div_denom {n d : ℤ} (n_ne_zero : n ≠ 0)
590+
(d_ne_zero : d ≠ 0) :
591+
∃ (c : ℤ), n = c * ((n : ℚ) / d).num ∧ (d : ℤ) = c * ((n : ℚ) / d).denom :=
592+
begin
593+
have : ((n : ℚ) / d) = rat.mk n d, by rw [←rat.mk_eq_div],
594+
exact rat.num_denom_mk n_ne_zero d_ne_zero this
595+
end
596+
589597
theorem coe_int_eq_of_int (z : ℤ) : ↑z = of_int z :=
590598
(coe_int_eq_mk z).trans (of_int_eq_mk z).symm
591599

src/data/rat/floor.lean

Lines changed: 103 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,8 @@ import algebra.floor
99
1010
## Summary
1111
12-
We define the `floor` function and the `floor_ring` instance on `ℚ`.
12+
We define the `floor` function and the `floor_ring` instance on `ℚ`. Some technical lemmas relating
13+
`floor` to integer division and modulo arithmetic are derived as well as some simple inequalities.
1314
1415
## Tags
1516
@@ -37,4 +38,105 @@ instance : floor_ring ℚ :=
3738

3839
protected lemma floor_def {q : ℚ} : ⌊q⌋ = q.num / q.denom := by { cases q, refl }
3940

41+
lemma floor_int_div_nat_eq_div {n : ℤ} {d : ℕ} : ⌊(↑n : ℚ) / (↑d : ℚ)⌋ = n / (↑d : ℤ) :=
42+
begin
43+
rw [rat.floor_def],
44+
cases decidable.em (d = 0) with d_eq_zero d_ne_zero,
45+
{ simp [d_eq_zero] },
46+
{ cases decidable.em (n = 0) with n_eq_zero n_ne_zero,
47+
{ simp [n_eq_zero] },
48+
{ set q := (n : ℚ) / d with q_eq,
49+
obtain ⟨c, n_eq_c_mul_num, d_eq_c_mul_denom⟩ : ∃ c, n = c * q.num ∧ (d : ℤ) = c * q.denom, by
50+
{ rw q_eq,
51+
exact_mod_cast (@rat.exists_eq_mul_div_num_and_eq_mul_div_denom n d n_ne_zero
52+
(by exact_mod_cast d_ne_zero)) },
53+
suffices : q.num / q.denom = c * q.num / (c * q.denom),
54+
by rwa [n_eq_c_mul_num, d_eq_c_mul_denom],
55+
suffices : c > 0, by solve_by_elim [int.mul_div_mul_of_pos],
56+
have q_denom_mul_c_pos : (0 : ℤ) < q.denom * c, by
57+
{ have : (d : ℤ) > 0, by exact_mod_cast (nat.pos_iff_ne_zero.elim_right d_ne_zero),
58+
rwa [d_eq_c_mul_denom, mul_comm] at this },
59+
suffices : (0 : ℤ) ≤ q.denom, from pos_of_mul_pos_left q_denom_mul_c_pos this,
60+
exact_mod_cast (le_of_lt q.pos) } }
61+
end
62+
63+
end rat
64+
65+
66+
lemma int.mod_nat_eq_sub_mul_floor_rat_div {n : ℤ} {d : ℕ} : n % d = n - d * ⌊(n : ℚ) / d⌋ :=
67+
by rw [(eq_sub_of_add_eq $ int.mod_add_div n d), rat.floor_int_div_nat_eq_div]
68+
69+
lemma nat.coprime_sub_mul_floor_rat_div_of_coprime {n d : ℕ} (n_coprime_d : n.coprime d) :
70+
((n : ℤ) - d * ⌊(n : ℚ)/ d⌋).nat_abs.coprime d :=
71+
begin
72+
have : (n : ℤ) % d = n - d * ⌊(n : ℚ)/ d⌋, from int.mod_nat_eq_sub_mul_floor_rat_div,
73+
rw ←this,
74+
have : d.coprime n, from n_coprime_d.symm,
75+
rwa [nat.coprime, nat.gcd_rec] at this
76+
end
77+
78+
79+
namespace rat
80+
81+
lemma num_lt_succ_floor_mul_denom (q : ℚ) : q.num < (⌊q⌋ + 1) * q.denom :=
82+
begin
83+
suffices : (q.num : ℚ) < (⌊q⌋ + 1) * q.denom, by exact_mod_cast this,
84+
suffices : (q.num : ℚ) < (q - fract q + 1) * q.denom, by
85+
{ have : (⌊q⌋ : ℚ) = q - fract q, from (eq_sub_of_add_eq $ floor_add_fract q),
86+
rwa this },
87+
suffices : (q.num : ℚ) < q.num + (1 - fract q) * q.denom, by
88+
{ have : (q - fract q + 1) * q.denom = q.num + (1 - fract q) * q.denom, calc
89+
(q - fract q + 1) * q.denom = (q + (1 - fract q)) * q.denom : by ring
90+
... = q * q.denom + (1 - fract q) * q.denom : by rw add_mul
91+
... = q.num + (1 - fract q) * q.denom : by simp,
92+
rwa this },
93+
suffices : 0 < (1 - fract q) * q.denom, by { rw ←sub_lt_iff_lt_add', simpa },
94+
have : 0 < 1 - fract q, by
95+
{ have : fract q < 1, from fract_lt_one q,
96+
have : 0 + fract q < 1, by simp [this],
97+
rwa lt_sub_iff_add_lt },
98+
exact mul_pos this (by exact_mod_cast q.pos)
99+
end
100+
101+
lemma fract_inv_num_lt_num_of_pos {q : ℚ} (q_pos : 0 < q): (fract q⁻¹).num < q.num :=
102+
begin
103+
-- we know that the numerator must be positive
104+
have q_num_pos : 0 < q.num, from rat.num_pos_iff_pos.elim_right q_pos,
105+
-- we will work with the absolute value of the numerator, which is equal to the numerator
106+
have q_num_abs_eq_q_num : (q.num.nat_abs : ℤ) = q.num, from
107+
(int.nat_abs_of_nonneg $ le_of_lt q_num_pos),
108+
set q_inv := (q.denom : ℚ) / q.num with q_inv_def,
109+
have q_inv_eq : q⁻¹ = q_inv, from rat.inv_def',
110+
suffices : (q_inv - ⌊q_inv⌋).num < q.num, by rwa q_inv_eq,
111+
suffices : ((q.denom - q.num * ⌊q_inv⌋ : ℚ) / q.num).num < q.num, by
112+
field_simp [this, (ne_of_gt q_num_pos)],
113+
suffices : (q.denom : ℤ) - q.num * ⌊q_inv⌋ < q.num, by
114+
{ -- use that `q.num` and `q.denom` are coprime to show that the numerator stays unreduced
115+
have : ((q.denom - q.num * ⌊q_inv⌋ : ℚ) / q.num).num = q.denom - q.num * ⌊q_inv⌋, by
116+
{ suffices : ((q.denom : ℤ) - q.num * ⌊q_inv⌋).nat_abs.coprime q.num.nat_abs, by
117+
exact_mod_cast (rat.num_div_eq_of_coprime q_num_pos this),
118+
have : (q.num.nat_abs : ℚ) = (q.num : ℚ), by exact_mod_cast q_num_abs_eq_q_num,
119+
have tmp := nat.coprime_sub_mul_floor_rat_div_of_coprime q.cop.symm,
120+
simpa only [this, q_num_abs_eq_q_num] using tmp },
121+
rwa this },
122+
-- to show the claim, start with the following inequality
123+
have q_inv_num_denom_ineq : q⁻¹.num - ⌊q⁻¹⌋ * q⁻¹.denom < q⁻¹.denom, by
124+
{ have : q⁻¹.num < (⌊q⁻¹⌋ + 1) * q⁻¹.denom, from rat.num_lt_succ_floor_mul_denom q⁻¹,
125+
have : q⁻¹.num < ⌊q⁻¹⌋ * q⁻¹.denom + q⁻¹.denom, by rwa [right_distrib, one_mul] at this,
126+
rwa [←sub_lt_iff_lt_add'] at this },
127+
-- use that `q.num` and `q.denom` are coprime to show that q_inv is the unreduced reciprocal
128+
-- of `q`
129+
have : q_inv.num = q.denom ∧ q_inv.denom = q.num.nat_abs, by
130+
{ have coprime_q_denom_q_num : q.denom.coprime q.num.nat_abs, from q.cop.symm,
131+
have : int.nat_abs q.denom = q.denom, by simp,
132+
rw ←this at coprime_q_denom_q_num,
133+
rw q_inv_def,
134+
split,
135+
{ exact_mod_cast (rat.num_div_eq_of_coprime q_num_pos coprime_q_denom_q_num) },
136+
{ suffices : (((q.denom : ℚ) / q.num).denom : ℤ) = q.num.nat_abs, by exact_mod_cast this,
137+
rw q_num_abs_eq_q_num,
138+
exact_mod_cast (rat.denom_div_eq_of_coprime q_num_pos coprime_q_denom_q_num) } },
139+
rwa [q_inv_eq, this.left, this.right, q_num_abs_eq_q_num, mul_comm] at q_inv_num_denom_ineq
140+
end
141+
40142
end rat

0 commit comments

Comments
 (0)