Skip to content

Commit

Permalink
feat(data/nat/basic): add a few lemmas (#14718)
Browse files Browse the repository at this point in the history
Add a few lemmas about sub and mod.
  • Loading branch information
user7230724 committed Jun 16, 2022
1 parent 3feb151 commit b05d845
Show file tree
Hide file tree
Showing 2 changed files with 12 additions and 0 deletions.
3 changes: 3 additions & 0 deletions src/algebra/order/sub.lean
Expand Up @@ -439,6 +439,9 @@ end
lemma tsub_tsub_tsub_cancel_right (h : c ≤ b) : (a - c) - (b - c) = a - b :=
by rw [tsub_tsub, add_tsub_cancel_of_le h]

lemma tsub_lt_of_lt (h : a < b) : a - c < b :=
lt_of_le_of_lt tsub_le_self h

/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/

namespace add_le_cancellable
Expand Down
9 changes: 9 additions & 0 deletions src/data/nat/basic.lean
Expand Up @@ -1289,6 +1289,15 @@ lemma dvd_left_injective : function.injective ((∣) : ℕ → ℕ → Prop) :=
lemma div_lt_div_of_lt_of_dvd {a b d : ℕ} (hdb : d ∣ b) (h : a < b) : a / d < b / d :=
by { rw nat.lt_div_iff_mul_lt hdb, exact lt_of_le_of_lt (mul_div_le a d) h }

lemma mul_add_mod (a b c : ℕ) : (a * b + c) % b = c % b :=
by simp [nat.add_mod]

lemma mul_add_mod_of_lt {a b c : ℕ} (h : c < b) : (a * b + c) % b = c :=
by rw [nat.mul_add_mod, nat.mod_eq_of_lt h]

lemma pred_eq_self_iff {n : ℕ} : n.pred = n ↔ n = 0 :=
by { cases n; simp [(nat.succ_ne_self _).symm] }

/-! ### `find` -/
section find

Expand Down

0 comments on commit b05d845

Please sign in to comment.