Skip to content

Commit

Permalink
fix Data.Int.Basic: fix zsmul on (#1217)
Browse files Browse the repository at this point in the history
This way two instances for `m • n`, `m n : ℤ`, are defeq. Also use defeq to golf some proofs.
  • Loading branch information
urkud committed Dec 26, 2022
1 parent be15128 commit e8ce8bc
Show file tree
Hide file tree
Showing 2 changed files with 23 additions and 23 deletions.
34 changes: 15 additions & 19 deletions Mathlib/Algebra/GroupPower/Lemmas.lean
Expand Up @@ -1218,33 +1218,29 @@ open Multiplicative

-- Porting note: the proof became a little roundabout while porting.
@[simp]
theorem Nat.to_add_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := by
induction' b with b ihs
· erw [_root_.pow_zero, toAdd_one, mul_zero]
· simp [*, _root_.pow_succ, add_comm, Nat.mul_succ]
#align nat.to_add_pow Nat.to_add_pow
theorem Nat.toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b :=
mul_comm _ _
#align nat.to_add_pow Nat.toAdd_pow

@[simp]
theorem Nat.of_add_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b :=
(Nat.to_add_pow _ _).symm
#align nat.of_add_mul Nat.of_add_mul
theorem Nat.ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b :=
(Nat.toAdd_pow _ _).symm
#align nat.of_add_mul Nat.ofAdd_mul

@[simp]
theorem Int.to_add_pow (a : Multiplicative ℤ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := by
induction b <;> simp [*, mul_add, pow_succ, add_comm]
#align int.to_add_pow Int.to_add_pow
theorem Int.toAdd_pow (a : Multiplicative ℤ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b :=
mul_comm _ _
#align int.to_add_pow Int.toAdd_pow

@[simp]
theorem Int.to_add_zpow (a : Multiplicative ℤ) (b : ℤ) : toAdd (a ^ b) = toAdd a * b :=
Int.induction_on b (by simp) (by simp (config := { contextual := true }) [zpow_add, mul_add])
(by
simp (config := { contextual := true }) [zpow_add, mul_add, sub_eq_add_neg])
#align int.to_add_zpow Int.to_add_zpow
theorem Int.toAdd_zpow (a : Multiplicative ℤ) (b : ℤ) : toAdd (a ^ b) = toAdd a * b :=
mul_comm _ _
#align int.to_add_zpow Int.toAdd_zpow

@[simp]
theorem Int.of_add_mul (a b : ℤ) : ofAdd (a * b) = ofAdd a ^ b :=
(Int.to_add_zpow _ _).symm
#align int.of_add_mul Int.of_add_mul
theorem Int.ofAdd_mul (a b : ℤ) : ofAdd (a * b) = ofAdd a ^ b :=
(Int.toAdd_zpow _ _).symm
#align int.of_add_mul Int.ofAdd_mul

end Multiplicative

Expand Down
12 changes: 8 additions & 4 deletions Mathlib/Data/Int/Basic.lean
Expand Up @@ -48,10 +48,14 @@ instance : CommRing ℤ where
add_left_neg := Int.add_left_neg
nsmul := (·*·)
nsmul_zero := Int.zero_mul
nsmul_succ n x := by
show ofNat (Nat.succ n) * x = x + ofNat n * x
simp only [ofNat_eq_coe]
rw [Int.ofNat_succ, Int.add_mul, Int.add_comm, Int.one_mul]
nsmul_succ n x :=
show (n + 1 : ℤ) * x = x + n * x
by rw [Int.add_mul, Int.add_comm, Int.one_mul]
zsmul := (·*·)
zsmul_zero' := Int.zero_mul
zsmul_succ' m n := by
simp only [ofNat_eq_coe, ofNat_succ, Int.add_mul, Int.add_comm, Int.one_mul]
zsmul_neg' m n := by simp only [negSucc_coe, ofNat_succ, Int.neg_mul]
sub_eq_add_neg _ _ := Int.sub_eq_add_neg
natCast := (·)
natCast_zero := rfl
Expand Down

0 comments on commit e8ce8bc

Please sign in to comment.