Morphisms of quivers #
structure
Prefunctor
(V : Type u₁)
[Quiver V]
(W : Type u₂)
[Quiver W]
:
Sort (max (max (max (u₁ + 1) (u₂ + 1)) v₁) v₂)
A morphism of quivers. As we will later have categorical functors extend this structure,
we call it a Prefunctor
.
- obj : V → W
The action of a (pre)functor on vertices/objects.
The action of a (pre)functor on edges/arrows/morphisms.
Instances For
theorem
Prefunctor.ext'
{V W : Type u}
[Quiver V]
[Quiver W]
{F G : V ⥤q W}
(h_obj : ∀ (X : V), F.obj X = G.obj X)
(h_map : ∀ (X Y : V) (f : X ⟶ Y), F.map f = Quiver.homOfEq (G.map f) ⋯ ⋯)
:
F = G
This may be a more useful form of Prefunctor.ext
.
Equations
- Prefunctor.instInhabited V = { default := 𝟭q V }
Notation for a prefunctor between quivers.
Equations
- Prefunctor.«term_⥤q_» = Lean.ParserDescr.trailingNode `Prefunctor.«term_⥤q_» 50 50 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⥤q ") (Lean.ParserDescr.cat `term 51))
Instances For
Notation for composition of prefunctors.
Equations
- Prefunctor.«term_⋙q_» = Lean.ParserDescr.trailingNode `Prefunctor.«term_⋙q_» 60 60 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⋙q ") (Lean.ParserDescr.cat `term 61))
Instances For
Notation for the identity prefunctor on a quiver.
Equations
- Prefunctor.«term𝟭q» = Lean.ParserDescr.node `Prefunctor.«term𝟭q» 1024 (Lean.ParserDescr.symbol "𝟭q")
Instances For
@[simp]
theorem
Prefunctor.homOfEq_map
{U : Type u_1}
{V : Type u_2}
[Quiver U]
[Quiver V]
(F : U ⥤q V)
{X Y : U}
(f : X ⟶ Y)
{X' Y' : U}
(hX : X = X')
(hY : Y = Y')
:
F.map (Quiver.homOfEq f hX hY) = Quiver.homOfEq (F.map f) ⋯ ⋯
Prefunctors commute with homOfEq
.