Skip to content

Commit

Permalink
perf(LinearAlgebra.TensorProduct.Basic): add TensorProduct.addMonoid (
Browse files Browse the repository at this point in the history
#11505)

We define `TensorProduct.addCommMonoid` in terms of `TensorProduct.addMonoid` to reduce unfolding.
  • Loading branch information
Amelia Livingston authored and dagurtomas committed Mar 22, 2024
1 parent 4bc8196 commit 8a0b662
Showing 1 changed file with 11 additions and 8 deletions.
19 changes: 11 additions & 8 deletions Mathlib/LinearAlgebra/TensorProduct/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -333,15 +333,18 @@ protected theorem add_smul (r s : R'') (x : M ⊗[R] N) : (r + s) • x = r •
rw [ihx, ihy, add_add_add_comm]
#align tensor_product.add_smul TensorProduct.add_smul

instance addMonoid : AddMonoid (M ⊗[R] N) :=
{ TensorProduct.addZeroClass _ _ with
toAddSemigroup := TensorProduct.addSemigroup _ _
toZero := (TensorProduct.addZeroClass _ _).toZero
nsmul := fun n v => n • v
nsmul_zero := by simp [TensorProduct.zero_smul]
nsmul_succ := by simp only [TensorProduct.one_smul, TensorProduct.add_smul, add_comm,
forall_const] }

instance addCommMonoid : AddCommMonoid (M ⊗[R] N) :=
{ TensorProduct.addCommSemigroup _ _,
TensorProduct.addZeroClass _ _ with
toAddSemigroup := TensorProduct.addSemigroup _ _
toZero := (TensorProduct.addZeroClass _ _).toZero
nsmul := fun n v => n • v
nsmul_zero := by simp [TensorProduct.zero_smul]
nsmul_succ := by simp only [TensorProduct.one_smul, TensorProduct.add_smul, add_comm,
forall_const] }
{ TensorProduct.addCommSemigroup _ _ with
toAddMonoid := TensorProduct.addMonoid }

instance leftDistribMulAction : DistribMulAction R' (M ⊗[R] N) :=
have : ∀ (r : R') (m : M) (n : N), r • m ⊗ₜ[R] n = (r • m) ⊗ₜ n := fun _ _ _ => rfl
Expand Down

0 comments on commit 8a0b662

Please sign in to comment.