Skip to content

Commit a45fd77

Browse files
committed
feat(Data/Fin): add lemmas (#8974)
- move `Nat.dvd_one` to `Data.Nat.Basic`; it should go to Std4; - rename `Fin.ofNat_eq_val` to `Fin.ofNat''_eq_cast`; - add `@[simp]` lemmas `Fin.val_nat_cast`, `Fin.nat_cast_self`, and `Fin.nat_cast_eq_zero`; - add `@[simp]` to `Fin.cast_nat_eq_last` and `ZMod.val_nat_cast`; - add `binomial_apply_last`, as the LHS of `binomial_apply_self` is no longer in simp normal form.
1 parent eb4573a commit a45fd77

File tree

5 files changed

+38
-42
lines changed

5 files changed

+38
-42
lines changed

Mathlib/Data/Fin/Basic.lean

Lines changed: 20 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -640,23 +640,23 @@ end Bit
640640
section OfNatCoe
641641

642642
@[simp]
643-
theorem ofNat_eq_val (n : ℕ) [NeZero n] (a : ℕ) : (Fin.ofNat'' a : Fin n) = a :=
643+
theorem ofNat''_eq_cast (n : ℕ) [NeZero n] (a : ℕ) : (Fin.ofNat'' a : Fin n) = a :=
644644
rfl
645-
#align fin.of_nat_eq_coe Fin.ofNat_eq_val
645+
#align fin.of_nat_eq_coe Fin.ofNat''_eq_cast
646+
647+
@[simp] lemma val_nat_cast (a n : ℕ) [NeZero n] : (a : Fin n).val = a % n := rfl
646648

647649
-- porting note: is this the right name for things involving `Nat.cast`?
648650
/-- Converting an in-range number to `Fin (n + 1)` produces a result
649651
whose value is the original number. -/
650-
theorem val_cast_of_lt {n : ℕ} [NeZero n] {a : ℕ} (h : a < n) : (a : Fin n).val = a := by
651-
rw [← ofNat_eq_val]
652-
exact Nat.mod_eq_of_lt h
652+
theorem val_cast_of_lt {n : ℕ} [NeZero n] {a : ℕ} (h : a < n) : (a : Fin n).val = a :=
653+
Nat.mod_eq_of_lt h
653654
#align fin.coe_val_of_lt Fin.val_cast_of_lt
654655

655656
/-- Converting the value of a `Fin (n + 1)` to `Fin (n + 1)` results
656657
in the same value. -/
657-
theorem cast_val_eq_self {n : ℕ} [NeZero n] (a : Fin n) : (a.val : Fin n) = a := by
658-
rw [Fin.eq_iff_veq]
659-
exact val_cast_of_lt a.isLt
658+
theorem cast_val_eq_self {n : ℕ} [NeZero n] (a : Fin n) : (a.val : Fin n) = a :=
659+
ext <| val_cast_of_lt a.isLt
660660
#align fin.coe_val_eq_self Fin.cast_val_eq_self
661661

662662
-- porting note: this is syntactically the same as `val_cast_of_lt`
@@ -665,9 +665,13 @@ theorem cast_val_eq_self {n : ℕ} [NeZero n] (a : Fin n) : (a.val : Fin n) = a
665665
-- porting note: this is syntactically the same as `cast_val_of_lt`
666666
#align fin.coe_coe_eq_self Fin.cast_val_eq_self
667667

668-
theorem cast_nat_eq_last (n) : (n : Fin (n + 1)) = Fin.last n := by
669-
rw [← Fin.ofNat_eq_val, Fin.ofNat'', Fin.last]
670-
simp only [Nat.mod_eq_of_lt n.lt_succ_self]
668+
@[simp] lemma nat_cast_self (n : ℕ) [NeZero n] : (n : Fin n) = 0 := by ext; simp
669+
670+
@[simp] lemma nat_cast_eq_zero {a n : ℕ} [NeZero n] : (a : Fin n) = 0 ↔ n ∣ a := by
671+
simp [eq_iff_veq, Nat.dvd_iff_mod_eq_zero]
672+
673+
@[simp]
674+
theorem cast_nat_eq_last (n) : (n : Fin (n + 1)) = Fin.last n := by ext; simp
671675
#align fin.coe_nat_eq_last Fin.cast_nat_eq_last
672676

673677
theorem le_val_last (i : Fin (n + 1)) : i ≤ n := by
@@ -682,19 +686,13 @@ end OfNatCoe
682686
#align fin.zero_ne_one Fin.zero_ne_one
683687

684688
@[simp]
685-
theorem zero_eq_one_iff [NeZero n] : (0 : Fin n) = 1 ↔ n = 1 := by
686-
constructor
687-
· intro h
688-
have := congr_arg ((↑) : Fin n → ℕ) h
689-
simp only [val_zero', val_one', @eq_comm _ 0, ← Nat.dvd_iff_mod_eq_zero] at this
690-
exact eq_one_of_dvd_one this
691-
· rintro rfl
692-
rfl
693-
#align fin.zero_eq_one_iff Fin.zero_eq_one_iff
689+
theorem one_eq_zero_iff [NeZero n] : (1 : Fin n) = 0 ↔ n = 1 := by
690+
rw [← Nat.cast_one, nat_cast_eq_zero, Nat.dvd_one]
691+
#align fin.one_eq_zero_iff Fin.one_eq_zero_iff
694692

695693
@[simp]
696-
theorem one_eq_zero_iff [NeZero n] : (1 : Fin n) = 0 ↔ n = 1 := by rw [eq_comm, zero_eq_one_iff]
697-
#align fin.one_eq_zero_iff Fin.one_eq_zero_iff
694+
theorem zero_eq_one_iff [NeZero n] : (0 : Fin n) = 1 ↔ n = 1 := by rw [eq_comm, one_eq_zero_iff]
695+
#align fin.zero_eq_one_iff Fin.zero_eq_one_iff
698696

699697
end Add
700698

Mathlib/Data/Nat/Basic.lean

Lines changed: 6 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -782,11 +782,14 @@ protected theorem mul_dvd_mul_iff_right {a b c : ℕ} (hc : 0 < c) : a * c ∣ b
782782
exists_congr fun d => by rw [Nat.mul_right_comm, mul_left_inj' hc.ne']
783783
#align nat.mul_dvd_mul_iff_right Nat.mul_dvd_mul_iff_right
784784

785+
@[simp] -- TODO: move to Std4
786+
theorem dvd_one {a : ℕ} : a ∣ 1 ↔ a = 1 :=
787+
⟨eq_one_of_dvd_one, fun h ↦ h.symm ▸ Nat.dvd_refl _⟩
788+
#align nat.dvd_one Nat.dvd_one
789+
785790
@[simp]
786791
theorem mod_mod_of_dvd (n : Nat) {m k : Nat} (h : m ∣ k) : n % k % m = n % m := by
787-
conv =>
788-
rhs
789-
rw [← mod_add_div n k]
792+
conv_rhs => rw [← mod_add_div n k]
790793
rcases h with ⟨t, rfl⟩
791794
rw [mul_assoc, add_mul_mod_self_left]
792795
#align nat.mod_mod_of_dvd Nat.mod_mod_of_dvd

Mathlib/Data/Nat/Order/Lemmas.lean

Lines changed: 0 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -93,12 +93,6 @@ protected lemma div_pos_iff (hb : b ≠ 0) : 0 < a / b ↔ b ≤ a := by
9393

9494
/-! ### `mod`, `dvd` -/
9595

96-
97-
@[simp]
98-
protected theorem dvd_one {n : ℕ} : n ∣ 1 ↔ n = 1 :=
99-
⟨eq_one_of_dvd_one, fun e => e.symm ▸ dvd_rfl⟩
100-
#align nat.dvd_one Nat.dvd_one
101-
10296
set_option linter.deprecated false in
10397
@[simp]
10498
protected theorem not_two_dvd_bit1 (n : ℕ) : ¬2 ∣ bit1 n := by

Mathlib/Data/ZMod/Basic.lean

Lines changed: 7 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -81,25 +81,23 @@ theorem val_mul' {m n : ZMod 0} : (m * n).val = m.val * n.val :=
8181
Int.natAbs_mul m n
8282
#align zmod.val_mul' ZMod.val_mul'
8383

84+
@[simp]
8485
theorem val_nat_cast {n : ℕ} (a : ℕ) : (a : ZMod n).val = a % n := by
8586
cases n
8687
· rw [Nat.mod_zero]
8788
exact Int.natAbs_ofNat a
88-
rw [← Fin.ofNat_eq_val]
89-
rfl
89+
· apply Fin.val_nat_cast
9090
#align zmod.val_nat_cast ZMod.val_nat_cast
9191

9292
theorem val_nat_cast_of_lt {n a : ℕ} (h : a < n) : (a : ZMod n).val = a := by
9393
rwa [val_nat_cast, Nat.mod_eq_of_lt]
9494

9595
instance charP (n : ℕ) : CharP (ZMod n) n where
96-
cast_eq_zero_iff' := by
97-
intro k
98-
cases' n with n
99-
· simp [zero_dvd_iff, Int.coe_nat_eq_zero, Nat.zero_eq]
100-
rw [Fin.eq_iff_veq]
101-
show (k : ZMod (n + 1)).val = (0 : ZMod (n + 1)).val ↔ _
102-
rw [val_nat_cast, val_zero, Nat.dvd_iff_mod_eq_zero]
96+
cast_eq_zero_iff' := by
97+
intro k
98+
cases' n with n
99+
· simp [zero_dvd_iff, Int.coe_nat_eq_zero, Nat.zero_eq]
100+
· exact Fin.nat_cast_eq_zero
103101

104102
@[simp]
105103
theorem addOrderOf_one (n : ℕ) : addOrderOf (1 : ZMod n) = n :=

Mathlib/Probability/ProbabilityMassFunction/Binomial.lean

Lines changed: 5 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -43,9 +43,12 @@ theorem binomial_apply_zero (p : ℝ≥0∞) (h : p ≤ 1) (n : ℕ) :
4343
simp [binomial_apply]
4444

4545
@[simp]
46+
theorem binomial_apply_last (p : ℝ≥0∞) (h : p ≤ 1) (n : ℕ) :
47+
binomial p h n (.last n) = p^n := by
48+
simp [binomial_apply]
49+
4650
theorem binomial_apply_self (p : ℝ≥0∞) (h : p ≤ 1) (n : ℕ) :
47-
binomial p h n n = p^n := by
48-
simp [binomial_apply, Nat.mod_eq_of_lt]
51+
binomial p h n n = p^n := by simp
4952

5053
/-- The binomial distribution on one coin is the bernoully distribution. -/
5154
theorem binomial_one_eq_bernoulli (p : ℝ≥0∞) (h : p ≤ 1) :

0 commit comments

Comments
 (0)