Skip to content

Commit

Permalink
refactor: Let Nat.mod reduce more
Browse files Browse the repository at this point in the history
this refined upon #4098 and makes `Nat.mod` reduce on even more
literals. The key observation that I missed earlier is that `if m ≤ n`
reduces better than `if n < m`.

Also see discussion at
leanprover-community/mathlib4#12853 (comment)
  • Loading branch information
nomeata committed May 13, 2024
1 parent 7db8e64 commit e44488d
Show file tree
Hide file tree
Showing 2 changed files with 18 additions and 15 deletions.
3 changes: 0 additions & 3 deletions src/Init/Data/Fin/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -208,7 +208,6 @@ theorem val_add_one {n : Nat} (i : Fin (n + 1)) :
| .inl h => cases Fin.eq_of_val_eq h; simp
| .inr h => simpa [Fin.ne_of_lt h] using val_add_one_of_lt h

unseal Nat.modCore in
@[simp] theorem val_two {n : Nat} : (2 : Fin (n + 3)).val = 2 := rfl

theorem add_one_pos (i : Fin (n + 1)) (h : i < Fin.last n) : (0 : Fin (n + 1)) < i + 1 := by
Expand Down Expand Up @@ -243,7 +242,6 @@ theorem succ_ne_zero {n} : ∀ k : Fin n, Fin.succ k ≠ 0

@[simp] theorem succ_zero_eq_one : Fin.succ (0 : Fin (n + 1)) = 1 := rfl

unseal Nat.modCore in
/-- Version of `succ_one_eq_two` to be used by `dsimp` -/
@[simp] theorem succ_one_eq_two : Fin.succ (1 : Fin (n + 2)) = 2 := rfl

Expand Down Expand Up @@ -395,7 +393,6 @@ theorem castSucc_lt_last (a : Fin n) : castSucc a < last n := a.is_lt

@[simp] theorem castSucc_zero : castSucc (0 : Fin (n + 1)) = 0 := rfl

unseal Nat.modCore in
@[simp] theorem castSucc_one {n : Nat} : castSucc (1 : Fin (n + 2)) = 1 := rfl

/-- `castSucc i` is positive when `i` is positive -/
Expand Down
30 changes: 18 additions & 12 deletions src/Init/Data/Nat/Div.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,28 +82,34 @@ decreasing_by apply div_rec_lemma; assumption

@[extern "lean_nat_mod"]
protected def mod : @& Nat → @& Nat → Nat
/- These four cases are not needed mathematically, they are just special cases of the
general case. However, it makes `0 % n = 0` etc. true definitionally rather than just
propositionally.
/-
Nat.modCore is defined by well-founded recursion and thus irreducible. Nevertheless it is
desireable if trivial `Nat.mod` calculations, namely
* `Nat.mod 0 m` for all `m`
* `Nat.mod n (m+n)` for concrete literals `n`
reduce definitionally.
This property is desirable for `Fin n` literals, as it means `(ofNat 0 : Fin n).val = 0` by
definition. This was true in lean3 and it simplified things for mathlib if it remains true. -/
definition.
-/
| 0, _ => 0
| 1, 0 => 0
| 1, 1 => 0
| 1, (_+2) => 1
| x@(_ + 2), y => Nat.modCore x y
| n@(_ + 1), m =>
if m ≤ n -- NB: if n < m does not reduce as well as `m ≤ n`!
then Nat.modCore n m
else n

instance instMod : Mod Nat := ⟨Nat.mod⟩

protected theorem modCore_eq_mod (x y : Nat) : Nat.modCore x y = x % y := by
show Nat.modCore x y = Nat.mod x y
match x, y with
| 0, y =>
rw [Nat.modCore]
exact if_neg fun ⟨hlt, hle⟩ => Nat.lt_irrefl _ (Nat.lt_of_lt_of_le hlt hle)
| 1, 0 => rw [Nat.modCore]; rfl
| 1, 1 => rw [Nat.modCore, Nat.modCore]; rfl
| 1, (_+2) => rw [Nat.modCore]; rfl
| (_ + 2), _ => rfl
| (x + 1), y =>
rw [Nat.mod]; dsimp
refine iteInduction (fun _ => rfl) (fun h => ?false) -- cannot use `split` this early yet
rw [Nat.modCore]
exact if_neg fun ⟨_hlt, hle⟩ => h hle

theorem mod_eq (x y : Nat) : x % y = if 0 < y ∧ y ≤ x then (x - y) % y else x := by
rw [←Nat.modCore_eq_mod, ←Nat.modCore_eq_mod, Nat.modCore]
Expand Down

0 comments on commit e44488d

Please sign in to comment.