Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: sufficient condition for StructuredArrow.pre to be an equivalence #6248

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
36 changes: 36 additions & 0 deletions Mathlib/CategoryTheory/Comma.lean
Original file line number Diff line number Diff line change
Expand Up @@ -300,6 +300,24 @@ def preLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) : Comma (F ⋙ L) R ⥤ Co
w := by simpa using f.w }
#align category_theory.comma.pre_left CategoryTheory.Comma.preLeft

instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [Faithful F] : Faithful (preLeft F L R) where
map_injective {X Y} f g h := hom_ext _ _ (F.map_injective (congrArg CommaMorphism.left h))
(by apply congrArg CommaMorphism.right h)

instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [Full F] : Full (preLeft F L R) where
preimage {X Y} f := CommaMorphism.mk (F.preimage f.left) f.right (by simpa using f.w)

instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [EssSurj F] : EssSurj (preLeft F L R) where
mem_essImage Y :=
⟨Comma.mk (F.objPreimage Y.left) Y.right ((L.mapIso (F.objObjPreimageIso _)).hom ≫ Y.hom),
⟨isoMk (F.objObjPreimageIso _) (Iso.refl _) (by simp)⟩⟩

/-- If `F` is an equivalence, then so is `preLeft F L R`. -/
noncomputable def isEquivalencePreLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [IsEquivalence F] :
IsEquivalence (preLeft F L R) :=
have := Equivalence.essSurj_of_equivalence F
Equivalence.ofFullyFaithfullyEssSurj _

/-- The functor `(F ⋙ L, R) ⥤ (L, R)` -/
@[simps]
def preRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) : Comma L (F ⋙ R) ⥤ Comma L R where
Expand All @@ -312,6 +330,24 @@ def preRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) : Comma L (F ⋙ R) ⥤ C
right := F.map f.right }
#align category_theory.comma.pre_right CategoryTheory.Comma.preRight

instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [Faithful F] : Faithful (preRight L F R) where
map_injective {X Y } f g h := hom_ext _ _ (by apply congrArg CommaMorphism.left h)
(F.map_injective (congrArg CommaMorphism.right h))

instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [Full F] : Full (preRight L F R) where
preimage {X Y} f := CommaMorphism.mk f.left (F.preimage f.right) (by simpa using f.w)

instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [EssSurj F] : EssSurj (preRight L F R) where
mem_essImage Y :=
⟨Comma.mk Y.left (F.objPreimage Y.right) (Y.hom ≫ (R.mapIso (F.objObjPreimageIso _)).inv),
⟨isoMk (Iso.refl _) (F.objObjPreimageIso _) (by simp [← R.map_comp])⟩⟩

/-- If `F` is an equivalence, then so is `preRight L F R`. -/
noncomputable def isEquivalencePreRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [IsEquivalence F] :
IsEquivalence (preRight L F R) :=
have := Equivalence.essSurj_of_equivalence F
Equivalence.ofFullyFaithfullyEssSurj _

/-- The functor `(L, R) ⥤ (L ⋙ F, R ⋙ F)` -/
@[simps]
def post (L : A ⥤ T) (R : B ⥤ T) (F : T ⥤ C) : Comma L R ⥤ Comma (L ⋙ F) (R ⋙ F) where
Expand Down
28 changes: 28 additions & 0 deletions Mathlib/CategoryTheory/StructuredArrow.lean
Original file line number Diff line number Diff line change
Expand Up @@ -244,6 +244,20 @@ def pre (S : D) (F : B ⥤ C) (G : C ⥤ D) : StructuredArrow S (F ⋙ G) ⥤ St
Comma.preRight _ F G
#align category_theory.structured_arrow.pre CategoryTheory.StructuredArrow.pre

instance (S : D) (F : B ⥤ C) (G : C ⥤ D) [Faithful F] : Faithful (pre S F G) :=
show Faithful (Comma.preRight _ _ _) from inferInstance

instance (S : D) (F : B ⥤ C) (G : C ⥤ D) [Full F] : Full (pre S F G) :=
show Full (Comma.preRight _ _ _) from inferInstance

instance (S : D) (F : B ⥤ C) (G : C ⥤ D) [EssSurj F] : EssSurj (pre S F G) :=
show EssSurj (Comma.preRight _ _ _) from inferInstance

/-- If `F` is an equivalence, then so is the functor `(S, F ⋙ G) ⥤ (S, G)`. -/
noncomputable def isEquivalencePre (S : D) (F : B ⥤ C) (G : C ⥤ D) [IsEquivalence F] :
IsEquivalence (pre S F G) :=
Comma.isEquivalencePreRight _ _ _

/-- The functor `(S, F) ⥤ (G(S), F ⋙ G)`. -/
@[simps]
def post (S : C) (F : B ⥤ C) (G : C ⥤ D) : StructuredArrow S F ⥤ StructuredArrow (G.obj S) (F ⋙ G)
Expand Down Expand Up @@ -468,6 +482,20 @@ def pre (F : B ⥤ C) (G : C ⥤ D) (S : D) : CostructuredArrow (F ⋙ G) S ⥤
Comma.preLeft F G _
#align category_theory.costructured_arrow.pre CategoryTheory.CostructuredArrow.pre

instance (F : B ⥤ C) (G : C ⥤ D) (S : D) [Faithful F] : Faithful (pre F G S) :=
show Faithful (Comma.preLeft _ _ _) from inferInstance

instance (F : B ⥤ C) (G : C ⥤ D) (S : D) [Full F] : Full (pre F G S) :=
show Full (Comma.preLeft _ _ _) from inferInstance

instance (F : B ⥤ C) (G : C ⥤ D) (S : D) [EssSurj F] : EssSurj (pre F G S) :=
show EssSurj (Comma.preLeft _ _ _) from inferInstance

/-- If `F` is an equivalence, then so is the functor `(F ⋙ G, S) ⥤ (G, S)`. -/
noncomputable def isEquivalencePre (F : B ⥤ C) (G : C ⥤ D) (S : D) [IsEquivalence F] :
IsEquivalence (pre F G S) :=
Comma.isEquivalencePreLeft _ _ _

/-- The functor `(F, S) ⥤ (F ⋙ G, G(S))`. -/
@[simps]
def post (F : B ⥤ C) (G : C ⥤ D) (S : C) :
Expand Down
Loading