Skip to content

Commit 13ae635

Browse files
committed
feat(CategoryTheory/Monoidal/Mon_): add simp lemmas (#10905)
1 parent 4dd3701 commit 13ae635

File tree

1 file changed

+34
-8
lines changed

1 file changed

+34
-8
lines changed

Mathlib/CategoryTheory/Monoidal/Mon_.lean

Lines changed: 34 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -469,6 +469,7 @@ theorem mul_rightUnitor {M : Mon_ C} :
469469
simp only [Category.assoc, Category.id_comp]
470470
#align Mon_.mul_right_unitor Mon_.mul_rightUnitor
471471

472+
@[simps tensorObj_X tensorHom_hom]
472473
instance monMonoidalStruct : MonoidalCategoryStruct (Mon_ C) :=
473474
let tensorObj (M N : Mon_ C) : Mon_ C :=
474475
{ X := M.X ⊗ N.X
@@ -497,14 +498,39 @@ instance monMonoidalStruct : MonoidalCategoryStruct (Mon_ C) :=
497498
leftUnitor := fun M ↦ isoOfIso (λ_ M.X) one_leftUnitor mul_leftUnitor
498499
rightUnitor := fun M ↦ isoOfIso (ρ_ M.X) one_rightUnitor mul_rightUnitor }
499500

500-
instance monMonoidal : MonoidalCategory (Mon_ C) := .ofTensorHom
501-
(tensor_id := by intros; ext; apply tensor_id)
502-
(tensor_comp := by intros; ext; apply tensor_comp)
503-
(associator_naturality := by intros; ext; dsimp; apply associator_naturality)
504-
(leftUnitor_naturality := by intros; ext; dsimp; apply leftUnitor_naturality)
505-
(rightUnitor_naturality := by intros; ext; dsimp; apply rightUnitor_naturality)
506-
(pentagon := by intros; ext; dsimp; apply pentagon)
507-
(triangle := by intros; ext; dsimp; apply triangle)
501+
@[simp]
502+
theorem tensorUnit_X : (𝟙_ (Mon_ C)).X = 𝟙_ C := rfl
503+
504+
@[simp]
505+
theorem whiskerLeft_hom {X Y : Mon_ C} (f : X ⟶ Y) (Z : Mon_ C) :
506+
(f ▷ Z).hom = f.hom ▷ Z.X := by
507+
rw [← tensorHom_id]; rfl
508+
509+
@[simp]
510+
theorem whiskerRight_hom (X : Mon_ C) {Y Z : Mon_ C} (f : Y ⟶ Z) :
511+
(X ◁ f).hom = X.X ◁ f.hom := by
512+
rw [← id_tensorHom]; rfl
513+
514+
@[simp]
515+
theorem leftUnitor_hom_hom (X : Mon_ C) : (λ_ X).hom.hom = (λ_ X.X).hom := rfl
516+
517+
@[simp]
518+
theorem leftUnitor_inv_hom (X : Mon_ C) : (λ_ X).inv.hom = (λ_ X.X).inv := rfl
519+
520+
@[simp]
521+
theorem rightUnitor_hom_hom (X : Mon_ C) : (ρ_ X).hom.hom = (ρ_ X.X).hom := rfl
522+
523+
@[simp]
524+
theorem rightUnitor_inv_hom (X : Mon_ C) : (ρ_ X).inv.hom = (ρ_ X.X).inv := rfl
525+
526+
@[simp]
527+
theorem associator_hom_hom (X Y Z : Mon_ C) : (α_ X Y Z).hom.hom = (α_ X.X Y.X Z.X).hom := rfl
528+
529+
@[simp]
530+
theorem associator_inv_hom (X Y Z : Mon_ C) : (α_ X Y Z).inv.hom = (α_ X.X Y.X Z.X).inv := rfl
531+
532+
instance monMonoidal : MonoidalCategory (Mon_ C) where
533+
tensorHom_def := by intros; ext; simp [tensorHom_def]
508534
#align Mon_.Mon_monoidal Mon_.monMonoidal
509535

510536
end Mon_

0 commit comments

Comments
 (0)