Skip to content

Commit bca7ccf

Browse files
authored
Merge pull request #1648 from patrick-nicodemus/monoidal-cat
2 parents 7079beb + b0cacec commit bca7ccf

File tree

2 files changed

+112
-0
lines changed

2 files changed

+112
-0
lines changed

theories/Basics/Utf8.v

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -57,6 +57,7 @@ Reserved Notation "x ≠ y" (at level 70).
5757
Reserved Notation "x ⇸ y" (at level 99, right associativity, y at level 200).
5858
Reserved Notation "x ↠ y" (at level 99, right associativity, y at level 200).
5959
Reserved Notation "x ↪ y" (at level 99, right associativity, y at level 200).
60+
Reserved Notation "A ⊗ B" (at level 45, left associativity).
6061
(* Reserved Notation "∀ x .. y , P" (at level 200, x binder, y binder, right associativity). *)
6162
Reserved Notation "x ∨ y" (at level 85, right associativity).
6263
(* Reserved Notation "x ⊔ y" (at level 85, right associativity). *)
Lines changed: 111 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,111 @@
1+
Require Import Basics Basics.Utf8 Basics.Tactics.
2+
Require Import implementations.list.
3+
Require Import Category.Core Category.Prod Category.Morphisms.
4+
Require Import NatCategory.
5+
Require Import Functor.Core Functor.Identity Functor.Composition.Core Functor.Prod.Core
6+
Functor.Utf8.
7+
Require Import NaturalTransformation.Core NaturalTransformation.Isomorphisms NaturalTransformation.Identity NaturalTransformation.Prod.
8+
Require Import NaturalTransformation.Composition.Core.
9+
Require Import FunctorCategory.Core FunctorCategory.Morphisms.
10+
Require Import ProductLaws.
11+
Require Import Cat.Core.
12+
13+
Set Universe Polymorphism.
14+
Set Implicit Arguments.
15+
Generalizable All Variables.
16+
Set Asymmetric Patterns.
17+
18+
Section MonoidalStructure.
19+
Context `{Funext}.
20+
21+
Local Notation "x --> y" := (morphism _ x y).
22+
23+
Section MonoidalCategoryConcepts.
24+
Variable C : PreCategory.
25+
Variable tensor : ((C * C) -> C)%category.
26+
Variable I : C.
27+
28+
Local Notation "A ⊗ B" := (tensor (Datatypes.pair A B)).
29+
30+
Local Open Scope functor_scope.
31+
Definition right_assoc := (tensor ∘ (Functor.Prod.pair 1 tensor) )%functor.
32+
Definition left_assoc := tensor ∘
33+
(Functor.Prod.pair tensor 1) ∘
34+
(Associativity.functor _ _ _).
35+
36+
Definition associator := NaturalIsomorphism right_assoc left_assoc.
37+
(* Orientation (A ⊗ B) ⊗ C -> A ⊗ (B ⊗ C) *)
38+
Definition pretensor (A : C) := Core.induced_snd tensor A.
39+
Definition I_pretensor := pretensor I.
40+
Definition posttensor (A : C) := Core.induced_fst tensor A.
41+
Definition I_posttensor := posttensor I.
42+
Definition left_unitor := NaturalIsomorphism I_pretensor 1.
43+
Definition right_unitor := NaturalIsomorphism I_posttensor 1.
44+
45+
Close Scope functor_scope.
46+
47+
Variable alpha : associator.
48+
Variable lambda : left_unitor.
49+
Variable rho : right_unitor.
50+
Notation alpha_nat_trans := ((@morphism_isomorphic
51+
(C * (C * C) -> C)%category right_assoc left_assoc) alpha).
52+
Notation lambda_nat_trans := ((@morphism_isomorphic _ _ _) lambda).
53+
Notation rho_nat_trans := ((@morphism_isomorphic _ _ _) rho).
54+
55+
Section coherence_laws.
56+
Variable a b c d : C.
57+
Local Definition P1 : (a ⊗ (b ⊗ (c ⊗ d))) --> (a ⊗ ((b ⊗ c) ⊗ d)).
58+
Proof.
59+
apply (morphism_of tensor); split; simpl.
60+
- exact (Core.identity a).
61+
- exact (alpha_nat_trans (b, (c, d))).
62+
Defined.
63+
64+
Local Definition P2 : a ⊗ ((b ⊗ c) ⊗ d) --> (a ⊗ (b ⊗ c)) ⊗ d
65+
:= alpha_nat_trans (a, (b ⊗ c, d)).
66+
Local Definition P3 : (a ⊗ (b ⊗ c)) ⊗ d --> ((a ⊗ b) ⊗ c ) ⊗ d.
67+
Proof.
68+
apply (morphism_of tensor); split; simpl.
69+
- exact (alpha_nat_trans (a,_)).
70+
- exact (Core.identity d).
71+
Defined.
72+
Local Definition P4 : a ⊗ (b ⊗ (c ⊗ d)) --> (a ⊗ b) ⊗ (c ⊗ d)
73+
:= alpha_nat_trans (a, (b, (c ⊗ d))).
74+
Local Definition P5 : (a ⊗ b) ⊗ (c ⊗ d) --> ((a ⊗ b) ⊗ c ) ⊗ d
75+
:= alpha_nat_trans (a ⊗ b,(c, d)).
76+
77+
Local Open Scope morphism_scope.
78+
Definition pentagon_eq := P3 o P2 o P1 = P5 o P4.
79+
Close Scope morphism_scope.
80+
81+
Local Definition Q1 : (a ⊗ (I ⊗ b)) --> a ⊗ b.
82+
Proof.
83+
apply (morphism_of tensor); split; simpl.
84+
- exact (Core.identity a).
85+
- exact (lambda_nat_trans _).
86+
Defined.
87+
88+
Local Definition Q2 : (a ⊗ (I ⊗ b)) --> a ⊗ b.
89+
Proof.
90+
refine (@Category.Core.compose _ _ ((a ⊗ I) ⊗ b) _ _ _).
91+
- apply (morphism_of tensor); split; simpl.
92+
+ exact (rho_nat_trans a).
93+
+ exact (Core.identity b).
94+
- exact (alpha_nat_trans (a,(I,b))).
95+
Defined.
96+
Definition triangle_eq := Q1 = Q2.
97+
End coherence_laws.
98+
End MonoidalCategoryConcepts.
99+
100+
Class MonoidalStructure (C : PreCategory) :=
101+
Build_MonoidalStructure {
102+
tensor : (C * C -> C)%category;
103+
I : C;
104+
alpha : associator tensor;
105+
lambda : left_unitor tensor I;
106+
rho : right_unitor tensor I;
107+
pentagon_eq_holds : forall a b c d : C, pentagon_eq alpha a b c d;
108+
triangle_eq_holds : forall a b : C, triangle_eq alpha lambda rho a b;
109+
}.
110+
111+
End MonoidalStructure.

0 commit comments

Comments
 (0)