Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(data/nat/digits): add last_digit_ne_zero #3544

Closed
wants to merge 8 commits into from
9 changes: 9 additions & 0 deletions src/data/list/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -630,6 +630,15 @@ theorem last_mem : ∀ {l : list α} (h : l ≠ []), last l h ∈ l
| [a] h := or.inl rfl
| (a::b::l) h := or.inr $ by { rw [last_cons_cons], exact last_mem (cons_ne_nil b l) }

lemma last_repeat_succ (a m : ℕ) :
(repeat a m.succ).last (ne_nil_of_length_eq_succ
(show (repeat a m.succ).length = m.succ, by rw length_repeat)) = a :=
begin
induction m with k IH,
{ simp },
{ simpa only [repeat_succ, last] }
end

/-! ### last' -/

@[simp] theorem last'_is_none :
Expand Down
112 changes: 85 additions & 27 deletions src/data/nat/digits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -133,6 +133,8 @@ begin
{ dsimp [of_digits], rw ih, },
end

@[simp] lemma of_digits_singleton {b n : ℕ} : of_digits b [n] = n := by simp [of_digits]

@[simp] lemma of_digits_one_cons {α : Type*} [semiring α] (h : ℕ) (L : list ℕ) :
of_digits (1 : α) (h :: L) = h + of_digits 1 L :=
by simp [of_digits]
Expand Down Expand Up @@ -236,10 +238,69 @@ begin
{ simp [of_digits, list.sum_cons, ih], }
end

/-! ### Inequalities
/-! ### Properties

This section contains various lemmas of properties relating to `digits` and `of_digits`. -/
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved

lemma digits_eq_nil_iff_eq_zero (b n : ℕ) : digits b n = [] ↔ n = 0 :=
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
begin
split,
{ intro h,
have : of_digits b (digits b n) = of_digits b [], by rw h,
convert this,
rw of_digits_digits },
{ rintro rfl,
simp }
end

lemma digits_ne_nil_iff_ne_zero (b n : ℕ) : digits b n ≠ [] ↔ n ≠ 0 :=
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
⟨λ h, h ∘ (digits_eq_nil_iff_eq_zero _ _).mpr,
λ h, h ∘ (digits_eq_nil_iff_eq_zero _ _).mp⟩
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved

This section contains various lemmas with inequalities relating to `digits` and `of_digits`.
-/
private lemma digits_last_aux1 {α : Type*} {L M : list α} (h : L = M) (p q) : L.last p = M.last q :=
by simp [h]

private lemma digits_last_aux2 (b n : ℕ) (h : 2 ≤ b) (w : 0 < n) :
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
digits b n = ((n % b) :: digits b (n / b)) :=
begin
rcases b with _|_|b,
{ finish },
{ norm_num at h },
rcases n with _|n,
{ norm_num at w },
simp,
end

lemma digits_last (b : ℕ) (h : 2 ≤ b) (m : ℕ) (hm : 0 < m) (p q) :
(digits b m).last p = (digits b (m/b)).last q :=
begin
rw digits_last_aux1 (digits_last_aux2 b m h hm),
{ rw list.last_cons, },
exact list.cons_ne_nil _ _,
end
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved

lemma last_digit_ne_zero (b m : ℕ) (hm : m ≠ 0) (p):
(digits b m).last p ≠ 0 :=
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This lemma would be more convenient if you remove the argument p and explicitly prove it using digits_ne_nil_iff_ne_zero.mpr hm.

Copy link
Collaborator Author

@shingtaklam1324 shingtaklam1324 Jul 25, 2020

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

To be honest that was my first approach to use digits_ne_nil_iff_ne_zero.mpr hm. I really struggled to get nat.strong_induction_on to work with it, hence I've generalized over the proof. (see the use of revert p in the proof). There might be a better approach but this is the best I could come up with.

I can mark this as private and add in a lemma with the proof provided, would that be ok? So change this to private lemma last_digit_ne_zero_aux and add

lemma last_digit_ne_zero {b m : ℕ} (hm : m ≠ 0) :
  (digits b m).last (digits_ne_nil_iff_ne_zero.mpr hm) ≠ 0 :=
last_digit_ne_zero_aux b m hm (digits_ne_nil_iff_ne_zero.mpr hm)

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I've done the above. So marked this one as private lemma last_digit_ne_zero_aux and then added another lemma that fills in the proof.

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

That works. Another option was to use as first tactic generalize : digits_ne_nil_iff_ne_zero.mpr hm = p,

begin
rcases b with _|_|b,
{ cases m; finish },
{ cases m, finish,
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
simp_rw [digits_one, list.last_repeat_succ 1 m],
norm_num },
revert hm p,
apply nat.strong_induction_on m,
intros n IH hn p,
have hnpos : 0 < n := nat.pos_of_ne_zero hn,
by_cases hnb : n < b + 2,
{ simp_rw [digits_of_lt b.succ.succ n hnpos hnb],
exact nat.pos_iff_ne_zero.mp hnpos },
{ rw digits_last b.succ.succ dec_trivial _ hnpos,
refine IH _ (nat.div_lt_self hnpos dec_trivial) _ _,
{ rw ←nat.pos_iff_ne_zero,
exact nat.div_pos (le_of_not_lt hnb) dec_trivial },
{ rw [digits_ne_nil_iff_ne_zero, ←nat.pos_iff_ne_zero],
exact nat.div_pos (le_of_not_lt hnb) dec_trivial } },
end

/-- The digits in the base b+2 expansion of n are all less than b+2 -/
lemma digits_lt_base' {b m : ℕ} : ∀ {d}, d ∈ digits (b+2) m → d < b+2 :=
Expand Down Expand Up @@ -330,33 +391,30 @@ end
lemma le_digits_len_le (b n m : ℕ) (h : n ≤ m) : (digits b n).length ≤ (digits b m).length :=
monotone_of_monotone_nat (digits_len_le_digits_len_succ b) h

-- future refactor to proof using lists (see lt_base_pow_length_digits)
-- See review in #3498. List proof depends on : m ≠ 0 → (digits b m) ≠ 0
lemma base_pow_length_digits_le' (b m : ℕ) : m ≠ 0 → (b + 2) ^ ((digits (b + 2) m).length) ≤ (b + 2) * m :=
lemma pow_length_le_mul_of_digits {b : ℕ} {l : list ℕ} (hl : l ≠ []) (hl2 : l.last hl ≠ 0):
(b + 2) ^ l.length ≤ (b + 2) * of_digits (b+2) l :=
begin
apply nat.strong_induction_on m,
clear m,
intros n IH npos,
cases n,
{ contradiction },
{ unfold digits at IH ⊢,
rw [digits_aux_def (b+2) (by simp) n.succ nat.succ_pos', list.length_cons],
specialize IH ((n.succ) / (b+2)) (nat.div_lt_self' n b),
rw [nat.pow_add, nat.pow_one, nat.mul_comm],
cases nat.lt_or_ge n.succ (b+2),
{ rw [nat.div_eq_of_lt h, digits_aux_zero, list.length, nat.pow_zero],
apply nat.mul_le_mul_left,
exact nat.succ_pos n },
{ have geb : (n.succ / (b + 2)) ≥ 1 := nat.div_pos h (by linarith),
specialize IH (by linarith [geb]),
replace IH := nat.mul_le_mul_left (b + 2) IH,
have IH' : (b + 2) * ((b + 2) * (n.succ / (b + 2))) ≤ (b + 2) * n.succ,
{ apply nat.mul_le_mul_left,
rw nat.mul_comm,
exact nat.div_mul_le_self n.succ (b + 2) },
exact le_trans IH IH' } }
rw [←list.init_append_last hl],
simp only [list.length_append, list.length, zero_add, list.length_init, of_digits_append, list.length_init,
of_digits_singleton, add_comm (l.length - 1), nat.pow_add, nat.pow_one],
apply nat.mul_le_mul_left,
refine le_trans _ (nat.le_add_left _ _),
have : 0 < l.last hl, { rwa [nat.pos_iff_ne_zero] },
convert nat.mul_le_mul_left _ this, rw [mul_one]
end

/-- Any non-zero natural number `m` is greater than
(b+2)^((number of digits in the base (b+2) representation of m) - 1)-/
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
lemma base_pow_length_digits_le' (b m : ℕ) (hm : m ≠ 0) :
(b + 2) ^ ((digits (b + 2) m).length) ≤ (b + 2) * m :=
begin
have : digits (b + 2) m ≠ [], from (digits_ne_nil_iff_ne_zero _ _).mpr hm,
convert pow_length_le_mul_of_digits this (last_digit_ne_zero _ _ _ this),
rwa of_digits_digits,
end

/-- Any non-zero natural number `m` is greater than
b^((number of digits in the base b representation of m) - 1)-/
shingtaklam1324 marked this conversation as resolved.
Show resolved Hide resolved
lemma base_pow_length_digits_le (b m : ℕ) (hb : 2 ≤ b): m ≠ 0 → b ^ ((digits b m).length) ≤ b * m :=
begin
rcases b with _ | _ | b; try { linarith },
Expand Down