Commit 2024-02-26 00:38 bbe9baad

View on Github →

chore: add lemmas for nat literals corresponding to lemmas for nat casts (#8006) I loogled for every occurrence of "cast", Nat and "natCast" and where the casted nat was n, and made sure there were corresponding @[simp] lemmas for 0, 1, and OfNat.ofNat n. This is necessary in general for simp confluence. Example:

import Mathlib
variable {α : Type*} [LinearOrderedRing α] (m n : ℕ) [m.AtLeastTwo] [n.AtLeastTwo]
example : ((OfNat.ofNat m : ℕ) : α) ≤ ((OfNat.ofNat n : ℕ) : α) ↔ (OfNat.ofNat m : ℕ) ≤ (OfNat.ofNat n : ℕ) := by
  simp only [Nat.cast_le] -- this `@[simp]` lemma can apply
example : ((OfNat.ofNat m : ℕ) : α) ≤ ((OfNat.ofNat n : ℕ) : α) ↔ (OfNat.ofNat m : α) ≤ (OfNat.ofNat n : α) := by
  simp only [Nat.cast_ofNat] -- and so can this one
example : (OfNat.ofNat m : α) ≤ (OfNat.ofNat n : α) ↔ (OfNat.ofNat m : ℕ) ≤ (OfNat.ofNat n : ℕ) := by
  simp -- fails! `simp` doesn't have a lemma to bridge their results. confluence issue.

As far as I know, the only file this PR leaves with ofNat gaps is PartENat.lean. #8002 is addressing that file in parallel.

Estimated changes

modified theorem ENat.coe_toNat_eq_self
added theorem ENat.ofNat_ne_top
added theorem ENat.recTopCoe_ofNat
added theorem ENat.recTopCoe_one
added theorem ENat.recTopCoe_zero
added theorem ENat.toNat_ofNat
added theorem ENat.top_ne_ofNat
added theorem ENat.top_sub_ofNat
added theorem ENat.top_sub_one
added theorem Nat.abs_ofNat
modified theorem Nat.cast_add_one_pos
added theorem Nat.cast_le_ofNat
added theorem Nat.cast_lt_ofNat
added theorem Nat.not_ofNat_le_one
added theorem Nat.not_ofNat_lt_one
added theorem Nat.ofNat_le
added theorem Nat.ofNat_le_cast
added theorem Nat.ofNat_lt
added theorem Nat.ofNat_lt_cast
added theorem Nat.ofNat_nonneg'
added theorem Nat.ofNat_nonneg
added theorem Nat.ofNat_pos'
added theorem Nat.ofNat_pos
added theorem Nat.one_le_ofNat
added theorem Nat.one_lt_ofNat
added theorem ofDual_ofNat
added theorem ofLex_ofNat
added theorem toDual_ofNat
added theorem toLex_ofNat