From 01c7d9eb2cc39f51491c395482a7bf80340eb3bf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Sun, 1 Jan 2023 17:57:13 +0100 Subject: [PATCH 01/66] Initial file copy from mathport --- Mathlib/CategoryTheory/Equivalence.lean | 815 ++++++++++++++++++++++++ 1 file changed, 815 insertions(+) create mode 100644 Mathlib/CategoryTheory/Equivalence.lean diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean new file mode 100644 index 0000000000000..ec6879fbcf67e --- /dev/null +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -0,0 +1,815 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Tim Baumann, Stephen Morgan, Scott Morrison, Floris van Doorn + +! This file was ported from Lean 3 source module category_theory.equivalence +! leanprover-community/mathlib commit 9aba7801eeecebb61f58a5763c2b6dd1b47dc6ef +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Functor.FullyFaithful +import Mathbin.CategoryTheory.FullSubcategory +import Mathbin.CategoryTheory.Whiskering +import Mathbin.CategoryTheory.EssentialImage +import Mathbin.Tactic.Slice + +/-! +# Equivalence of categories + +An equivalence of categories `C` and `D` is a pair of functors `F : C ⥤ D` and `G : D ⥤ C` such +that `η : 𝟭 C ≅ F ⋙ G` and `ε : G ⋙ F ≅ 𝟭 D`. In many situations, equivalences are a better +notion of "sameness" of categories than the stricter isomorphims of categories. + +Recall that one way to express that two functors `F : C ⥤ D` and `G : D ⥤ C` are adjoint is using +two natural transformations `η : 𝟭 C ⟶ F ⋙ G` and `ε : G ⋙ F ⟶ 𝟭 D`, called the unit and the +counit, such that the compositions `F ⟶ FGF ⟶ F` and `G ⟶ GFG ⟶ G` are the identity. Unfortunately, +it is not the case that the natural isomorphisms `η` and `ε` in the definition of an equivalence +automatically give an adjunction. However, it is true that +* if one of the two compositions is the identity, then so is the other, and +* given an equivalence of categories, it is always possible to refine `η` in such a way that the + identities are satisfied. + +For this reason, in mathlib we define an equivalence to be a "half-adjoint equivalence", which is +a tuple `(F, G, η, ε)` as in the first paragraph such that the composite `F ⟶ FGF ⟶ F` is the +identity. By the remark above, this already implies that the tuple is an "adjoint equivalence", +i.e., that the composite `G ⟶ GFG ⟶ G` is also the identity. + +We also define essentially surjective functors and show that a functor is an equivalence if and only +if it is full, faithful and essentially surjective. + +## Main definitions + +* `equivalence`: bundled (half-)adjoint equivalences of categories +* `is_equivalence`: type class on a functor `F` containing the data of the inverse `G` as well as + the natural isomorphisms `η` and `ε`. +* `ess_surj`: type class on a functor `F` containing the data of the preimages and the isomorphisms + `F.obj (preimage d) ≅ d`. + +## Main results + +* `equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence +* `is_equivalence.equiv_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence +iff `G` is. +* `equivalence.of_fully_faithfully_ess_surj`: a fully faithful essentially surjective functor is an + equivalence. + +## Notations + +We write `C ≌ D` (`\backcong`, not to be confused with `≅`/`\cong`) for a bundled equivalence. + +-/ + + +namespace CategoryTheory + +open CategoryTheory.Functor NatIso Category + +-- declare the `v`'s first; see `category_theory.category` for an explanation +universe v₁ v₂ v₃ u₁ u₂ u₃ + +/-- We define an equivalence as a (half)-adjoint equivalence, a pair of functors with + a unit and counit which are natural isomorphisms and the triangle law `Fη ≫ εF = 1`, or in other + words the composite `F ⟶ FGF ⟶ F` is the identity. + + In `unit_inverse_comp`, we show that this is actually an adjoint equivalence, i.e., that the + composite `G ⟶ GFG ⟶ G` is also the identity. + + The triangle equation is written as a family of equalities between morphisms, it is more + complicated if we write it as an equality of natural transformations, because then we would have + to insert natural transformations like `F ⟶ F1`. + +See +-/ +structure Equivalence (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] where mk' :: + Functor : C ⥤ D + inverse : D ⥤ C + unitIso : 𝟭 C ≅ Functor ⋙ inverse + counitIso : inverse ⋙ Functor ≅ 𝟭 D + functor_unit_iso_comp' : + ∀ X : C, + Functor.map ((unit_iso.Hom : 𝟭 C ⟶ Functor ⋙ inverse).app X) ≫ + counit_iso.Hom.app (Functor.obj X) = + 𝟙 (Functor.obj X) := by + obviously +#align category_theory.equivalence CategoryTheory.Equivalence + +restate_axiom equivalence.functor_unit_iso_comp' + +-- mathport name: «expr ≌ » +infixr:10 " ≌ " => Equivalence + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +namespace Equivalence + +/-- The unit of an equivalence of categories. -/ +abbrev unit (e : C ≌ D) : 𝟭 C ⟶ e.Functor ⋙ e.inverse := + e.unitIso.Hom +#align category_theory.equivalence.unit CategoryTheory.Equivalence.unit + +/-- The counit of an equivalence of categories. -/ +abbrev counit (e : C ≌ D) : e.inverse ⋙ e.Functor ⟶ 𝟭 D := + e.counitIso.Hom +#align category_theory.equivalence.counit CategoryTheory.Equivalence.counit + +/-- The inverse of the unit of an equivalence of categories. -/ +abbrev unitInv (e : C ≌ D) : e.Functor ⋙ e.inverse ⟶ 𝟭 C := + e.unitIso.inv +#align category_theory.equivalence.unit_inv CategoryTheory.Equivalence.unitInv + +/-- The inverse of the counit of an equivalence of categories. -/ +abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.Functor := + e.counitIso.inv +#align category_theory.equivalence.counit_inv CategoryTheory.Equivalence.counitInv + +/- While these abbreviations are convenient, they also cause some trouble, +preventing structure projections from unfolding. -/ +@[simp] +theorem equivalence_mk'_unit (functor inverse unit_iso counit_iso f) : + (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).Unit = unit_iso.Hom := + rfl +#align + category_theory.equivalence.equivalence_mk'_unit CategoryTheory.Equivalence.equivalence_mk'_unit + +@[simp] +theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : + (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.Hom := + rfl +#align + category_theory.equivalence.equivalence_mk'_counit CategoryTheory.Equivalence.equivalence_mk'_counit + +@[simp] +theorem equivalence_mk'_unit_inv (functor inverse unit_iso counit_iso f) : + (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := + rfl +#align + category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unit_inv + +@[simp] +theorem equivalence_mk'_counit_inv (functor inverse unit_iso counit_iso f) : + (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := + rfl +#align + category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counit_inv + +@[simp] +theorem functor_unit_comp (e : C ≌ D) (X : C) : + e.Functor.map (e.Unit.app X) ≫ e.counit.app (e.Functor.obj X) = 𝟙 (e.Functor.obj X) := + e.functor_unit_iso_comp X +#align category_theory.equivalence.functor_unit_comp CategoryTheory.Equivalence.functor_unit_comp + +@[simp] +theorem counit_inv_functor_comp (e : C ≌ D) (X : C) : + e.counitInv.app (e.Functor.obj X) ≫ e.Functor.map (e.unitInv.app X) = 𝟙 (e.Functor.obj X) := + by + erw [iso.inv_eq_inv (e.functor.map_iso (e.unit_iso.app X) ≪≫ e.counit_iso.app (e.functor.obj X)) + (iso.refl _)] + exact e.functor_unit_comp X +#align + category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counit_inv_functor_comp + +theorem counit_inv_app_functor (e : C ≌ D) (X : C) : + e.counitInv.app (e.Functor.obj X) = e.Functor.map (e.Unit.app X) := + by + symm + erw [← iso.comp_hom_eq_id (e.counit_iso.app _), functor_unit_comp] + rfl +#align + category_theory.equivalence.counit_inv_app_functor CategoryTheory.Equivalence.counit_inv_app_functor + +theorem counit_app_functor (e : C ≌ D) (X : C) : + e.counit.app (e.Functor.obj X) = e.Functor.map (e.unitInv.app X) := + by + erw [← iso.hom_comp_eq_id (e.functor.map_iso (e.unit_iso.app X)), functor_unit_comp] + rfl +#align category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor + +/-- The other triangle equality. The proof follows the following proof in Globular: + http://globular.science/1905.001 -/ +@[simp] +theorem unit_inverse_comp (e : C ≌ D) (Y : D) : + e.Unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := + by + rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counit_inv_functor_comp, map_comp] + dsimp + rw [← iso.hom_inv_id_assoc (e.unit_iso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] + slice_lhs 2 3 => erw [e.unit.naturality] + slice_lhs 1 2 => erw [e.unit.naturality] + slice_lhs 4 4 => + rw [← iso.hom_inv_id_assoc (e.inverse.map_iso (e.counit_iso.app _)) (e.unit_inv.app _)] + slice_lhs 3 4 => + erw [← map_comp e.inverse, e.counit.naturality] + erw [(e.counit_iso.app _).hom_inv_id, map_id] + erw [id_comp] + slice_lhs 2 3 => erw [← map_comp e.inverse, e.counit_iso.inv.naturality, map_comp] + slice_lhs 3 4 => erw [e.unit_inv.naturality] + slice_lhs 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unit_iso.app _).hom_inv_id, map_id] + erw [id_comp] + slice_lhs 3 4 => erw [← e.unit_inv.naturality] + slice_lhs 2 3 => + erw [← map_comp e.inverse, ← e.counit_iso.inv.naturality, (e.counit_iso.app _).hom_inv_id, + map_id] + erw [id_comp, (e.unit_iso.app _).hom_inv_id]; rfl +#align category_theory.equivalence.unit_inverse_comp CategoryTheory.Equivalence.unit_inverse_comp + +@[simp] +theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : + e.inverse.map (e.counitInv.app Y) ≫ e.unitInv.app (e.inverse.obj Y) = 𝟙 (e.inverse.obj Y) := + by + erw [iso.inv_eq_inv (e.unit_iso.app (e.inverse.obj Y) ≪≫ e.inverse.map_iso (e.counit_iso.app Y)) + (iso.refl _)] + exact e.unit_inverse_comp Y +#align + category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp + +theorem unit_app_inverse (e : C ≌ D) (Y : D) : + e.Unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := + by + erw [← iso.comp_hom_eq_id (e.inverse.map_iso (e.counit_iso.app Y)), unit_inverse_comp] + rfl +#align category_theory.equivalence.unit_app_inverse CategoryTheory.Equivalence.unit_app_inverse + +theorem unit_inv_app_inverse (e : C ≌ D) (Y : D) : + e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := + by + symm + erw [← iso.hom_comp_eq_id (e.unit_iso.app _), unit_inverse_comp] + rfl +#align + category_theory.equivalence.unit_inv_app_inverse CategoryTheory.Equivalence.unit_inv_app_inverse + +@[simp] +theorem fun_inv_map (e : C ≌ D) (X Y : D) (f : X ⟶ Y) : + e.Functor.map (e.inverse.map f) = e.counit.app X ≫ f ≫ e.counitInv.app Y := + (NatIso.naturality_2 e.counitIso f).symm +#align category_theory.equivalence.fun_inv_map CategoryTheory.Equivalence.fun_inv_map + +@[simp] +theorem inv_fun_map (e : C ≌ D) (X Y : C) (f : X ⟶ Y) : + e.inverse.map (e.Functor.map f) = e.unitInv.app X ≫ f ≫ e.Unit.app Y := + (NatIso.naturality_1 e.unitIso f).symm +#align category_theory.equivalence.inv_fun_map CategoryTheory.Equivalence.inv_fun_map + +section + +-- In this section we convert an arbitrary equivalence to a half-adjoint equivalence. +variable {F : C ⥤ D} {G : D ⥤ C} (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) + +/-- If `η : 𝟭 C ≅ F ⋙ G` is part of a (not necessarily half-adjoint) equivalence, we can upgrade it +to a refined natural isomorphism `adjointify_η η : 𝟭 C ≅ F ⋙ G` which exhibits the properties +required for a half-adjoint equivalence. See `equivalence.mk`. -/ +def adjointifyη : 𝟭 C ≅ F ⋙ G := + calc + 𝟭 C ≅ F ⋙ G := η + _ ≅ F ⋙ 𝟭 D ⋙ G := isoWhiskerLeft F (leftUnitor G).symm + _ ≅ F ⋙ (G ⋙ F) ⋙ G := isoWhiskerLeft F (isoWhiskerRight ε.symm G) + _ ≅ F ⋙ G ⋙ F ⋙ G := isoWhiskerLeft F (associator G F G) + _ ≅ (F ⋙ G) ⋙ F ⋙ G := (associator F G (F ⋙ G)).symm + _ ≅ 𝟭 C ⋙ F ⋙ G := isoWhiskerRight η.symm (F ⋙ G) + _ ≅ F ⋙ G := leftUnitor (F ⋙ G) + +#align category_theory.equivalence.adjointify_η CategoryTheory.Equivalence.adjointifyη + +theorem adjointify_η_ε (X : C) : + F.map ((adjointifyη η ε).Hom.app X) ≫ ε.Hom.app (F.obj X) = 𝟙 (F.obj X) := + by + dsimp [adjointify_η]; simp + have := ε.hom.naturality (F.map (η.inv.app X)); dsimp at this; rw [this]; clear this + rw [← assoc _ _ (F.map _)] + have := ε.hom.naturality (ε.inv.app <| F.obj X); dsimp at this; rw [this]; clear this + have := (ε.app <| F.obj X).hom_inv_id; dsimp at this; rw [this]; clear this + rw [id_comp]; have := (F.map_iso <| η.app X).hom_inv_id; dsimp at this; rw [this] +#align category_theory.equivalence.adjointify_η_ε CategoryTheory.Equivalence.adjointify_η_ε + +end + +/-- Every equivalence of categories consisting of functors `F` and `G` such that `F ⋙ G` and + `G ⋙ F` are naturally isomorphic to identity functors can be transformed into a half-adjoint + equivalence without changing `F` or `G`. -/ +protected def mk (F : C ⥤ D) (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : C ≌ D := + ⟨F, G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ +#align category_theory.equivalence.mk CategoryTheory.Equivalence.mk + +/-- Equivalence of categories is reflexive. -/ +@[refl, simps] +def refl : C ≌ C := + ⟨𝟭 C, 𝟭 C, Iso.refl _, Iso.refl _, fun X => Category.id_comp _⟩ +#align category_theory.equivalence.refl CategoryTheory.Equivalence.refl + +instance : Inhabited (C ≌ C) := + ⟨refl⟩ + +/-- Equivalence of categories is symmetric. -/ +@[symm, simps] +def symm (e : C ≌ D) : D ≌ C := + ⟨e.inverse, e.Functor, e.counitIso.symm, e.unitIso.symm, e.inverse_counit_inv_comp⟩ +#align category_theory.equivalence.symm CategoryTheory.Equivalence.symm + +variable {E : Type u₃} [Category.{v₃} E] + +/-- Equivalence of categories is transitive. -/ +@[trans, simps] +def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E + where + Functor := e.Functor ⋙ f.Functor + inverse := f.inverse ⋙ e.inverse + unitIso := by + refine' iso.trans e.unit_iso _ + exact iso_whisker_left e.functor (iso_whisker_right f.unit_iso e.inverse) + counitIso := by + refine' iso.trans _ f.counit_iso + exact iso_whisker_left f.inverse (iso_whisker_right e.counit_iso f.functor) + -- We wouldn't have needed to give this proof if we'd used `equivalence.mk`, + -- but we choose to avoid using that here, for the sake of good structure projection `simp` + -- lemmas. + functor_unit_iso_comp' X := by + dsimp + rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counit_inv_app_functor, fun_inv_map, + iso.inv_hom_id_app_assoc, assoc, iso.inv_hom_id_app, counit_app_functor, ← functor.map_comp] + erw [comp_id, iso.hom_inv_id_app, Functor.map_id] +#align category_theory.equivalence.trans CategoryTheory.Equivalence.trans + +/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic +functor. -/ +def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.Functor ⋙ e.inverse ⋙ F ≅ F := + (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.unitIso.symm F ≪≫ F.leftUnitor +#align category_theory.equivalence.fun_inv_id_assoc CategoryTheory.Equivalence.funInvIdAssoc + +@[simp] +theorem fun_inv_id_assoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : + (funInvIdAssoc e F).Hom.app X = F.map (e.unitInv.app X) := + by + dsimp [fun_inv_id_assoc] + tidy +#align + category_theory.equivalence.fun_inv_id_assoc_hom_app CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app + +@[simp] +theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : + (funInvIdAssoc e F).inv.app X = F.map (e.Unit.app X) := + by + dsimp [fun_inv_id_assoc] + tidy +#align + category_theory.equivalence.fun_inv_id_assoc_inv_app CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app + +/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic +functor. -/ +def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.Functor ⋙ F ≅ F := + (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.counitIso F ≪≫ F.leftUnitor +#align category_theory.equivalence.inv_fun_id_assoc CategoryTheory.Equivalence.invFunIdAssoc + +@[simp] +theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : + (invFunIdAssoc e F).Hom.app X = F.map (e.counit.app X) := + by + dsimp [inv_fun_id_assoc] + tidy +#align + category_theory.equivalence.inv_fun_id_assoc_hom_app CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app + +@[simp] +theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : + (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := + by + dsimp [inv_fun_id_assoc] + tidy +#align + category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app + +/-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ +@[simps Functor inverse unitIso counitIso] +def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := + Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.Functor) + (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by tidy)) + (NatIso.ofComponents (fun F => e.invFunIdAssoc F) (by tidy)) +#align category_theory.equivalence.congr_left CategoryTheory.Equivalence.congrLeft + +/-- If `C` is equivalent to `D`, then `E ⥤ C` is equivalent to `E ⥤ D`. -/ +@[simps Functor inverse unitIso counitIso] +def congrRight (e : C ≌ D) : E ⥤ C ≌ E ⥤ D := + Equivalence.mk ((whiskeringRight _ _ _).obj e.Functor) ((whiskeringRight _ _ _).obj e.inverse) + (NatIso.ofComponents + (fun F => F.rightUnitor.symm ≪≫ isoWhiskerLeft F e.unitIso ≪≫ Functor.associator _ _ _) + (by tidy)) + (NatIso.ofComponents + (fun F => Functor.associator _ _ _ ≪≫ isoWhiskerLeft F e.counitIso ≪≫ F.rightUnitor) + (by tidy)) +#align category_theory.equivalence.congr_right CategoryTheory.Equivalence.congrRight + +section CancellationLemmas + +variable (e : C ≌ D) + +/- We need special forms of `cancel_nat_iso_hom_right(_assoc)` and +`cancel_nat_iso_inv_right(_assoc)` for units and counits, because neither `simp` or `rw` will apply +those lemmas in this setting without providing `e.unit_iso` (or similar) as an explicit argument. +We also provide the lemmas for length four compositions, since they're occasionally useful. +(e.g. in proving that equivalences take monos to monos) -/ +@[simp] +theorem cancel_unit_right {X Y : C} (f f' : X ⟶ Y) : + f ≫ e.Unit.app Y = f' ≫ e.Unit.app Y ↔ f = f' := by simp only [cancel_mono] +#align category_theory.equivalence.cancel_unit_right CategoryTheory.Equivalence.cancel_unit_right + +@[simp] +theorem cancel_unit_inv_right {X Y : C} (f f' : X ⟶ e.inverse.obj (e.Functor.obj Y)) : + f ≫ e.unitInv.app Y = f' ≫ e.unitInv.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_unit_inv_right CategoryTheory.Equivalence.cancel_unit_inv_right + +@[simp] +theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.Functor.obj (e.inverse.obj Y)) : + f ≫ e.counit.app Y = f' ≫ e.counit.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_counit_right CategoryTheory.Equivalence.cancel_counit_right + +@[simp] +theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : + f ≫ e.counitInv.app Y = f' ≫ e.counitInv.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right CategoryTheory.Equivalence.cancel_counit_inv_right + +@[simp] +theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : + f ≫ g ≫ e.Unit.app Y = f' ≫ g' ≫ e.Unit.app Y ↔ f ≫ g = f' ≫ g' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc + +@[simp] +theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') + (g' : X' ⟶ Y) : f ≫ g ≫ e.counitInv.app Y = f' ≫ g' ≫ e.counitInv.app Y ↔ f ≫ g = f' ≫ g' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right_assoc CategoryTheory.Equivalence.cancel_counit_inv_right_assoc + +@[simp] +theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) + (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : + f ≫ g ≫ h ≫ e.Unit.app Z = f' ≫ g' ≫ h' ≫ e.Unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' + +@[simp] +theorem cancel_counit_inv_right_assoc' {W X X' Y Y' Z : D} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) + (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : + f ≫ g ≫ h ≫ e.counitInv.app Z = f' ≫ g' ≫ h' ≫ e.counitInv.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := + by simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right_assoc' CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' + +end CancellationLemmas + +section + +-- There's of course a monoid structure on `C ≌ C`, +-- but let's not encourage using it. +-- The power structure is nevertheless useful. +/-- Natural number powers of an auto-equivalence. Use `(^)` instead. -/ +def powNat (e : C ≌ C) : ℕ → (C ≌ C) + | 0 => Equivalence.refl + | 1 => e + | n + 2 => e.trans (pow_nat (n + 1)) +#align category_theory.equivalence.pow_nat CategoryTheory.Equivalence.powNat + +/-- Powers of an auto-equivalence. Use `(^)` instead. -/ +def pow (e : C ≌ C) : ℤ → (C ≌ C) + | Int.ofNat n => e.powNat n + | Int.negSucc n => e.symm.powNat (n + 1) +#align category_theory.equivalence.pow CategoryTheory.Equivalence.pow + +instance : Pow (C ≌ C) ℤ := + ⟨pow⟩ + +@[simp] +theorem pow_zero (e : C ≌ C) : e ^ (0 : ℤ) = equivalence.refl := + rfl +#align category_theory.equivalence.pow_zero CategoryTheory.Equivalence.pow_zero + +@[simp] +theorem pow_one (e : C ≌ C) : e ^ (1 : ℤ) = e := + rfl +#align category_theory.equivalence.pow_one CategoryTheory.Equivalence.pow_one + +@[simp] +theorem pow_neg_one (e : C ≌ C) : e ^ (-1 : ℤ) = e.symm := + rfl +#align category_theory.equivalence.pow_neg_one CategoryTheory.Equivalence.pow_neg_one + +-- TODO as necessary, add the natural isomorphisms `(e^a).trans e^b ≅ e^(a+b)`. +-- At this point, we haven't even defined the category of equivalences. +end + +end Equivalence + +/-- A functor that is part of a (half) adjoint equivalence -/ +class IsEquivalence (F : C ⥤ D) where mk' :: + inverse : D ⥤ C + unitIso : 𝟭 C ≅ F ⋙ inverse + counitIso : inverse ⋙ F ≅ 𝟭 D + functor_unit_iso_comp' : + ∀ X : C, + F.map ((unit_iso.Hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counit_iso.Hom.app (F.obj X) = + 𝟙 (F.obj X) := by + obviously +#align category_theory.is_equivalence CategoryTheory.IsEquivalence + +restate_axiom is_equivalence.functor_unit_iso_comp' + +attribute [simp, reassoc.1] is_equivalence.functor_unit_iso_comp + +namespace IsEquivalence + +instance ofEquivalence (F : C ≌ D) : IsEquivalence F.Functor := + { F with } +#align category_theory.is_equivalence.of_equivalence CategoryTheory.IsEquivalence.ofEquivalence + +instance ofEquivalenceInverse (F : C ≌ D) : IsEquivalence F.inverse := + IsEquivalence.ofEquivalence F.symm +#align + category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.IsEquivalence.ofEquivalenceInverse + +open Equivalence + +/-- To see that a functor is an equivalence, it suffices to provide an inverse functor `G` such that + `F ⋙ G` and `G ⋙ F` are naturally isomorphic to identity functors. -/ +protected def mk {F : C ⥤ D} (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : IsEquivalence F := + ⟨G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ +#align category_theory.is_equivalence.mk CategoryTheory.IsEquivalence.mk + +end IsEquivalence + +namespace Functor + +/-- Interpret a functor that is an equivalence as an equivalence. -/ +def asEquivalence (F : C ⥤ D) [IsEquivalence F] : C ≌ D := + ⟨F, IsEquivalence.inverse F, IsEquivalence.unitIso, IsEquivalence.counitIso, + IsEquivalence.functor_unit_iso_comp⟩ +#align category_theory.functor.as_equivalence CategoryTheory.Functor.asEquivalence + +instance isEquivalenceRefl : IsEquivalence (𝟭 C) := + IsEquivalence.ofEquivalence Equivalence.refl +#align category_theory.functor.is_equivalence_refl CategoryTheory.Functor.isEquivalenceRefl + +/-- The inverse functor of a functor that is an equivalence. -/ +def inv (F : C ⥤ D) [IsEquivalence F] : D ⥤ C := + IsEquivalence.inverse F +#align category_theory.functor.inv CategoryTheory.Functor.inv + +instance isEquivalenceInv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv := + IsEquivalence.ofEquivalence F.asEquivalence.symm +#align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalenceInv + +@[simp] +theorem as_equivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.Functor = F := + rfl +#align category_theory.functor.as_equivalence_functor CategoryTheory.Functor.as_equivalence_functor + +@[simp] +theorem as_equivalence_inverse (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.inverse = inv F := + rfl +#align category_theory.functor.as_equivalence_inverse CategoryTheory.Functor.as_equivalence_inverse + +@[simp] +theorem as_equivalence_unit {F : C ⥤ D} [h : IsEquivalence F] : + F.asEquivalence.unitIso = @IsEquivalence.unitIso _ _ h := + rfl +#align category_theory.functor.as_equivalence_unit CategoryTheory.Functor.as_equivalence_unit + +@[simp] +theorem as_equivalence_counit {F : C ⥤ D} [IsEquivalence F] : + F.asEquivalence.counitIso = is_equivalence.counit_iso := + rfl +#align category_theory.functor.as_equivalence_counit CategoryTheory.Functor.as_equivalence_counit + +@[simp] +theorem inv_inv (F : C ⥤ D) [IsEquivalence F] : inv (inv F) = F := + rfl +#align category_theory.functor.inv_inv CategoryTheory.Functor.inv_inv + +variable {E : Type u₃} [Category.{v₃} E] + +instance isEquivalenceTrans (F : C ⥤ D) (G : D ⥤ E) [IsEquivalence F] [IsEquivalence G] : + IsEquivalence (F ⋙ G) := + IsEquivalence.ofEquivalence (Equivalence.trans (asEquivalence F) (asEquivalence G)) +#align category_theory.functor.is_equivalence_trans CategoryTheory.Functor.isEquivalenceTrans + +end Functor + +namespace Equivalence + +@[simp] +theorem functor_inv (E : C ≌ D) : E.Functor.inv = E.inverse := + rfl +#align category_theory.equivalence.functor_inv CategoryTheory.Equivalence.functor_inv + +@[simp] +theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.Functor := + rfl +#align category_theory.equivalence.inverse_inv CategoryTheory.Equivalence.inverse_inv + +@[simp] +theorem functor_as_equivalence (E : C ≌ D) : E.Functor.asEquivalence = E := + by + cases E + congr +#align + category_theory.equivalence.functor_as_equivalence CategoryTheory.Equivalence.functor_as_equivalence + +@[simp] +theorem inverse_as_equivalence (E : C ≌ D) : E.inverse.asEquivalence = E.symm := + by + cases E + congr +#align + category_theory.equivalence.inverse_as_equivalence CategoryTheory.Equivalence.inverse_as_equivalence + +end Equivalence + +namespace IsEquivalence + +@[simp] +theorem fun_inv_map (F : C ⥤ D) [IsEquivalence F] (X Y : D) (f : X ⟶ Y) : + F.map (F.inv.map f) = F.asEquivalence.counit.app X ≫ f ≫ F.asEquivalence.counitInv.app Y := + by + erw [nat_iso.naturality_2] + rfl +#align category_theory.is_equivalence.fun_inv_map CategoryTheory.IsEquivalence.fun_inv_map + +@[simp] +theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : + F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.Unit.app Y := + by + erw [nat_iso.naturality_1] + rfl +#align category_theory.is_equivalence.inv_fun_map CategoryTheory.IsEquivalence.inv_fun_map + +/-- When a functor `F` is an equivalence of categories, and `G` is isomorphic to `F`, then +`G` is also an equivalence of categories. -/ +@[simps] +def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G + where + inverse := hF.inverse + unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) + counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso + functor_unit_iso_comp' X := by + dsimp [nat_iso.hcomp] + erw [id_comp, F.map_id, comp_id] + apply (cancel_epi (e.hom.app X)).mp + slice_lhs 1 2 => rw [← e.hom.naturality] + slice_lhs 2 3 => rw [← nat_trans.vcomp_app', e.hom_inv_id] + simp only [nat_trans.id_app, id_comp, comp_id, F.map_comp, assoc] + erw [hF.counit_iso.hom.naturality] + slice_lhs 1 2 => rw [functor_unit_iso_comp] + simp only [functor.id_map, id_comp] +#align category_theory.is_equivalence.of_iso CategoryTheory.IsEquivalence.ofIso + +/-- Compatibility of `of_iso` with the composition of isomorphisms of functors -/ +theorem of_iso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEquivalence F) : + ofIso e' (ofIso e hF) = ofIso (e ≪≫ e') hF := + by + dsimp [of_iso] + congr 1 <;> ext X <;> dsimp [nat_iso.hcomp] + · simp only [id_comp, assoc, functor.map_comp] + · simp only [Functor.map_id, comp_id, id_comp, assoc] +#align category_theory.is_equivalence.of_iso_trans CategoryTheory.IsEquivalence.of_iso_trans + +/-- Compatibility of `of_iso` with identity isomorphisms of functors -/ +theorem of_iso_refl (F : C ⥤ D) (hF : IsEquivalence F) : ofIso (Iso.refl F) hF = hF := + by + rcases hF with ⟨Finv, Funit, Fcounit, Fcomp⟩ + dsimp [of_iso] + congr 1 <;> ext X <;> dsimp [nat_iso.hcomp] + · simp only [comp_id, map_id] + · simp only [id_comp, map_id] +#align category_theory.is_equivalence.of_iso_refl CategoryTheory.IsEquivalence.of_iso_refl + +/-- When `F` and `G` are two isomorphic functors, then `F` is an equivalence iff `G` is. -/ +@[simps] +def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence G + where + toFun := ofIso e + invFun := ofIso e.symm + left_inv hF := by rw [of_iso_trans, iso.self_symm_id, of_iso_refl] + right_inv hF := by rw [of_iso_trans, iso.symm_self_id, of_iso_refl] +#align category_theory.is_equivalence.equiv_of_iso CategoryTheory.IsEquivalence.equivOfIso + +/-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ +@[simp] +def cancelCompRight {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : IsEquivalence G) + (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence F := + ofIso (Functor.associator F G G.inv ≪≫ NatIso.hcomp (Iso.refl F) hG.unitIso.symm ≪≫ rightUnitor F) + (Functor.isEquivalenceTrans (F ⋙ G) G.inv) +#align category_theory.is_equivalence.cancel_comp_right CategoryTheory.IsEquivalence.cancelCompRight + +/-- If `F` and `F ⋙ G` are equivalence of categories, then `G` is also an equivalence. -/ +@[simp] +def cancelCompLeft {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hF : IsEquivalence F) + (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence G := + ofIso + ((Functor.associator F.inv F G).symm ≪≫ NatIso.hcomp hF.counitIso (Iso.refl G) ≪≫ leftUnitor G) + (Functor.isEquivalenceTrans F.inv (F ⋙ G)) +#align category_theory.is_equivalence.cancel_comp_left CategoryTheory.IsEquivalence.cancelCompLeft + +end IsEquivalence + +namespace Equivalence + +/-- An equivalence is essentially surjective. + +See . +-/ +theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := + ⟨fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩ +#align + category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.ess_surj_of_equivalence + +-- see Note [lower instance priority] +/-- An equivalence is faithful. + +See . +-/ +instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F + where map_injective' X Y f g w := + by + have p := congr_arg (@CategoryTheory.Functor.map _ _ _ _ F.inv _ _) w + simpa only [cancel_epi, cancel_mono, is_equivalence.inv_fun_map] using p +#align + category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithful_of_equivalence + +-- see Note [lower instance priority] +/-- An equivalence is full. + +See . +-/ +instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F + where + preimage X Y f := F.asEquivalence.Unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y + witness' X Y f := + F.inv.map_injective <| by + simpa only [is_equivalence.inv_fun_map, assoc, iso.inv_hom_id_app_assoc, + iso.inv_hom_id_app] using comp_id _ +#align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.fullOfEquivalence + +@[simps] +private noncomputable def equivalence_inverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C + where + obj X := F.objPreimage X + map X Y f := F.preimage ((F.objObjPreimageIso X).Hom ≫ f ≫ (F.objObjPreimageIso Y).inv) + map_id' X := by apply F.map_injective; tidy + map_comp' X Y Z f g := by apply F.map_injective <;> simp +#align + category_theory.equivalence.equivalence_inverse category_theory.equivalence.equivalence_inverse + +/-- A functor which is full, faithful, and essentially surjective is an equivalence. + +See . +-/ +noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : + IsEquivalence F := + IsEquivalence.mk (equivalenceInverse F) + (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) + fun X Y f => by + apply F.map_injective + obviously) + (NatIso.ofComponents F.objObjPreimageIso (by tidy)) +#align + category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj + +@[simp] +theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : + e.Functor.map f = e.Functor.map g ↔ f = g := + ⟨fun h => e.Functor.map_injective h, fun h => h ▸ rfl⟩ +#align + category_theory.equivalence.functor_map_inj_iff CategoryTheory.Equivalence.functor_map_inj_iff + +@[simp] +theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : + e.inverse.map f = e.inverse.map g ↔ f = g := + functor_map_inj_iff e.symm f g +#align + category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff + +instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) + where mem_ess_image Y := ⟨e.symm Y, by simp⟩ +#align + category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor + +noncomputable instance inducedFunctorOfEquiv {C' : Type _} (e : C' ≃ D) : + IsEquivalence (inducedFunctor e) := + Equivalence.ofFullyFaithfullyEssSurj _ +#align + category_theory.equivalence.induced_functor_of_equiv CategoryTheory.Equivalence.inducedFunctorOfEquiv + +noncomputable instance fullyFaithfulToEssImage (F : C ⥤ D) [Full F] [Faithful F] : + IsEquivalence F.toEssImage := + ofFullyFaithfullyEssSurj F.toEssImage +#align + category_theory.equivalence.fully_faithful_to_ess_image CategoryTheory.Equivalence.fullyFaithfulToEssImage + +end Equivalence + +end CategoryTheory + From bc9467dcbd72b32db46289f388ecfe931d10aeca Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Sun, 1 Jan 2023 17:57:13 +0100 Subject: [PATCH 02/66] Mathbin -> Mathlib; add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Equivalence.lean | 10 +++++----- 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index cda9b954af115..0ea10c2a77b1f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -144,6 +144,7 @@ import Mathlib.CategoryTheory.Category.Basic import Mathlib.CategoryTheory.Category.KleisliCat import Mathlib.CategoryTheory.Category.RelCat import Mathlib.CategoryTheory.ConcreteCategory.Bundled +import Mathlib.CategoryTheory.Equivalence import Mathlib.CategoryTheory.EssentialImage import Mathlib.CategoryTheory.FullSubcategory import Mathlib.CategoryTheory.Functor.Basic diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index ec6879fbcf67e..80691b4c42a3a 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -8,11 +8,11 @@ Authors: Tim Baumann, Stephen Morgan, Scott Morrison, Floris van Doorn ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Functor.FullyFaithful -import Mathbin.CategoryTheory.FullSubcategory -import Mathbin.CategoryTheory.Whiskering -import Mathbin.CategoryTheory.EssentialImage -import Mathbin.Tactic.Slice +import Mathlib.CategoryTheory.Functor.FullyFaithful +import Mathlib.CategoryTheory.FullSubcategory +import Mathlib.CategoryTheory.Whiskering +import Mathlib.CategoryTheory.EssentialImage +import Mathlib.Tactic.Slice /-! # Equivalence of categories From afd9aed6550ed5f7bc5eb86cc41ba21af45bfa20 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Henrik=20B=C3=B6ving?= Date: Mon, 2 Jan 2023 13:21:50 +0100 Subject: [PATCH 03/66] push it as far as possible --- Mathlib/CategoryTheory/Equivalence.lean | 243 ++++++++++++------------ Mathlib/CategoryTheory/Iso.lean | 3 + 2 files changed, 121 insertions(+), 125 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 80691b4c42a3a..ae41d1168b7ba 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -12,8 +12,9 @@ import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.FullSubcategory import Mathlib.CategoryTheory.Whiskering import Mathlib.CategoryTheory.EssentialImage -import Mathlib.Tactic.Slice - +-- TODO: import Mathlib.Tactic.Slice +-- TODO: remove +set_option autoImplicit false /-! # Equivalence of categories @@ -40,18 +41,18 @@ if it is full, faithful and essentially surjective. ## Main definitions -* `equivalence`: bundled (half-)adjoint equivalences of categories -* `is_equivalence`: type class on a functor `F` containing the data of the inverse `G` as well as +* `Equivalence`: bundled (half-)adjoint equivalences of categories +* `IsEquivalence`: type class on a functor `F` containing the data of the inverse `G` as well as the natural isomorphisms `η` and `ε`. -* `ess_surj`: type class on a functor `F` containing the data of the preimages and the isomorphisms +* `EssSurj`: type class on a functor `F` containing the data of the preimages and the isomorphisms `F.obj (preimage d) ≅ d`. ## Main results -* `equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence -* `is_equivalence.equiv_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence +* `Equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence +* `IsEquivalence.equiv_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence iff `G` is. -* `equivalence.of_fully_faithfully_ess_surj`: a fully faithful essentially surjective functor is an +* `Equivalence.of_fully_faithfully_essSurj`: a fully faithful essentially surjective functor is an equivalence. ## Notations @@ -60,12 +61,11 @@ We write `C ≌ D` (`\backcong`, not to be confused with `≅`/`\cong`) for a bu -/ - namespace CategoryTheory open CategoryTheory.Functor NatIso Category --- declare the `v`'s first; see `category_theory.category` for an explanation +-- declare the `v`'s first; see `CategoryTheory.Category` for an explanation universe v₁ v₂ v₃ u₁ u₂ u₃ /-- We define an equivalence as a (half)-adjoint equivalence, a pair of functors with @@ -82,19 +82,14 @@ universe v₁ v₂ v₃ u₁ u₂ u₃ See -/ structure Equivalence (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] where mk' :: - Functor : C ⥤ D + functor : C ⥤ D inverse : D ⥤ C - unitIso : 𝟭 C ≅ Functor ⋙ inverse - counitIso : inverse ⋙ Functor ≅ 𝟭 D - functor_unit_iso_comp' : - ∀ X : C, - Functor.map ((unit_iso.Hom : 𝟭 C ⟶ Functor ⋙ inverse).app X) ≫ - counit_iso.Hom.app (Functor.obj X) = - 𝟙 (Functor.obj X) := by - obviously + unitIso : 𝟭 C ≅ functor ⋙ inverse + counitIso : inverse ⋙ functor ≅ 𝟭 D + functor_unit_iso_comp' : ∀ X : C, functor.map (unitIso.hom.app X) ≫ counitIso.hom.app (functor.obj X) = 𝟙 (functor.obj X) := by aesop_cat #align category_theory.equivalence CategoryTheory.Equivalence -restate_axiom equivalence.functor_unit_iso_comp' +restate_axiom Equivalence.functor_unit_iso_comp' -- mathport name: «expr ≌ » infixr:10 " ≌ " => Equivalence @@ -104,22 +99,22 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] namespace Equivalence /-- The unit of an equivalence of categories. -/ -abbrev unit (e : C ≌ D) : 𝟭 C ⟶ e.Functor ⋙ e.inverse := - e.unitIso.Hom +abbrev unit (e : C ≌ D) : 𝟭 C ⟶ e.functor ⋙ e.inverse := + e.unitIso.hom #align category_theory.equivalence.unit CategoryTheory.Equivalence.unit /-- The counit of an equivalence of categories. -/ -abbrev counit (e : C ≌ D) : e.inverse ⋙ e.Functor ⟶ 𝟭 D := - e.counitIso.Hom +abbrev counit (e : C ≌ D) : e.inverse ⋙ e.functor ⟶ 𝟭 D := + e.counitIso.hom #align category_theory.equivalence.counit CategoryTheory.Equivalence.counit /-- The inverse of the unit of an equivalence of categories. -/ -abbrev unitInv (e : C ≌ D) : e.Functor ⋙ e.inverse ⟶ 𝟭 C := +abbrev unitInv (e : C ≌ D) : e.functor ⋙ e.inverse ⟶ 𝟭 C := e.unitIso.inv #align category_theory.equivalence.unit_inv CategoryTheory.Equivalence.unitInv /-- The inverse of the counit of an equivalence of categories. -/ -abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.Functor := +abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.functor := e.counitIso.inv #align category_theory.equivalence.counit_inv CategoryTheory.Equivalence.counitInv @@ -127,61 +122,61 @@ abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.Functor := preventing structure projections from unfolding. -/ @[simp] theorem equivalence_mk'_unit (functor inverse unit_iso counit_iso f) : - (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).Unit = unit_iso.Hom := + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unit = unit_iso.hom := rfl #align category_theory.equivalence.equivalence_mk'_unit CategoryTheory.Equivalence.equivalence_mk'_unit @[simp] theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : - (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.Hom := + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.hom := rfl #align category_theory.equivalence.equivalence_mk'_counit CategoryTheory.Equivalence.equivalence_mk'_counit @[simp] -theorem equivalence_mk'_unit_inv (functor inverse unit_iso counit_iso f) : - (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := +theorem equivalence_mk'_unitInv (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := rfl #align - category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unit_inv + category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unitInv @[simp] -theorem equivalence_mk'_counit_inv (functor inverse unit_iso counit_iso f) : - (⟨Functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := +theorem equivalence_mk'_counitInv (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := rfl #align - category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counit_inv + category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counitInv @[simp] theorem functor_unit_comp (e : C ≌ D) (X : C) : - e.Functor.map (e.Unit.app X) ≫ e.counit.app (e.Functor.obj X) = 𝟙 (e.Functor.obj X) := + e.functor.map (e.unit.app X) ≫ e.counit.app (e.functor.obj X) = 𝟙 (e.functor.obj X) := e.functor_unit_iso_comp X #align category_theory.equivalence.functor_unit_comp CategoryTheory.Equivalence.functor_unit_comp @[simp] -theorem counit_inv_functor_comp (e : C ≌ D) (X : C) : - e.counitInv.app (e.Functor.obj X) ≫ e.Functor.map (e.unitInv.app X) = 𝟙 (e.Functor.obj X) := +theorem counitInv_functor_comp (e : C ≌ D) (X : C) : + e.counitInv.app (e.functor.obj X) ≫ e.functor.map (e.unitInv.app X) = 𝟙 (e.functor.obj X) := by - erw [iso.inv_eq_inv (e.functor.map_iso (e.unit_iso.app X) ≪≫ e.counit_iso.app (e.functor.obj X)) - (iso.refl _)] + erw [Iso.inv_eq_inv (e.functor.mapIso (e.unitIso.app X) ≪≫ e.counitIso.app (e.functor.obj X)) + (Iso.refl _)] exact e.functor_unit_comp X #align - category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counit_inv_functor_comp + category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counitInv_functor_comp theorem counit_inv_app_functor (e : C ≌ D) (X : C) : - e.counitInv.app (e.Functor.obj X) = e.Functor.map (e.Unit.app X) := + e.counitInv.app (e.functor.obj X) = e.functor.map (e.unit.app X) := by symm - erw [← iso.comp_hom_eq_id (e.counit_iso.app _), functor_unit_comp] + erw [← Iso.comp_hom_eq_id (e.counitIso.app _), functor_unit_comp] rfl #align category_theory.equivalence.counit_inv_app_functor CategoryTheory.Equivalence.counit_inv_app_functor theorem counit_app_functor (e : C ≌ D) (X : C) : - e.counit.app (e.Functor.obj X) = e.Functor.map (e.unitInv.app X) := + e.counit.app (e.functor.obj X) = e.functor.map (e.unitInv.app X) := by - erw [← iso.hom_comp_eq_id (e.functor.map_iso (e.unit_iso.app X)), functor_unit_comp] + erw [← Iso.hom_comp_eq_id (e.functor.mapIso (e.unitIso.app X)), functor_unit_comp] rfl #align category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor @@ -217,37 +212,36 @@ theorem unit_inverse_comp (e : C ≌ D) (Y : D) : theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : e.inverse.map (e.counitInv.app Y) ≫ e.unitInv.app (e.inverse.obj Y) = 𝟙 (e.inverse.obj Y) := by - erw [iso.inv_eq_inv (e.unit_iso.app (e.inverse.obj Y) ≪≫ e.inverse.map_iso (e.counit_iso.app Y)) - (iso.refl _)] + erw [Iso.inv_eq_inv (e.unitIso.app (e.inverse.obj Y) ≪≫ e.inverse.mapIso (e.counitIso.app Y)) + (Iso.refl _)] exact e.unit_inverse_comp Y #align category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp theorem unit_app_inverse (e : C ≌ D) (Y : D) : - e.Unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := + e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := by - erw [← iso.comp_hom_eq_id (e.inverse.map_iso (e.counit_iso.app Y)), unit_inverse_comp] - rfl + erw [← Iso.comp_hom_eq_id (e.inverse.mapIso (e.counitIso.app Y)), unit_inverse_comp] #align category_theory.equivalence.unit_app_inverse CategoryTheory.Equivalence.unit_app_inverse theorem unit_inv_app_inverse (e : C ≌ D) (Y : D) : e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := by symm - erw [← iso.hom_comp_eq_id (e.unit_iso.app _), unit_inverse_comp] + erw [← Iso.hom_comp_eq_id (e.unitIso.app _), unit_inverse_comp] rfl #align category_theory.equivalence.unit_inv_app_inverse CategoryTheory.Equivalence.unit_inv_app_inverse @[simp] theorem fun_inv_map (e : C ≌ D) (X Y : D) (f : X ⟶ Y) : - e.Functor.map (e.inverse.map f) = e.counit.app X ≫ f ≫ e.counitInv.app Y := + e.functor.map (e.inverse.map f) = e.counit.app X ≫ f ≫ e.counitInv.app Y := (NatIso.naturality_2 e.counitIso f).symm #align category_theory.equivalence.fun_inv_map CategoryTheory.Equivalence.fun_inv_map @[simp] theorem inv_fun_map (e : C ≌ D) (X Y : C) (f : X ⟶ Y) : - e.inverse.map (e.Functor.map f) = e.unitInv.app X ≫ f ≫ e.Unit.app Y := + e.inverse.map (e.functor.map f) = e.unitInv.app X ≫ f ≫ e.unit.app Y := (NatIso.naturality_1 e.unitIso f).symm #align category_theory.equivalence.inv_fun_map CategoryTheory.Equivalence.inv_fun_map @@ -272,9 +266,9 @@ def adjointifyη : 𝟭 C ≅ F ⋙ G := #align category_theory.equivalence.adjointify_η CategoryTheory.Equivalence.adjointifyη theorem adjointify_η_ε (X : C) : - F.map ((adjointifyη η ε).Hom.app X) ≫ ε.Hom.app (F.obj X) = 𝟙 (F.obj X) := + F.map ((adjointifyη η ε).hom.app X) ≫ ε.hom.app (F.obj X) = 𝟙 (F.obj X) := by - dsimp [adjointify_η]; simp + dsimp [adjointifyη]; simp have := ε.hom.naturality (F.map (η.inv.app X)); dsimp at this; rw [this]; clear this rw [← assoc _ _ (F.map _)] have := ε.hom.naturality (ε.inv.app <| F.obj X); dsimp at this; rw [this]; clear this @@ -303,7 +297,7 @@ instance : Inhabited (C ≌ C) := /-- Equivalence of categories is symmetric. -/ @[symm, simps] def symm (e : C ≌ D) : D ≌ C := - ⟨e.inverse, e.Functor, e.counitIso.symm, e.unitIso.symm, e.inverse_counit_inv_comp⟩ + ⟨e.inverse, e.functor, e.counitIso.symm, e.unitIso.symm, e.inverse_counit_inv_comp⟩ #align category_theory.equivalence.symm CategoryTheory.Equivalence.symm variable {E : Type u₃} [Category.{v₃} E] @@ -312,60 +306,60 @@ variable {E : Type u₃} [Category.{v₃} E] @[trans, simps] def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E where - Functor := e.Functor ⋙ f.Functor + functor := e.functor ⋙ f.functor inverse := f.inverse ⋙ e.inverse unitIso := by - refine' iso.trans e.unit_iso _ - exact iso_whisker_left e.functor (iso_whisker_right f.unit_iso e.inverse) + refine' Iso.trans e.unitIso _ + exact isoWhiskerLeft e.functor (isoWhiskerRight f.unitIso e.inverse) counitIso := by - refine' iso.trans _ f.counit_iso - exact iso_whisker_left f.inverse (iso_whisker_right e.counit_iso f.functor) + refine' Iso.trans _ f.counitIso + exact isoWhiskerLeft f.inverse (isoWhiskerRight e.counitIso f.functor) -- We wouldn't have needed to give this proof if we'd used `equivalence.mk`, -- but we choose to avoid using that here, for the sake of good structure projection `simp` -- lemmas. functor_unit_iso_comp' X := by dsimp rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counit_inv_app_functor, fun_inv_map, - iso.inv_hom_id_app_assoc, assoc, iso.inv_hom_id_app, counit_app_functor, ← functor.map_comp] - erw [comp_id, iso.hom_inv_id_app, Functor.map_id] + Iso.inv_hom_id_app_assoc, assoc, Iso.inv_hom_id_app, counit_app_functor, ← Functor.map_comp] + erw [comp_id, Iso.hom_inv_id_app, Functor.map_id] #align category_theory.equivalence.trans CategoryTheory.Equivalence.trans /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ -def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.Functor ⋙ e.inverse ⋙ F ≅ F := +def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F ≅ F := (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.unitIso.symm F ≪≫ F.leftUnitor #align category_theory.equivalence.fun_inv_id_assoc CategoryTheory.Equivalence.funInvIdAssoc @[simp] theorem fun_inv_id_assoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : - (funInvIdAssoc e F).Hom.app X = F.map (e.unitInv.app X) := + (funInvIdAssoc e F).hom.app X = F.map (e.unitInv.app X) := by - dsimp [fun_inv_id_assoc] - tidy + dsimp [funInvIdAssoc] + aesop_cat #align category_theory.equivalence.fun_inv_id_assoc_hom_app CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app @[simp] theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : - (funInvIdAssoc e F).inv.app X = F.map (e.Unit.app X) := + (funInvIdAssoc e F).inv.app X = F.map (e.unit.app X) := by - dsimp [fun_inv_id_assoc] - tidy + dsimp [funInvIdAssoc] + aesop_cat #align category_theory.equivalence.fun_inv_id_assoc_inv_app CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ -def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.Functor ⋙ F ≅ F := +def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.functor ⋙ F ≅ F := (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.counitIso F ≪≫ F.leftUnitor #align category_theory.equivalence.inv_fun_id_assoc CategoryTheory.Equivalence.invFunIdAssoc @[simp] theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : - (invFunIdAssoc e F).Hom.app X = F.map (e.counit.app X) := + (invFunIdAssoc e F).hom.app X = F.map (e.counit.app X) := by - dsimp [inv_fun_id_assoc] - tidy + dsimp [invFunIdAssoc] + aesop_cat #align category_theory.equivalence.inv_fun_id_assoc_hom_app CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app @@ -373,29 +367,29 @@ theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := by - dsimp [inv_fun_id_assoc] - tidy + dsimp [invFunIdAssoc] + aesop_cat #align category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app /-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ -@[simps Functor inverse unitIso counitIso] +@[simps functor inverse unitIso counitIso] def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := - Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.Functor) - (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by tidy)) - (NatIso.ofComponents (fun F => e.invFunIdAssoc F) (by tidy)) + Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.functor) + (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by aesop_cat)) + (NatIso.ofComponents (fun F => e.invFunIdAssoc F) (by aesop_cat)) #align category_theory.equivalence.congr_left CategoryTheory.Equivalence.congrLeft /-- If `C` is equivalent to `D`, then `E ⥤ C` is equivalent to `E ⥤ D`. -/ -@[simps Functor inverse unitIso counitIso] +@[simps functor inverse unitIso counitIso] def congrRight (e : C ≌ D) : E ⥤ C ≌ E ⥤ D := - Equivalence.mk ((whiskeringRight _ _ _).obj e.Functor) ((whiskeringRight _ _ _).obj e.inverse) + Equivalence.mk ((whiskeringRight _ _ _).obj e.functor) ((whiskeringRight _ _ _).obj e.inverse) (NatIso.ofComponents (fun F => F.rightUnitor.symm ≪≫ isoWhiskerLeft F e.unitIso ≪≫ Functor.associator _ _ _) - (by tidy)) + (by aesop_cat)) (NatIso.ofComponents (fun F => Functor.associator _ _ _ ≪≫ isoWhiskerLeft F e.counitIso ≪≫ F.rightUnitor) - (by tidy)) + (by aesop_cat)) #align category_theory.equivalence.congr_right CategoryTheory.Equivalence.congrRight section CancellationLemmas @@ -409,17 +403,17 @@ We also provide the lemmas for length four compositions, since they're occasiona (e.g. in proving that equivalences take monos to monos) -/ @[simp] theorem cancel_unit_right {X Y : C} (f f' : X ⟶ Y) : - f ≫ e.Unit.app Y = f' ≫ e.Unit.app Y ↔ f = f' := by simp only [cancel_mono] + f ≫ e.unit.app Y = f' ≫ e.unit.app Y ↔ f = f' := by simp only [cancel_mono] #align category_theory.equivalence.cancel_unit_right CategoryTheory.Equivalence.cancel_unit_right @[simp] -theorem cancel_unit_inv_right {X Y : C} (f f' : X ⟶ e.inverse.obj (e.Functor.obj Y)) : +theorem cancel_unit_inv_right {X Y : C} (f f' : X ⟶ e.inverse.obj (e.functor.obj Y)) : f ≫ e.unitInv.app Y = f' ≫ e.unitInv.app Y ↔ f = f' := by simp only [cancel_mono] #align category_theory.equivalence.cancel_unit_inv_right CategoryTheory.Equivalence.cancel_unit_inv_right @[simp] -theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.Functor.obj (e.inverse.obj Y)) : +theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.functor.obj (e.inverse.obj Y)) : f ≫ e.counit.app Y = f' ≫ e.counit.app Y ↔ f = f' := by simp only [cancel_mono] #align category_theory.equivalence.cancel_counit_right CategoryTheory.Equivalence.cancel_counit_right @@ -432,7 +426,7 @@ theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : @[simp] theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : - f ≫ g ≫ e.Unit.app Y = f' ≫ g' ≫ e.Unit.app Y ↔ f ≫ g = f' ≫ g' := by + f ≫ g ≫ e.unit.app Y = f' ≫ g' ≫ e.unit.app Y ↔ f ≫ g = f' ≫ g' := by simp only [← category.assoc, cancel_mono] #align category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc @@ -447,7 +441,7 @@ theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) @[simp] theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : - f ≫ g ≫ h ≫ e.Unit.app Z = f' ≫ g' ≫ h' ≫ e.Unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by + f ≫ g ≫ h ≫ e.unit.app Z = f' ≫ g' ≫ h' ≫ e.unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by simp only [← category.assoc, cancel_mono] #align category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' @@ -471,7 +465,7 @@ section def powNat (e : C ≌ C) : ℕ → (C ≌ C) | 0 => Equivalence.refl | 1 => e - | n + 2 => e.trans (pow_nat (n + 1)) + | n + 2 => e.trans (powNat e (n + 1)) #align category_theory.equivalence.pow_nat CategoryTheory.Equivalence.powNat /-- Powers of an auto-equivalence. Use `(^)` instead. -/ @@ -484,7 +478,7 @@ instance : Pow (C ≌ C) ℤ := ⟨pow⟩ @[simp] -theorem pow_zero (e : C ≌ C) : e ^ (0 : ℤ) = equivalence.refl := +theorem pow_zero (e : C ≌ C) : e ^ (0 : ℤ) = Equivalence.refl := rfl #align category_theory.equivalence.pow_zero CategoryTheory.Equivalence.pow_zero @@ -511,18 +505,18 @@ class IsEquivalence (F : C ⥤ D) where mk' :: counitIso : inverse ⋙ F ≅ 𝟭 D functor_unit_iso_comp' : ∀ X : C, - F.map ((unit_iso.Hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counit_iso.Hom.app (F.obj X) = + F.map ((unitIso.hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counitIso.hom.app (F.obj X) = 𝟙 (F.obj X) := by - obviously + aesop_cat #align category_theory.is_equivalence CategoryTheory.IsEquivalence -restate_axiom is_equivalence.functor_unit_iso_comp' +restate_axiom IsEquivalence.functor_unit_iso_comp' attribute [simp, reassoc.1] is_equivalence.functor_unit_iso_comp namespace IsEquivalence -instance ofEquivalence (F : C ≌ D) : IsEquivalence F.Functor := +instance ofEquivalence (F : C ≌ D) : IsEquivalence F.functor := { F with } #align category_theory.is_equivalence.of_equivalence CategoryTheory.IsEquivalence.ofEquivalence @@ -563,7 +557,7 @@ instance isEquivalenceInv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv #align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalenceInv @[simp] -theorem as_equivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.Functor = F := +theorem as_equivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.functor = F := rfl #align category_theory.functor.as_equivalence_functor CategoryTheory.Functor.as_equivalence_functor @@ -573,14 +567,14 @@ theorem as_equivalence_inverse (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence #align category_theory.functor.as_equivalence_inverse CategoryTheory.Functor.as_equivalence_inverse @[simp] -theorem as_equivalence_unit {F : C ⥤ D} [h : IsEquivalence F] : - F.asEquivalence.unitIso = @IsEquivalence.unitIso _ _ h := +theorem as_equivalence_unit {F : C ⥤ D} [IsEquivalence F] : + F.asEquivalence.unitIso = IsEquivalence.unitIso := rfl #align category_theory.functor.as_equivalence_unit CategoryTheory.Functor.as_equivalence_unit @[simp] theorem as_equivalence_counit {F : C ⥤ D} [IsEquivalence F] : - F.asEquivalence.counitIso = is_equivalence.counit_iso := + F.asEquivalence.counitIso = IsEquivalence.counitIso := rfl #align category_theory.functor.as_equivalence_counit CategoryTheory.Functor.as_equivalence_counit @@ -601,17 +595,17 @@ end Functor namespace Equivalence @[simp] -theorem functor_inv (E : C ≌ D) : E.Functor.inv = E.inverse := +theorem functor_inv (E : C ≌ D) : E.functor.inv = E.inverse := rfl #align category_theory.equivalence.functor_inv CategoryTheory.Equivalence.functor_inv @[simp] -theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.Functor := +theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.functor := rfl #align category_theory.equivalence.inverse_inv CategoryTheory.Equivalence.inverse_inv @[simp] -theorem functor_as_equivalence (E : C ≌ D) : E.Functor.asEquivalence = E := +theorem functor_as_equivalence (E : C ≌ D) : E.functor.asEquivalence = E := by cases E congr @@ -634,15 +628,15 @@ namespace IsEquivalence theorem fun_inv_map (F : C ⥤ D) [IsEquivalence F] (X Y : D) (f : X ⟶ Y) : F.map (F.inv.map f) = F.asEquivalence.counit.app X ≫ f ≫ F.asEquivalence.counitInv.app Y := by - erw [nat_iso.naturality_2] + erw [NatIso.naturality_2] rfl #align category_theory.is_equivalence.fun_inv_map CategoryTheory.IsEquivalence.fun_inv_map @[simp] theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : - F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.Unit.app Y := + F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.unit.app Y := by - erw [nat_iso.naturality_1] + erw [NatIso.naturality_1] rfl #align category_theory.is_equivalence.inv_fun_map CategoryTheory.IsEquivalence.inv_fun_map @@ -670,9 +664,9 @@ def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G theorem of_iso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEquivalence F) : ofIso e' (ofIso e hF) = ofIso (e ≪≫ e') hF := by - dsimp [of_iso] - congr 1 <;> ext X <;> dsimp [nat_iso.hcomp] - · simp only [id_comp, assoc, functor.map_comp] + dsimp [ofIso] + congr 1 <;> ext X <;> dsimp [NatIso.hcomp] + · simp only [id_comp, assoc, Functor.map_comp] · simp only [Functor.map_id, comp_id, id_comp, assoc] #align category_theory.is_equivalence.of_iso_trans CategoryTheory.IsEquivalence.of_iso_trans @@ -680,8 +674,8 @@ theorem of_iso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEqui theorem of_iso_refl (F : C ⥤ D) (hF : IsEquivalence F) : ofIso (Iso.refl F) hF = hF := by rcases hF with ⟨Finv, Funit, Fcounit, Fcomp⟩ - dsimp [of_iso] - congr 1 <;> ext X <;> dsimp [nat_iso.hcomp] + dsimp [ofIso] + congr 1 <;> ext X <;> dsimp [NatIso.hcomp] · simp only [comp_id, map_id] · simp only [id_comp, map_id] #align category_theory.is_equivalence.of_iso_refl CategoryTheory.IsEquivalence.of_iso_refl @@ -692,8 +686,8 @@ def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence where toFun := ofIso e invFun := ofIso e.symm - left_inv hF := by rw [of_iso_trans, iso.self_symm_id, of_iso_refl] - right_inv hF := by rw [of_iso_trans, iso.symm_self_id, of_iso_refl] + left_inv hF := by rw [of_iso_trans, Iso.self_symm_id, of_iso_refl] + right_inv hF := by rw [of_iso_trans, Iso.symm_self_id, of_iso_refl] #align category_theory.is_equivalence.equiv_of_iso CategoryTheory.IsEquivalence.equivOfIso /-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ @@ -731,9 +725,8 @@ theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := See . -/ -instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F - where map_injective' X Y f g w := - by +instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F where + map_injective X Y f g w := by have p := congr_arg (@CategoryTheory.Functor.map _ _ _ _ F.inv _ _) w simpa only [cancel_epi, cancel_mono, is_equivalence.inv_fun_map] using p #align @@ -746,22 +739,22 @@ See . -/ instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F where - preimage X Y f := F.asEquivalence.Unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y - witness' X Y f := + preimage X Y f := F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y + witness X Y f := F.inv.map_injective <| by simpa only [is_equivalence.inv_fun_map, assoc, iso.inv_hom_id_app_assoc, iso.inv_hom_id_app] using comp_id _ #align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.fullOfEquivalence @[simps] -private noncomputable def equivalence_inverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C +private noncomputable def equivalenceInverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C where obj X := F.objPreimage X map X Y f := F.preimage ((F.objObjPreimageIso X).Hom ≫ f ≫ (F.objObjPreimageIso Y).inv) - map_id' X := by apply F.map_injective; tidy - map_comp' X Y Z f g := by apply F.map_injective <;> simp + map_id X := by apply F.map_injective; aesop_cat + map_comp X Y Z f g := by apply F.map_injective <;> simp #align - category_theory.equivalence.equivalence_inverse category_theory.equivalence.equivalence_inverse + category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse /-- A functor which is full, faithful, and essentially surjective is an equivalence. @@ -773,15 +766,15 @@ noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [ (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) fun X Y f => by apply F.map_injective - obviously) - (NatIso.ofComponents F.objObjPreimageIso (by tidy)) + aesop_cat) + (NatIso.ofComponents F.objObjPreimageIso (by aesop_cat)) #align category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj @[simp] theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : - e.Functor.map f = e.Functor.map g ↔ f = g := - ⟨fun h => e.Functor.map_injective h, fun h => h ▸ rfl⟩ + e.functor.map f = e.functor.map g ↔ f = g := + ⟨fun h => e.functor.map_injective h, fun h => h ▸ rfl⟩ #align category_theory.equivalence.functor_map_inj_iff CategoryTheory.Equivalence.functor_map_inj_iff @@ -793,7 +786,7 @@ theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) - where mem_ess_image Y := ⟨e.symm Y, by simp⟩ + where mem_essImage Y := ⟨e.symm Y, by simp⟩ #align category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index b749044fa66cd..da44222129822 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -145,6 +145,9 @@ def trans (α : X ≅ Y) (β : Y ≅ Z) : X ≅ Z where inv := β.inv ≫ α.inv #align category_theory.iso.trans CategoryTheory.Iso.trans +instance : Trans (Iso : C → C → Type v) (Iso : C → C → Type v) (Iso : C → C → Type v) where + trans := Iso.trans + /-- Notation for composition of isomorphisms. -/ infixr:80 " ≪≫ " => Iso.trans -- type as `\ll \gg`. From f0f65020500fe71f52cca6477686ec43400ee9a2 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 3 Feb 2023 17:00:44 -0500 Subject: [PATCH 04/66] seems to work --- Mathlib/Tactic/Core.lean | 28 ++++++++++++++++ Mathlib/Tactic/Slice.lean | 70 +++++++++++++++++++++++++++++++++++++++ 2 files changed, 98 insertions(+) create mode 100644 Mathlib/Tactic/Slice.lean diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 17866714d4752..9a2648c1fb895 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -95,6 +95,16 @@ end Lean namespace Lean.Elab.Tactic +/-- +Variant on `rewriteTarget` that behaves more like the Lean 3 version +-/ +def rewriteTarget' (stx : Syntax) (symm : Bool) : TacticM Unit := do + Term.withSynthesize <| withMainContext do + let e ← elabTerm stx none true + let r ← (← getMainGoal).rewrite (← getMainTarget) e symm (config := {}) + let mvarId' ← (← getMainGoal).replaceTargetEq r.eNew r.eqProof + replaceMainGoal ([mvarId']) + /-- Run a tactic on all goals, and always succeeds. @@ -133,6 +143,24 @@ def iterateAtMost : Nat → m Unit → m Unit | 0, _ => pure () | n + 1, tac => try tac; iterateAtMost n tac catch _ => pure () +/-- `iterateExactly' n t` executes `t` `n` times. If any iteration fails, the whole tactic fails. +-/ +def iterateExactly' : Nat → m Unit → m Unit +| 0, _ => pure () +| n+1, tac => tac *> iterateExactly' n tac + +/-- +`iterateRange m n t`: Repeat the given tactic at least `m` times and +at most `n` times or until `t` fails. Fails if `t` does not run at least `m` times. +-/ +def iterateRange : Nat → Nat → m Unit → m Unit +| 0, 0, _ => pure () +| 0, b, tac => iterateAtMost b tac +| (a+1), n, tac => do + let _ ← tac + let _ ← iterateRange a (n-1) tac + pure () + /-- Repeats a tactic until it fails. Always succeeds. -/ partial def iterateUntilFailure (tac : m Unit) : m Unit := try tac; iterateUntilFailure tac catch _ => pure () diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean new file mode 100644 index 0000000000000..0eaf24c8dbe43 --- /dev/null +++ b/Mathlib/Tactic/Slice.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2018 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison + +-/ +import Mathlib.CategoryTheory.Category.Basic +import Mathlib.Tactic.Conv + +open CategoryTheory +open Lean Parser.Tactic Elab Command Elab.Tactic Meta + +-- TODO someone might like to generalise this tactic to work with other associative structures. +namespace Tactic + +variable [Monad m] [MonadExceptOf Exception m] + +partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α) := do + try + let a ← tac + let l ← iterateUntilFailureWithResults tac + pure (a :: l) + catch _ => pure [] +#align tactic.repeat_with_results Tactic.iterateUntilFailureWithResults + +def iterateUntilFailureCount {α : Type} (tac : m α) : m ℕ := do + let r ← iterateUntilFailureWithResults tac + return r.length +#align tactic.repeat_count Tactic.iterateUntilFailureCount + +end Tactic + +namespace Conv + +open Tactic + +variable [Monad m] [MonadExceptOf Exception m] + +def evalSlice (a b : Nat) : TacticM Unit := do + iterateUntilFailure do + ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + iterateRange (a - 1) (a - 1) do + evalTactic (← `(conv| congr)) + evalTactic (← `(tactic| rotate_left)) + let k ← iterateUntilFailureCount + <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) + let c := k+1+a-b + iterateRange c c <| evalTactic (← `(conv| congr)) + iterateUntilFailure do + ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + +elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat + +open Lean Parser.Tactic Parser.Tactic.Conv Elab.Tactic Meta +syntax (name := sliceLHS) "sliceLHS" num num " => " convSeq : tactic +macro_rules + | `(tactic| sliceLHS $a $b => $seq) => + `(tactic| conv => lhs; slice $a $b; ($seq:convSeq)) + +syntax (name := sliceRHS) "sliceRHS" num num " => " convSeq : tactic +macro_rules + | `(tactic| sliceRHS $a $b => $seq) => + `(tactic| conv => rhs; slice $a $b; ($seq:convSeq)) + +-- add_tactic_doc +-- { Name := "slice" +-- category := DocCategory.tactic +-- declNames := [`tactic.interactive.slice_lhs, `tactic.interactive.slice_rhs] +-- tags := ["category theory"] } +-- From 91f541a86962d4423416dd7573883a95bbb71934 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 7 Feb 2023 14:00:33 -0500 Subject: [PATCH 05/66] add example and equivalence file for testing --- Mathlib/CategoryTheory/Equivalence.lean | 809 ++++++++++++++++++++++++ Mathlib/Tactic/Core.lean | 1 + Mathlib/Tactic/Slice.lean | 19 +- lean-toolchain | 2 +- 4 files changed, 827 insertions(+), 4 deletions(-) create mode 100644 Mathlib/CategoryTheory/Equivalence.lean diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean new file mode 100644 index 0000000000000..673184c6cb4fb --- /dev/null +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -0,0 +1,809 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Tim Baumann, Stephen Morgan, Scott Morrison, Floris van Doorn + +! This file was ported from Lean 3 source module category_theory.equivalence +! leanprover-community/mathlib commit 9aba7801eeecebb61f58a5763c2b6dd1b47dc6ef +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathlib.CategoryTheory.Functor.FullyFaithful +import Mathlib.CategoryTheory.FullSubcategory +import Mathlib.CategoryTheory.Whiskering +import Mathlib.CategoryTheory.EssentialImage +import Mathlib.Tactic.Slice +-- TODO: remove +set_option autoImplicit false +/-! +# Equivalence of categories + +An equivalence of categories `C` and `D` is a pair of functors `F : C ⥤ D` and `G : D ⥤ C` such +that `η : 𝟭 C ≅ F ⋙ G` and `ε : G ⋙ F ≅ 𝟭 D`. In many situations, equivalences are a better +notion of "sameness" of categories than the stricter isomorphims of categories. + +Recall that one way to express that two functors `F : C ⥤ D` and `G : D ⥤ C` are adjoint is using +two natural transformations `η : 𝟭 C ⟶ F ⋙ G` and `ε : G ⋙ F ⟶ 𝟭 D`, called the unit and the +counit, such that the compositions `F ⟶ FGF ⟶ F` and `G ⟶ GFG ⟶ G` are the identity. Unfortunately, +it is not the case that the natural isomorphisms `η` and `ε` in the definition of an equivalence +automatically give an adjunction. However, it is true that +* if one of the two compositions is the identity, then so is the other, and +* given an equivalence of categories, it is always possible to refine `η` in such a way that the + identities are satisfied. + +For this reason, in mathlib we define an equivalence to be a "half-adjoint equivalence", which is +a tuple `(F, G, η, ε)` as in the first paragraph such that the composite `F ⟶ FGF ⟶ F` is the +identity. By the remark above, this already implies that the tuple is an "adjoint equivalence", +i.e., that the composite `G ⟶ GFG ⟶ G` is also the identity. + +We also define essentially surjective functors and show that a functor is an equivalence if and only +if it is full, faithful and essentially surjective. + +## Main definitions + +* `Equivalence`: bundled (half-)adjoint equivalences of categories +* `IsEquivalence`: type class on a functor `F` containing the data of the inverse `G` as well as + the natural isomorphisms `η` and `ε`. +* `EssSurj`: type class on a functor `F` containing the data of the preimages and the isomorphisms + `F.obj (preimage d) ≅ d`. + +## Main results + +* `Equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence +* `IsEquivalence.equiv_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence +iff `G` is. +* `Equivalence.of_fully_faithfully_essSurj`: a fully faithful essentially surjective functor is an + equivalence. + +## Notations + +We write `C ≌ D` (`\backcong`, not to be confused with `≅`/`\cong`) for a bundled equivalence. + +-/ + +namespace CategoryTheory + +open CategoryTheory.Functor NatIso Category + +-- declare the `v`'s first; see `CategoryTheory.Category` for an explanation +universe v₁ v₂ v₃ u₁ u₂ u₃ + +/-- We define an equivalence as a (half)-adjoint equivalence, a pair of functors with + a unit and counit which are natural isomorphisms and the triangle law `Fη ≫ εF = 1`, or in other + words the composite `F ⟶ FGF ⟶ F` is the identity. + + In `unit_inverse_comp`, we show that this is actually an adjoint equivalence, i.e., that the + composite `G ⟶ GFG ⟶ G` is also the identity. + + The triangle equation is written as a family of equalities between morphisms, it is more + complicated if we write it as an equality of natural transformations, because then we would have + to insert natural transformations like `F ⟶ F1`. + +See +-/ +structure Equivalence (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] where mk' :: + functor : C ⥤ D + inverse : D ⥤ C + unitIso : 𝟭 C ≅ functor ⋙ inverse + counitIso : inverse ⋙ functor ≅ 𝟭 D + functor_unit_iso_comp' : ∀ X : C, functor.map (unitIso.hom.app X) ≫ counitIso.hom.app (functor.obj X) = 𝟙 (functor.obj X) := by aesop_cat +#align category_theory.equivalence CategoryTheory.Equivalence + +restate_axiom Equivalence.functor_unit_iso_comp' + +-- mathport name: «expr ≌ » +infixr:10 " ≌ " => Equivalence + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +namespace Equivalence + +/-- The unit of an equivalence of categories. -/ +abbrev unit (e : C ≌ D) : 𝟭 C ⟶ e.functor ⋙ e.inverse := + e.unitIso.hom +#align category_theory.equivalence.unit CategoryTheory.Equivalence.unit + +/-- The counit of an equivalence of categories. -/ +abbrev counit (e : C ≌ D) : e.inverse ⋙ e.functor ⟶ 𝟭 D := + e.counitIso.hom +#align category_theory.equivalence.counit CategoryTheory.Equivalence.counit + +/-- The inverse of the unit of an equivalence of categories. -/ +abbrev unitInv (e : C ≌ D) : e.functor ⋙ e.inverse ⟶ 𝟭 C := + e.unitIso.inv +#align category_theory.equivalence.unit_inv CategoryTheory.Equivalence.unitInv + +/-- The inverse of the counit of an equivalence of categories. -/ +abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.functor := + e.counitIso.inv +#align category_theory.equivalence.counit_inv CategoryTheory.Equivalence.counitInv + +/- While these abbreviations are convenient, they also cause some trouble, +preventing structure projections from unfolding. -/ +@[simp] +theorem equivalence_mk'_unit (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unit = unit_iso.hom := + rfl +#align + category_theory.equivalence.equivalence_mk'_unit CategoryTheory.Equivalence.equivalence_mk'_unit + +@[simp] +theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.hom := + rfl +#align + category_theory.equivalence.equivalence_mk'_counit CategoryTheory.Equivalence.equivalence_mk'_counit + +@[simp] +theorem equivalence_mk'_unitInv (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := + rfl +#align + category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unitInv + +@[simp] +theorem equivalence_mk'_counitInv (functor inverse unit_iso counit_iso f) : + (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := + rfl +#align + category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counitInv + +@[simp] +theorem functor_unit_comp (e : C ≌ D) (X : C) : + e.functor.map (e.unit.app X) ≫ e.counit.app (e.functor.obj X) = 𝟙 (e.functor.obj X) := + e.functor_unit_iso_comp X +#align category_theory.equivalence.functor_unit_comp CategoryTheory.Equivalence.functor_unit_comp + +@[simp] +theorem counitInv_functor_comp (e : C ≌ D) (X : C) : + e.counitInv.app (e.functor.obj X) ≫ e.functor.map (e.unitInv.app X) = 𝟙 (e.functor.obj X) := + by + erw [Iso.inv_eq_inv (e.functor.mapIso (e.unitIso.app X) ≪≫ e.counitIso.app (e.functor.obj X)) + (Iso.refl _)] + exact e.functor_unit_comp X +#align + category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counitInv_functor_comp + +theorem counit_inv_app_functor (e : C ≌ D) (X : C) : + e.counitInv.app (e.functor.obj X) = e.functor.map (e.unit.app X) := + by + symm + erw [← Iso.comp_hom_eq_id (e.counitIso.app _), functor_unit_comp] + rfl +#align + category_theory.equivalence.counit_inv_app_functor CategoryTheory.Equivalence.counit_inv_app_functor + +theorem counit_app_functor (e : C ≌ D) (X : C) : + e.counit.app (e.functor.obj X) = e.functor.map (e.unitInv.app X) := + by + erw [← Iso.hom_comp_eq_id (e.functor.mapIso (e.unitIso.app X)), functor_unit_comp] + rfl +#align category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor + +/-- The other triangle equality. The proof follows the following proof in Globular: + http://globular.science/1905.001 -/ +@[simp] +theorem unit_inverse_comp (e : C ≌ D) (Y : D) : + e.unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := + by + rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counitInv_functor_comp, map_comp] + dsimp + rw [← Iso.hom_inv_id_assoc (e.unitIso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] + sliceLHS 2 3 => erw [e.unit.naturality] + sliceLHS 1 2 => erw [e.unit.naturality] + sliceLHS 4 4 => + rw [← Iso.hom_inv_id_assoc (e.inverse.mapIso (e.counitIso.app _)) (e.unitInv.app _)] + sliceLHS 3 4 => + erw [← map_comp e.inverse, e.counit.naturality] + erw [(e.counitIso.app _).hom_inv_id, map_id] + erw [id_comp] + sliceLHS 2 3 => erw [← map_comp e.inverse, e.counitIso.inv.naturality, map_comp] + sliceLHS 3 4 => erw [e.unitInv.naturality] + sliceLHS 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unitIso.app _).hom_inv_id, map_id] + erw [id_comp] + sliceLHS 3 4 => erw [← e.unitInv.naturality] + sliceLHS 2 3 => + erw [← map_comp e.inverse, ← e.counitIso.inv.naturality, (e.counitIso.app _).hom_inv_id, + map_id] + erw [id_comp, (e.unitIso.app _).hom_inv_id]; rfl + rfl +#align category_theory.equivalence.unit_inverse_comp CategoryTheory.Equivalence.unit_inverse_comp + +@[simp] +theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : + e.inverse.map (e.counitInv.app Y) ≫ e.unitInv.app (e.inverse.obj Y) = 𝟙 (e.inverse.obj Y) := + by + erw [Iso.inv_eq_inv (e.unitIso.app (e.inverse.obj Y) ≪≫ e.inverse.mapIso (e.counitIso.app Y)) + (Iso.refl _)] + exact e.unit_inverse_comp Y +#align + category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp + +theorem unit_app_inverse (e : C ≌ D) (Y : D) : + e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := + by + erw [← Iso.comp_hom_eq_id (e.inverse.mapIso (e.counitIso.app Y)), unit_inverse_comp] +#align category_theory.equivalence.unit_app_inverse CategoryTheory.Equivalence.unit_app_inverse + +theorem unit_inv_app_inverse (e : C ≌ D) (Y : D) : + e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := + by + symm + erw [← Iso.hom_comp_eq_id (e.unitIso.app _), unit_inverse_comp] + rfl +#align + category_theory.equivalence.unit_inv_app_inverse CategoryTheory.Equivalence.unit_inv_app_inverse + +@[simp] +theorem fun_inv_map (e : C ≌ D) (X Y : D) (f : X ⟶ Y) : + e.functor.map (e.inverse.map f) = e.counit.app X ≫ f ≫ e.counitInv.app Y := + (NatIso.naturality_2 e.counitIso f).symm +#align category_theory.equivalence.fun_inv_map CategoryTheory.Equivalence.fun_inv_map + +@[simp] +theorem inv_fun_map (e : C ≌ D) (X Y : C) (f : X ⟶ Y) : + e.inverse.map (e.functor.map f) = e.unitInv.app X ≫ f ≫ e.unit.app Y := + (NatIso.naturality_1 e.unitIso f).symm +#align category_theory.equivalence.inv_fun_map CategoryTheory.Equivalence.inv_fun_map + +section + +-- In this section we convert an arbitrary equivalence to a half-adjoint equivalence. +variable {F : C ⥤ D} {G : D ⥤ C} (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) + +/-- If `η : 𝟭 C ≅ F ⋙ G` is part of a (not necessarily half-adjoint) equivalence, we can upgrade it +to a refined natural isomorphism `adjointify_η η : 𝟭 C ≅ F ⋙ G` which exhibits the properties +required for a half-adjoint equivalence. See `equivalence.mk`. -/ +def adjointifyη : 𝟭 C ≅ F ⋙ G := + calc + 𝟭 C ≅ F ⋙ G := η + _ ≅ F ⋙ 𝟭 D ⋙ G := isoWhiskerLeft F (leftUnitor G).symm + _ ≅ F ⋙ (G ⋙ F) ⋙ G := isoWhiskerLeft F (isoWhiskerRight ε.symm G) + _ ≅ F ⋙ G ⋙ F ⋙ G := isoWhiskerLeft F (associator G F G) + _ ≅ (F ⋙ G) ⋙ F ⋙ G := (associator F G (F ⋙ G)).symm + _ ≅ 𝟭 C ⋙ F ⋙ G := isoWhiskerRight η.symm (F ⋙ G) + _ ≅ F ⋙ G := leftUnitor (F ⋙ G) + +#align category_theory.equivalence.adjointify_η CategoryTheory.Equivalence.adjointifyη + +theorem adjointify_η_ε (X : C) : + F.map ((adjointifyη η ε).hom.app X) ≫ ε.hom.app (F.obj X) = 𝟙 (F.obj X) := + by + dsimp [adjointifyη]; simp + have := ε.hom.naturality (F.map (η.inv.app X)); dsimp at this; rw [this]; clear this + rw [← assoc _ _ (F.map _)] + have := ε.hom.naturality (ε.inv.app <| F.obj X); dsimp at this; rw [this]; clear this + have := (ε.app <| F.obj X).hom_inv_id; dsimp at this; rw [this]; clear this + rw [id_comp]; have := (F.map_iso <| η.app X).hom_inv_id; dsimp at this; rw [this] +#align category_theory.equivalence.adjointify_η_ε CategoryTheory.Equivalence.adjointify_η_ε + +end + +/-- Every equivalence of categories consisting of functors `F` and `G` such that `F ⋙ G` and + `G ⋙ F` are naturally isomorphic to identity functors can be transformed into a half-adjoint + equivalence without changing `F` or `G`. -/ +protected def mk (F : C ⥤ D) (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : C ≌ D := + ⟨F, G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ +#align category_theory.equivalence.mk CategoryTheory.Equivalence.mk + +/-- Equivalence of categories is reflexive. -/ +@[refl, simps] +def refl : C ≌ C := + ⟨𝟭 C, 𝟭 C, Iso.refl _, Iso.refl _, fun X => Category.id_comp _⟩ +#align category_theory.equivalence.refl CategoryTheory.Equivalence.refl + +instance : Inhabited (C ≌ C) := + ⟨refl⟩ + +/-- Equivalence of categories is symmetric. -/ +@[symm, simps] +def symm (e : C ≌ D) : D ≌ C := + ⟨e.inverse, e.functor, e.counitIso.symm, e.unitIso.symm, e.inverse_counit_inv_comp⟩ +#align category_theory.equivalence.symm CategoryTheory.Equivalence.symm + +variable {E : Type u₃} [Category.{v₃} E] + +/-- Equivalence of categories is transitive. -/ +@[trans, simps] +def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E + where + functor := e.functor ⋙ f.functor + inverse := f.inverse ⋙ e.inverse + unitIso := by + refine' Iso.trans e.unitIso _ + exact isoWhiskerLeft e.functor (isoWhiskerRight f.unitIso e.inverse) + counitIso := by + refine' Iso.trans _ f.counitIso + exact isoWhiskerLeft f.inverse (isoWhiskerRight e.counitIso f.functor) + -- We wouldn't have needed to give this proof if we'd used `equivalence.mk`, + -- but we choose to avoid using that here, for the sake of good structure projection `simp` + -- lemmas. + functor_unit_iso_comp' X := by + dsimp + rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counit_inv_app_functor, fun_inv_map, + Iso.inv_hom_id_app_assoc, assoc, Iso.inv_hom_id_app, counit_app_functor, ← Functor.map_comp] + erw [comp_id, Iso.hom_inv_id_app, Functor.map_id] +#align category_theory.equivalence.trans CategoryTheory.Equivalence.trans + +/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic +functor. -/ +def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F ≅ F := + (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.unitIso.symm F ≪≫ F.leftUnitor +#align category_theory.equivalence.fun_inv_id_assoc CategoryTheory.Equivalence.funInvIdAssoc + +@[simp] +theorem fun_inv_id_assoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : + (funInvIdAssoc e F).hom.app X = F.map (e.unitInv.app X) := + by + dsimp [funInvIdAssoc] + aesop_cat +#align + category_theory.equivalence.fun_inv_id_assoc_hom_app CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app + +@[simp] +theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : + (funInvIdAssoc e F).inv.app X = F.map (e.unit.app X) := + by + dsimp [funInvIdAssoc] + aesop_cat +#align + category_theory.equivalence.fun_inv_id_assoc_inv_app CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app + +/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic +functor. -/ +def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.functor ⋙ F ≅ F := + (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.counitIso F ≪≫ F.leftUnitor +#align category_theory.equivalence.inv_fun_id_assoc CategoryTheory.Equivalence.invFunIdAssoc + +@[simp] +theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : + (invFunIdAssoc e F).hom.app X = F.map (e.counit.app X) := + by + dsimp [invFunIdAssoc] + aesop_cat +#align + category_theory.equivalence.inv_fun_id_assoc_hom_app CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app + +@[simp] +theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : + (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := + by + dsimp [invFunIdAssoc] + aesop_cat +#align + category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app + +/-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ +@[simps functor inverse unitIso counitIso] +def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := + Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.functor) + (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by aesop_cat)) + (NatIso.ofComponents (fun F => e.invFunIdAssoc F) (by aesop_cat)) +#align category_theory.equivalence.congr_left CategoryTheory.Equivalence.congrLeft + +/-- If `C` is equivalent to `D`, then `E ⥤ C` is equivalent to `E ⥤ D`. -/ +@[simps functor inverse unitIso counitIso] +def congrRight (e : C ≌ D) : E ⥤ C ≌ E ⥤ D := + Equivalence.mk ((whiskeringRight _ _ _).obj e.functor) ((whiskeringRight _ _ _).obj e.inverse) + (NatIso.ofComponents + (fun F => F.rightUnitor.symm ≪≫ isoWhiskerLeft F e.unitIso ≪≫ Functor.associator _ _ _) + (by aesop_cat)) + (NatIso.ofComponents + (fun F => Functor.associator _ _ _ ≪≫ isoWhiskerLeft F e.counitIso ≪≫ F.rightUnitor) + (by aesop_cat)) +#align category_theory.equivalence.congr_right CategoryTheory.Equivalence.congrRight + +section CancellationLemmas + +variable (e : C ≌ D) + +/- We need special forms of `cancel_nat_iso_hom_right(_assoc)` and +`cancel_nat_iso_inv_right(_assoc)` for units and counits, because neither `simp` or `rw` will apply +those lemmas in this setting without providing `e.unit_iso` (or similar) as an explicit argument. +We also provide the lemmas for length four compositions, since they're occasionally useful. +(e.g. in proving that equivalences take monos to monos) -/ +@[simp] +theorem cancel_unit_right {X Y : C} (f f' : X ⟶ Y) : + f ≫ e.unit.app Y = f' ≫ e.unit.app Y ↔ f = f' := by simp only [cancel_mono] +#align category_theory.equivalence.cancel_unit_right CategoryTheory.Equivalence.cancel_unit_right + +@[simp] +theorem cancel_unit_inv_right {X Y : C} (f f' : X ⟶ e.inverse.obj (e.functor.obj Y)) : + f ≫ e.unitInv.app Y = f' ≫ e.unitInv.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_unit_inv_right CategoryTheory.Equivalence.cancel_unit_inv_right + +@[simp] +theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.functor.obj (e.inverse.obj Y)) : + f ≫ e.counit.app Y = f' ≫ e.counit.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_counit_right CategoryTheory.Equivalence.cancel_counit_right + +@[simp] +theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : + f ≫ e.counitInv.app Y = f' ≫ e.counitInv.app Y ↔ f = f' := by simp only [cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right CategoryTheory.Equivalence.cancel_counit_inv_right + +@[simp] +theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : + f ≫ g ≫ e.unit.app Y = f' ≫ g' ≫ e.unit.app Y ↔ f ≫ g = f' ≫ g' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc + +@[simp] +theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') + (g' : X' ⟶ Y) : f ≫ g ≫ e.counitInv.app Y = f' ≫ g' ≫ e.counitInv.app Y ↔ f ≫ g = f' ≫ g' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right_assoc CategoryTheory.Equivalence.cancel_counit_inv_right_assoc + +@[simp] +theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) + (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : + f ≫ g ≫ h ≫ e.unit.app Z = f' ≫ g' ≫ h' ≫ e.unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by + simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' + +@[simp] +theorem cancel_counit_inv_right_assoc' {W X X' Y Y' Z : D} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) + (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : + f ≫ g ≫ h ≫ e.counitInv.app Z = f' ≫ g' ≫ h' ≫ e.counitInv.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := + by simp only [← category.assoc, cancel_mono] +#align + category_theory.equivalence.cancel_counit_inv_right_assoc' CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' + +end CancellationLemmas + +section + +-- There's of course a monoid structure on `C ≌ C`, +-- but let's not encourage using it. +-- The power structure is nevertheless useful. +/-- Natural number powers of an auto-equivalence. Use `(^)` instead. -/ +def powNat (e : C ≌ C) : ℕ → (C ≌ C) + | 0 => Equivalence.refl + | 1 => e + | n + 2 => e.trans (powNat e (n + 1)) +#align category_theory.equivalence.pow_nat CategoryTheory.Equivalence.powNat + +/-- Powers of an auto-equivalence. Use `(^)` instead. -/ +def pow (e : C ≌ C) : ℤ → (C ≌ C) + | Int.ofNat n => e.powNat n + | Int.negSucc n => e.symm.powNat (n + 1) +#align category_theory.equivalence.pow CategoryTheory.Equivalence.pow + +instance : Pow (C ≌ C) ℤ := + ⟨pow⟩ + +@[simp] +theorem pow_zero (e : C ≌ C) : e ^ (0 : ℤ) = Equivalence.refl := + rfl +#align category_theory.equivalence.pow_zero CategoryTheory.Equivalence.pow_zero + +@[simp] +theorem pow_one (e : C ≌ C) : e ^ (1 : ℤ) = e := + rfl +#align category_theory.equivalence.pow_one CategoryTheory.Equivalence.pow_one + +@[simp] +theorem pow_neg_one (e : C ≌ C) : e ^ (-1 : ℤ) = e.symm := + rfl +#align category_theory.equivalence.pow_neg_one CategoryTheory.Equivalence.pow_neg_one + +-- TODO as necessary, add the natural isomorphisms `(e^a).trans e^b ≅ e^(a+b)`. +-- At this point, we haven't even defined the category of equivalences. +end + +end Equivalence + +/-- A functor that is part of a (half) adjoint equivalence -/ +class IsEquivalence (F : C ⥤ D) where mk' :: + inverse : D ⥤ C + unitIso : 𝟭 C ≅ F ⋙ inverse + counitIso : inverse ⋙ F ≅ 𝟭 D + functor_unit_iso_comp' : + ∀ X : C, + F.map ((unitIso.hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counitIso.hom.app (F.obj X) = + 𝟙 (F.obj X) := by + aesop_cat +#align category_theory.is_equivalence CategoryTheory.IsEquivalence + +restate_axiom IsEquivalence.functor_unit_iso_comp' + +attribute [simp, reassoc.1] is_equivalence.functor_unit_iso_comp + +namespace IsEquivalence + +instance ofEquivalence (F : C ≌ D) : IsEquivalence F.functor := + { F with } +#align category_theory.is_equivalence.of_equivalence CategoryTheory.IsEquivalence.ofEquivalence + +instance ofEquivalenceInverse (F : C ≌ D) : IsEquivalence F.inverse := + IsEquivalence.ofEquivalence F.symm +#align + category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.IsEquivalence.ofEquivalenceInverse + +open Equivalence + +/-- To see that a functor is an equivalence, it suffices to provide an inverse functor `G` such that + `F ⋙ G` and `G ⋙ F` are naturally isomorphic to identity functors. -/ +protected def mk {F : C ⥤ D} (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : IsEquivalence F := + ⟨G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ +#align category_theory.is_equivalence.mk CategoryTheory.IsEquivalence.mk + +end IsEquivalence + +namespace Functor + +/-- Interpret a functor that is an equivalence as an equivalence. -/ +def asEquivalence (F : C ⥤ D) [IsEquivalence F] : C ≌ D := + ⟨F, IsEquivalence.inverse F, IsEquivalence.unitIso, IsEquivalence.counitIso, + IsEquivalence.functor_unit_iso_comp⟩ +#align category_theory.functor.as_equivalence CategoryTheory.Functor.asEquivalence + +instance isEquivalenceRefl : IsEquivalence (𝟭 C) := + IsEquivalence.ofEquivalence Equivalence.refl +#align category_theory.functor.is_equivalence_refl CategoryTheory.Functor.isEquivalenceRefl + +/-- The inverse functor of a functor that is an equivalence. -/ +def inv (F : C ⥤ D) [IsEquivalence F] : D ⥤ C := + IsEquivalence.inverse F +#align category_theory.functor.inv CategoryTheory.Functor.inv + +instance isEquivalenceInv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv := + IsEquivalence.ofEquivalence F.asEquivalence.symm +#align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalenceInv + +@[simp] +theorem as_equivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.functor = F := + rfl +#align category_theory.functor.as_equivalence_functor CategoryTheory.Functor.as_equivalence_functor + +@[simp] +theorem as_equivalence_inverse (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.inverse = inv F := + rfl +#align category_theory.functor.as_equivalence_inverse CategoryTheory.Functor.as_equivalence_inverse + +@[simp] +theorem as_equivalence_unit {F : C ⥤ D} [IsEquivalence F] : + F.asEquivalence.unitIso = IsEquivalence.unitIso := + rfl +#align category_theory.functor.as_equivalence_unit CategoryTheory.Functor.as_equivalence_unit + +@[simp] +theorem as_equivalence_counit {F : C ⥤ D} [IsEquivalence F] : + F.asEquivalence.counitIso = IsEquivalence.counitIso := + rfl +#align category_theory.functor.as_equivalence_counit CategoryTheory.Functor.as_equivalence_counit + +@[simp] +theorem inv_inv (F : C ⥤ D) [IsEquivalence F] : inv (inv F) = F := + rfl +#align category_theory.functor.inv_inv CategoryTheory.Functor.inv_inv + +variable {E : Type u₃} [Category.{v₃} E] + +instance isEquivalenceTrans (F : C ⥤ D) (G : D ⥤ E) [IsEquivalence F] [IsEquivalence G] : + IsEquivalence (F ⋙ G) := + IsEquivalence.ofEquivalence (Equivalence.trans (asEquivalence F) (asEquivalence G)) +#align category_theory.functor.is_equivalence_trans CategoryTheory.Functor.isEquivalenceTrans + +end Functor + +namespace Equivalence + +@[simp] +theorem functor_inv (E : C ≌ D) : E.functor.inv = E.inverse := + rfl +#align category_theory.equivalence.functor_inv CategoryTheory.Equivalence.functor_inv + +@[simp] +theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.functor := + rfl +#align category_theory.equivalence.inverse_inv CategoryTheory.Equivalence.inverse_inv + +@[simp] +theorem functor_as_equivalence (E : C ≌ D) : E.functor.asEquivalence = E := + by + cases E + congr +#align + category_theory.equivalence.functor_as_equivalence CategoryTheory.Equivalence.functor_as_equivalence + +@[simp] +theorem inverse_as_equivalence (E : C ≌ D) : E.inverse.asEquivalence = E.symm := + by + cases E + congr +#align + category_theory.equivalence.inverse_as_equivalence CategoryTheory.Equivalence.inverse_as_equivalence + +end Equivalence + +namespace IsEquivalence + +@[simp] +theorem fun_inv_map (F : C ⥤ D) [IsEquivalence F] (X Y : D) (f : X ⟶ Y) : + F.map (F.inv.map f) = F.asEquivalence.counit.app X ≫ f ≫ F.asEquivalence.counitInv.app Y := + by + erw [NatIso.naturality_2] + rfl +#align category_theory.is_equivalence.fun_inv_map CategoryTheory.IsEquivalence.fun_inv_map + +@[simp] +theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : + F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.unit.app Y := + by + erw [NatIso.naturality_1] + rfl +#align category_theory.is_equivalence.inv_fun_map CategoryTheory.IsEquivalence.inv_fun_map + +/-- When a functor `F` is an equivalence of categories, and `G` is isomorphic to `F`, then +`G` is also an equivalence of categories. -/ +@[simps] +def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G + where + inverse := hF.inverse + unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) + counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso + functor_unit_iso_comp' X := by + dsimp [NatIso.hcomp] + erw [id_comp, F.map_id, comp_id] + apply (cancel_epi (e.hom.app X)).mp + sliceLHS 1 2 => rw [← e.hom.naturality] + sliceLHS 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] + simp only [NatTrans.id_app, id_comp, comp_id, F.map_comp, assoc] + erw [hF.counitIso.hom.naturality] + sliceLHS 1 2 => rw [functor_unit_iso_comp] + simp only [functor.id_map, id_comp] +#align category_theory.is_equivalence.of_iso CategoryTheory.IsEquivalence.ofIso + +/-- Compatibility of `of_iso` with the composition of isomorphisms of functors -/ +theorem of_iso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEquivalence F) : + ofIso e' (ofIso e hF) = ofIso (e ≪≫ e') hF := + by + dsimp [ofIso] + congr 1 <;> ext X <;> dsimp [NatIso.hcomp] + · simp only [id_comp, assoc, Functor.map_comp] + · simp only [Functor.map_id, comp_id, id_comp, assoc] +#align category_theory.is_equivalence.of_iso_trans CategoryTheory.IsEquivalence.of_iso_trans + +/-- Compatibility of `of_iso` with identity isomorphisms of functors -/ +theorem of_iso_refl (F : C ⥤ D) (hF : IsEquivalence F) : ofIso (Iso.refl F) hF = hF := + by + rcases hF with ⟨Finv, Funit, Fcounit, Fcomp⟩ + dsimp [ofIso] + congr 1 <;> ext X <;> dsimp [NatIso.hcomp] + · simp only [comp_id, map_id] + · simp only [id_comp, map_id] +#align category_theory.is_equivalence.of_iso_refl CategoryTheory.IsEquivalence.of_iso_refl + +/-- When `F` and `G` are two isomorphic functors, then `F` is an equivalence iff `G` is. -/ +@[simps] +def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence G + where + toFun := ofIso e + invFun := ofIso e.symm + left_inv hF := by rw [of_iso_trans, Iso.self_symm_id, of_iso_refl] + right_inv hF := by rw [of_iso_trans, Iso.symm_self_id, of_iso_refl] +#align category_theory.is_equivalence.equiv_of_iso CategoryTheory.IsEquivalence.equivOfIso + +/-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ +@[simp] +def cancelCompRight {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : IsEquivalence G) + (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence F := + ofIso (Functor.associator F G G.inv ≪≫ NatIso.hcomp (Iso.refl F) hG.unitIso.symm ≪≫ rightUnitor F) + (Functor.isEquivalenceTrans (F ⋙ G) G.inv) +#align category_theory.is_equivalence.cancel_comp_right CategoryTheory.IsEquivalence.cancelCompRight + +/-- If `F` and `F ⋙ G` are equivalence of categories, then `G` is also an equivalence. -/ +@[simp] +def cancelCompLeft {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hF : IsEquivalence F) + (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence G := + ofIso + ((Functor.associator F.inv F G).symm ≪≫ NatIso.hcomp hF.counitIso (Iso.refl G) ≪≫ leftUnitor G) + (Functor.isEquivalenceTrans F.inv (F ⋙ G)) +#align category_theory.is_equivalence.cancel_comp_left CategoryTheory.IsEquivalence.cancelCompLeft + +end IsEquivalence + +namespace Equivalence + +/-- An equivalence is essentially surjective. + +See . +-/ +theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := + ⟨fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩ +#align + category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.ess_surj_of_equivalence + +-- see Note [lower instance priority] +/-- An equivalence is faithful. + +See . +-/ +instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F where + map_injective X Y f g w := by + have p := congr_arg (@CategoryTheory.Functor.map _ _ _ _ F.inv _ _) w + simpa only [cancel_epi, cancel_mono, is_equivalence.inv_fun_map] using p +#align + category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithful_of_equivalence + +-- see Note [lower instance priority] +/-- An equivalence is full. + +See . +-/ +instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F + where + preimage X Y f := F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y + witness X Y f := + F.inv.map_injective <| by + simpa only [is_equivalence.inv_fun_map, assoc, iso.inv_hom_id_app_assoc, + iso.inv_hom_id_app] using comp_id _ +#align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.fullOfEquivalence + +@[simps] +private noncomputable def equivalenceInverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C + where + obj X := F.objPreimage X + map X Y f := F.preimage ((F.objObjPreimageIso X).Hom ≫ f ≫ (F.objObjPreimageIso Y).inv) + map_id X := by apply F.map_injective; aesop_cat + map_comp X Y Z f g := by apply F.map_injective <;> simp +#align + category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse + +/-- A functor which is full, faithful, and essentially surjective is an equivalence. + +See . +-/ +noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : + IsEquivalence F := + IsEquivalence.mk (equivalenceInverse F) + (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) + fun X Y f => by + apply F.map_injective + aesop_cat) + (NatIso.ofComponents F.objObjPreimageIso (by aesop_cat)) +#align + category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj + +@[simp] +theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : + e.functor.map f = e.functor.map g ↔ f = g := + ⟨fun h => e.functor.map_injective h, fun h => h ▸ rfl⟩ +#align + category_theory.equivalence.functor_map_inj_iff CategoryTheory.Equivalence.functor_map_inj_iff + +@[simp] +theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : + e.inverse.map f = e.inverse.map g ↔ f = g := + functor_map_inj_iff e.symm f g +#align + category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff + +instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) + where mem_essImage Y := ⟨e.symm Y, by simp⟩ +#align + category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor + +noncomputable instance inducedFunctorOfEquiv {C' : Type _} (e : C' ≃ D) : + IsEquivalence (inducedFunctor e) := + Equivalence.ofFullyFaithfullyEssSurj _ +#align + category_theory.equivalence.induced_functor_of_equiv CategoryTheory.Equivalence.inducedFunctorOfEquiv + +noncomputable instance fullyFaithfulToEssImage (F : C ⥤ D) [Full F] [Faithful F] : + IsEquivalence F.toEssImage := + ofFullyFaithfullyEssSurj F.toEssImage +#align + category_theory.equivalence.fully_faithful_to_ess_image CategoryTheory.Equivalence.fullyFaithfulToEssImage + +end Equivalence + +end CategoryTheory + diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 9a2648c1fb895..72fd2d376e88d 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -105,6 +105,7 @@ def rewriteTarget' (stx : Syntax) (symm : Bool) : TacticM Unit := do let mvarId' ← (← getMainGoal).replaceTargetEq r.eNew r.eqProof replaceMainGoal ([mvarId']) + /-- Run a tactic on all goals, and always succeeds. diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index 0eaf24c8dbe43..7c5932ca553d5 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -37,17 +37,17 @@ open Tactic variable [Monad m] [MonadExceptOf Exception m] def evalSlice (a b : Nat) : TacticM Unit := do - iterateUntilFailure do + iterateUntilFailure do ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) iterateRange (a - 1) (a - 1) do evalTactic (← `(conv| congr)) evalTactic (← `(tactic| rotate_left)) let k ← iterateUntilFailureCount <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) - let c := k+1+a-b + let c := k+1+a-b iterateRange c c <| evalTactic (← `(conv| congr)) iterateUntilFailure do - ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat @@ -62,6 +62,19 @@ macro_rules | `(tactic| sliceRHS $a $b => $seq) => `(tactic| conv => rhs; slice $a $b; ($seq:convSeq)) +variable (C : Type) [Category C] (X Y Z W U : C) +variable (f₁ f₂ : X ⟶ Y) (g : Y ⟶ Z) (h : Z ⟶ W) (l : W ⟶ U) +-- +-- example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ l := by +-- conv => +-- lhs +-- slice 1 4 +-- conv => +-- lhs +-- slice 1 1 +-- rw [h₁] +-- rfl + -- add_tactic_doc -- { Name := "slice" -- category := DocCategory.tactic diff --git a/lean-toolchain b/lean-toolchain index c76d115ed8958..04ab6c3b09caa 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2023-01-29 +leanprover/lean4:nightly-2023-02-03 From 92e65beac07e193b862f7f15a68bdb38af99da6f Mon Sep 17 00:00:00 2001 From: thorimur <68410468+thorimur@users.noreply.github.com> Date: Tue, 7 Feb 2023 16:09:09 -0500 Subject: [PATCH 06/66] test: create `slice.lean` test file --- Mathlib/Tactic/Slice.lean | 35 +++++++++++------------------------ test/slice.lean | 16 ++++++++++++++++ 2 files changed, 27 insertions(+), 24 deletions(-) create mode 100644 test/slice.lean diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index 7c5932ca553d5..b58b9a8df03b9 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -15,17 +15,17 @@ namespace Tactic variable [Monad m] [MonadExceptOf Exception m] -partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α) := do - try +partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α) := do + try let a ← tac - let l ← iterateUntilFailureWithResults tac - pure (a :: l) - catch _ => pure [] + let l ← iterateUntilFailureWithResults tac + pure (a :: l) + catch _ => pure [] #align tactic.repeat_with_results Tactic.iterateUntilFailureWithResults def iterateUntilFailureCount {α : Type} (tac : m α) : m ℕ := do let r ← iterateUntilFailureWithResults tac - return r.length + return r.length #align tactic.repeat_count Tactic.iterateUntilFailureCount end Tactic @@ -37,16 +37,16 @@ open Tactic variable [Monad m] [MonadExceptOf Exception m] def evalSlice (a b : Nat) : TacticM Unit := do - iterateUntilFailure do + iterateUntilFailure do ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) iterateRange (a - 1) (a - 1) do evalTactic (← `(conv| congr)) evalTactic (← `(tactic| rotate_left)) - let k ← iterateUntilFailureCount - <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) - let c := k+1+a-b + let k ← iterateUntilFailureCount + <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) + let c := k+1+a-b iterateRange c c <| evalTactic (← `(conv| congr)) - iterateUntilFailure do + iterateUntilFailure do ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat @@ -62,19 +62,6 @@ macro_rules | `(tactic| sliceRHS $a $b => $seq) => `(tactic| conv => rhs; slice $a $b; ($seq:convSeq)) -variable (C : Type) [Category C] (X Y Z W U : C) -variable (f₁ f₂ : X ⟶ Y) (g : Y ⟶ Z) (h : Z ⟶ W) (l : W ⟶ U) --- --- example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ l := by --- conv => --- lhs --- slice 1 4 --- conv => --- lhs --- slice 1 1 --- rw [h₁] --- rfl - -- add_tactic_doc -- { Name := "slice" -- category := DocCategory.tactic diff --git a/test/slice.lean b/test/slice.lean new file mode 100644 index 0000000000000..e1d8162dbfed6 --- /dev/null +++ b/test/slice.lean @@ -0,0 +1,16 @@ +import Mathlib.Tactic.Slice + +open CategoryTheory + +variable (C : Type) [Category C] (X Y Z W U : C) +variable (f₁ f₂ : X ⟶ Y) (g : Y ⟶ Z) (h : Z ⟶ W) (l : W ⟶ U) + +example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ l := by + conv => + lhs + slice 1 4 + conv => + lhs + slice 1 1 + rw [h₁] + rfl From eedfb6167e983bd14ec76e3cbead7978121cc103 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 09:36:08 -0500 Subject: [PATCH 07/66] remove equivalence.lean --- Mathlib/CategoryTheory/Equivalence.lean | 809 ---------------------- Mathlib/CategoryTheory/Functor/Basic.lean | 1 + Mathlib/Tactic/Core.lean | 1 - Mathlib/Tactic/Slice.lean | 19 +- 4 files changed, 11 insertions(+), 819 deletions(-) delete mode 100644 Mathlib/CategoryTheory/Equivalence.lean diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean deleted file mode 100644 index 673184c6cb4fb..0000000000000 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ /dev/null @@ -1,809 +0,0 @@ -/- -Copyright (c) 2017 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tim Baumann, Stephen Morgan, Scott Morrison, Floris van Doorn - -! This file was ported from Lean 3 source module category_theory.equivalence -! leanprover-community/mathlib commit 9aba7801eeecebb61f58a5763c2b6dd1b47dc6ef -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. --/ -import Mathlib.CategoryTheory.Functor.FullyFaithful -import Mathlib.CategoryTheory.FullSubcategory -import Mathlib.CategoryTheory.Whiskering -import Mathlib.CategoryTheory.EssentialImage -import Mathlib.Tactic.Slice --- TODO: remove -set_option autoImplicit false -/-! -# Equivalence of categories - -An equivalence of categories `C` and `D` is a pair of functors `F : C ⥤ D` and `G : D ⥤ C` such -that `η : 𝟭 C ≅ F ⋙ G` and `ε : G ⋙ F ≅ 𝟭 D`. In many situations, equivalences are a better -notion of "sameness" of categories than the stricter isomorphims of categories. - -Recall that one way to express that two functors `F : C ⥤ D` and `G : D ⥤ C` are adjoint is using -two natural transformations `η : 𝟭 C ⟶ F ⋙ G` and `ε : G ⋙ F ⟶ 𝟭 D`, called the unit and the -counit, such that the compositions `F ⟶ FGF ⟶ F` and `G ⟶ GFG ⟶ G` are the identity. Unfortunately, -it is not the case that the natural isomorphisms `η` and `ε` in the definition of an equivalence -automatically give an adjunction. However, it is true that -* if one of the two compositions is the identity, then so is the other, and -* given an equivalence of categories, it is always possible to refine `η` in such a way that the - identities are satisfied. - -For this reason, in mathlib we define an equivalence to be a "half-adjoint equivalence", which is -a tuple `(F, G, η, ε)` as in the first paragraph such that the composite `F ⟶ FGF ⟶ F` is the -identity. By the remark above, this already implies that the tuple is an "adjoint equivalence", -i.e., that the composite `G ⟶ GFG ⟶ G` is also the identity. - -We also define essentially surjective functors and show that a functor is an equivalence if and only -if it is full, faithful and essentially surjective. - -## Main definitions - -* `Equivalence`: bundled (half-)adjoint equivalences of categories -* `IsEquivalence`: type class on a functor `F` containing the data of the inverse `G` as well as - the natural isomorphisms `η` and `ε`. -* `EssSurj`: type class on a functor `F` containing the data of the preimages and the isomorphisms - `F.obj (preimage d) ≅ d`. - -## Main results - -* `Equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence -* `IsEquivalence.equiv_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence -iff `G` is. -* `Equivalence.of_fully_faithfully_essSurj`: a fully faithful essentially surjective functor is an - equivalence. - -## Notations - -We write `C ≌ D` (`\backcong`, not to be confused with `≅`/`\cong`) for a bundled equivalence. - --/ - -namespace CategoryTheory - -open CategoryTheory.Functor NatIso Category - --- declare the `v`'s first; see `CategoryTheory.Category` for an explanation -universe v₁ v₂ v₃ u₁ u₂ u₃ - -/-- We define an equivalence as a (half)-adjoint equivalence, a pair of functors with - a unit and counit which are natural isomorphisms and the triangle law `Fη ≫ εF = 1`, or in other - words the composite `F ⟶ FGF ⟶ F` is the identity. - - In `unit_inverse_comp`, we show that this is actually an adjoint equivalence, i.e., that the - composite `G ⟶ GFG ⟶ G` is also the identity. - - The triangle equation is written as a family of equalities between morphisms, it is more - complicated if we write it as an equality of natural transformations, because then we would have - to insert natural transformations like `F ⟶ F1`. - -See --/ -structure Equivalence (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] where mk' :: - functor : C ⥤ D - inverse : D ⥤ C - unitIso : 𝟭 C ≅ functor ⋙ inverse - counitIso : inverse ⋙ functor ≅ 𝟭 D - functor_unit_iso_comp' : ∀ X : C, functor.map (unitIso.hom.app X) ≫ counitIso.hom.app (functor.obj X) = 𝟙 (functor.obj X) := by aesop_cat -#align category_theory.equivalence CategoryTheory.Equivalence - -restate_axiom Equivalence.functor_unit_iso_comp' - --- mathport name: «expr ≌ » -infixr:10 " ≌ " => Equivalence - -variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] - -namespace Equivalence - -/-- The unit of an equivalence of categories. -/ -abbrev unit (e : C ≌ D) : 𝟭 C ⟶ e.functor ⋙ e.inverse := - e.unitIso.hom -#align category_theory.equivalence.unit CategoryTheory.Equivalence.unit - -/-- The counit of an equivalence of categories. -/ -abbrev counit (e : C ≌ D) : e.inverse ⋙ e.functor ⟶ 𝟭 D := - e.counitIso.hom -#align category_theory.equivalence.counit CategoryTheory.Equivalence.counit - -/-- The inverse of the unit of an equivalence of categories. -/ -abbrev unitInv (e : C ≌ D) : e.functor ⋙ e.inverse ⟶ 𝟭 C := - e.unitIso.inv -#align category_theory.equivalence.unit_inv CategoryTheory.Equivalence.unitInv - -/-- The inverse of the counit of an equivalence of categories. -/ -abbrev counitInv (e : C ≌ D) : 𝟭 D ⟶ e.inverse ⋙ e.functor := - e.counitIso.inv -#align category_theory.equivalence.counit_inv CategoryTheory.Equivalence.counitInv - -/- While these abbreviations are convenient, they also cause some trouble, -preventing structure projections from unfolding. -/ -@[simp] -theorem equivalence_mk'_unit (functor inverse unit_iso counit_iso f) : - (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unit = unit_iso.hom := - rfl -#align - category_theory.equivalence.equivalence_mk'_unit CategoryTheory.Equivalence.equivalence_mk'_unit - -@[simp] -theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : - (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.hom := - rfl -#align - category_theory.equivalence.equivalence_mk'_counit CategoryTheory.Equivalence.equivalence_mk'_counit - -@[simp] -theorem equivalence_mk'_unitInv (functor inverse unit_iso counit_iso f) : - (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := - rfl -#align - category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unitInv - -@[simp] -theorem equivalence_mk'_counitInv (functor inverse unit_iso counit_iso f) : - (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := - rfl -#align - category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counitInv - -@[simp] -theorem functor_unit_comp (e : C ≌ D) (X : C) : - e.functor.map (e.unit.app X) ≫ e.counit.app (e.functor.obj X) = 𝟙 (e.functor.obj X) := - e.functor_unit_iso_comp X -#align category_theory.equivalence.functor_unit_comp CategoryTheory.Equivalence.functor_unit_comp - -@[simp] -theorem counitInv_functor_comp (e : C ≌ D) (X : C) : - e.counitInv.app (e.functor.obj X) ≫ e.functor.map (e.unitInv.app X) = 𝟙 (e.functor.obj X) := - by - erw [Iso.inv_eq_inv (e.functor.mapIso (e.unitIso.app X) ≪≫ e.counitIso.app (e.functor.obj X)) - (Iso.refl _)] - exact e.functor_unit_comp X -#align - category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counitInv_functor_comp - -theorem counit_inv_app_functor (e : C ≌ D) (X : C) : - e.counitInv.app (e.functor.obj X) = e.functor.map (e.unit.app X) := - by - symm - erw [← Iso.comp_hom_eq_id (e.counitIso.app _), functor_unit_comp] - rfl -#align - category_theory.equivalence.counit_inv_app_functor CategoryTheory.Equivalence.counit_inv_app_functor - -theorem counit_app_functor (e : C ≌ D) (X : C) : - e.counit.app (e.functor.obj X) = e.functor.map (e.unitInv.app X) := - by - erw [← Iso.hom_comp_eq_id (e.functor.mapIso (e.unitIso.app X)), functor_unit_comp] - rfl -#align category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor - -/-- The other triangle equality. The proof follows the following proof in Globular: - http://globular.science/1905.001 -/ -@[simp] -theorem unit_inverse_comp (e : C ≌ D) (Y : D) : - e.unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := - by - rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counitInv_functor_comp, map_comp] - dsimp - rw [← Iso.hom_inv_id_assoc (e.unitIso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] - sliceLHS 2 3 => erw [e.unit.naturality] - sliceLHS 1 2 => erw [e.unit.naturality] - sliceLHS 4 4 => - rw [← Iso.hom_inv_id_assoc (e.inverse.mapIso (e.counitIso.app _)) (e.unitInv.app _)] - sliceLHS 3 4 => - erw [← map_comp e.inverse, e.counit.naturality] - erw [(e.counitIso.app _).hom_inv_id, map_id] - erw [id_comp] - sliceLHS 2 3 => erw [← map_comp e.inverse, e.counitIso.inv.naturality, map_comp] - sliceLHS 3 4 => erw [e.unitInv.naturality] - sliceLHS 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unitIso.app _).hom_inv_id, map_id] - erw [id_comp] - sliceLHS 3 4 => erw [← e.unitInv.naturality] - sliceLHS 2 3 => - erw [← map_comp e.inverse, ← e.counitIso.inv.naturality, (e.counitIso.app _).hom_inv_id, - map_id] - erw [id_comp, (e.unitIso.app _).hom_inv_id]; rfl - rfl -#align category_theory.equivalence.unit_inverse_comp CategoryTheory.Equivalence.unit_inverse_comp - -@[simp] -theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : - e.inverse.map (e.counitInv.app Y) ≫ e.unitInv.app (e.inverse.obj Y) = 𝟙 (e.inverse.obj Y) := - by - erw [Iso.inv_eq_inv (e.unitIso.app (e.inverse.obj Y) ≪≫ e.inverse.mapIso (e.counitIso.app Y)) - (Iso.refl _)] - exact e.unit_inverse_comp Y -#align - category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp - -theorem unit_app_inverse (e : C ≌ D) (Y : D) : - e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := - by - erw [← Iso.comp_hom_eq_id (e.inverse.mapIso (e.counitIso.app Y)), unit_inverse_comp] -#align category_theory.equivalence.unit_app_inverse CategoryTheory.Equivalence.unit_app_inverse - -theorem unit_inv_app_inverse (e : C ≌ D) (Y : D) : - e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := - by - symm - erw [← Iso.hom_comp_eq_id (e.unitIso.app _), unit_inverse_comp] - rfl -#align - category_theory.equivalence.unit_inv_app_inverse CategoryTheory.Equivalence.unit_inv_app_inverse - -@[simp] -theorem fun_inv_map (e : C ≌ D) (X Y : D) (f : X ⟶ Y) : - e.functor.map (e.inverse.map f) = e.counit.app X ≫ f ≫ e.counitInv.app Y := - (NatIso.naturality_2 e.counitIso f).symm -#align category_theory.equivalence.fun_inv_map CategoryTheory.Equivalence.fun_inv_map - -@[simp] -theorem inv_fun_map (e : C ≌ D) (X Y : C) (f : X ⟶ Y) : - e.inverse.map (e.functor.map f) = e.unitInv.app X ≫ f ≫ e.unit.app Y := - (NatIso.naturality_1 e.unitIso f).symm -#align category_theory.equivalence.inv_fun_map CategoryTheory.Equivalence.inv_fun_map - -section - --- In this section we convert an arbitrary equivalence to a half-adjoint equivalence. -variable {F : C ⥤ D} {G : D ⥤ C} (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) - -/-- If `η : 𝟭 C ≅ F ⋙ G` is part of a (not necessarily half-adjoint) equivalence, we can upgrade it -to a refined natural isomorphism `adjointify_η η : 𝟭 C ≅ F ⋙ G` which exhibits the properties -required for a half-adjoint equivalence. See `equivalence.mk`. -/ -def adjointifyη : 𝟭 C ≅ F ⋙ G := - calc - 𝟭 C ≅ F ⋙ G := η - _ ≅ F ⋙ 𝟭 D ⋙ G := isoWhiskerLeft F (leftUnitor G).symm - _ ≅ F ⋙ (G ⋙ F) ⋙ G := isoWhiskerLeft F (isoWhiskerRight ε.symm G) - _ ≅ F ⋙ G ⋙ F ⋙ G := isoWhiskerLeft F (associator G F G) - _ ≅ (F ⋙ G) ⋙ F ⋙ G := (associator F G (F ⋙ G)).symm - _ ≅ 𝟭 C ⋙ F ⋙ G := isoWhiskerRight η.symm (F ⋙ G) - _ ≅ F ⋙ G := leftUnitor (F ⋙ G) - -#align category_theory.equivalence.adjointify_η CategoryTheory.Equivalence.adjointifyη - -theorem adjointify_η_ε (X : C) : - F.map ((adjointifyη η ε).hom.app X) ≫ ε.hom.app (F.obj X) = 𝟙 (F.obj X) := - by - dsimp [adjointifyη]; simp - have := ε.hom.naturality (F.map (η.inv.app X)); dsimp at this; rw [this]; clear this - rw [← assoc _ _ (F.map _)] - have := ε.hom.naturality (ε.inv.app <| F.obj X); dsimp at this; rw [this]; clear this - have := (ε.app <| F.obj X).hom_inv_id; dsimp at this; rw [this]; clear this - rw [id_comp]; have := (F.map_iso <| η.app X).hom_inv_id; dsimp at this; rw [this] -#align category_theory.equivalence.adjointify_η_ε CategoryTheory.Equivalence.adjointify_η_ε - -end - -/-- Every equivalence of categories consisting of functors `F` and `G` such that `F ⋙ G` and - `G ⋙ F` are naturally isomorphic to identity functors can be transformed into a half-adjoint - equivalence without changing `F` or `G`. -/ -protected def mk (F : C ⥤ D) (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : C ≌ D := - ⟨F, G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ -#align category_theory.equivalence.mk CategoryTheory.Equivalence.mk - -/-- Equivalence of categories is reflexive. -/ -@[refl, simps] -def refl : C ≌ C := - ⟨𝟭 C, 𝟭 C, Iso.refl _, Iso.refl _, fun X => Category.id_comp _⟩ -#align category_theory.equivalence.refl CategoryTheory.Equivalence.refl - -instance : Inhabited (C ≌ C) := - ⟨refl⟩ - -/-- Equivalence of categories is symmetric. -/ -@[symm, simps] -def symm (e : C ≌ D) : D ≌ C := - ⟨e.inverse, e.functor, e.counitIso.symm, e.unitIso.symm, e.inverse_counit_inv_comp⟩ -#align category_theory.equivalence.symm CategoryTheory.Equivalence.symm - -variable {E : Type u₃} [Category.{v₃} E] - -/-- Equivalence of categories is transitive. -/ -@[trans, simps] -def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E - where - functor := e.functor ⋙ f.functor - inverse := f.inverse ⋙ e.inverse - unitIso := by - refine' Iso.trans e.unitIso _ - exact isoWhiskerLeft e.functor (isoWhiskerRight f.unitIso e.inverse) - counitIso := by - refine' Iso.trans _ f.counitIso - exact isoWhiskerLeft f.inverse (isoWhiskerRight e.counitIso f.functor) - -- We wouldn't have needed to give this proof if we'd used `equivalence.mk`, - -- but we choose to avoid using that here, for the sake of good structure projection `simp` - -- lemmas. - functor_unit_iso_comp' X := by - dsimp - rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counit_inv_app_functor, fun_inv_map, - Iso.inv_hom_id_app_assoc, assoc, Iso.inv_hom_id_app, counit_app_functor, ← Functor.map_comp] - erw [comp_id, Iso.hom_inv_id_app, Functor.map_id] -#align category_theory.equivalence.trans CategoryTheory.Equivalence.trans - -/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic -functor. -/ -def funInvIdAssoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F ≅ F := - (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.unitIso.symm F ≪≫ F.leftUnitor -#align category_theory.equivalence.fun_inv_id_assoc CategoryTheory.Equivalence.funInvIdAssoc - -@[simp] -theorem fun_inv_id_assoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : - (funInvIdAssoc e F).hom.app X = F.map (e.unitInv.app X) := - by - dsimp [funInvIdAssoc] - aesop_cat -#align - category_theory.equivalence.fun_inv_id_assoc_hom_app CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app - -@[simp] -theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : - (funInvIdAssoc e F).inv.app X = F.map (e.unit.app X) := - by - dsimp [funInvIdAssoc] - aesop_cat -#align - category_theory.equivalence.fun_inv_id_assoc_inv_app CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app - -/-- Composing a functor with both functors of an equivalence yields a naturally isomorphic -functor. -/ -def invFunIdAssoc (e : C ≌ D) (F : D ⥤ E) : e.inverse ⋙ e.functor ⋙ F ≅ F := - (Functor.associator _ _ _).symm ≪≫ isoWhiskerRight e.counitIso F ≪≫ F.leftUnitor -#align category_theory.equivalence.inv_fun_id_assoc CategoryTheory.Equivalence.invFunIdAssoc - -@[simp] -theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : - (invFunIdAssoc e F).hom.app X = F.map (e.counit.app X) := - by - dsimp [invFunIdAssoc] - aesop_cat -#align - category_theory.equivalence.inv_fun_id_assoc_hom_app CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app - -@[simp] -theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : - (invFunIdAssoc e F).inv.app X = F.map (e.counitInv.app X) := - by - dsimp [invFunIdAssoc] - aesop_cat -#align - category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app - -/-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ -@[simps functor inverse unitIso counitIso] -def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := - Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.functor) - (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by aesop_cat)) - (NatIso.ofComponents (fun F => e.invFunIdAssoc F) (by aesop_cat)) -#align category_theory.equivalence.congr_left CategoryTheory.Equivalence.congrLeft - -/-- If `C` is equivalent to `D`, then `E ⥤ C` is equivalent to `E ⥤ D`. -/ -@[simps functor inverse unitIso counitIso] -def congrRight (e : C ≌ D) : E ⥤ C ≌ E ⥤ D := - Equivalence.mk ((whiskeringRight _ _ _).obj e.functor) ((whiskeringRight _ _ _).obj e.inverse) - (NatIso.ofComponents - (fun F => F.rightUnitor.symm ≪≫ isoWhiskerLeft F e.unitIso ≪≫ Functor.associator _ _ _) - (by aesop_cat)) - (NatIso.ofComponents - (fun F => Functor.associator _ _ _ ≪≫ isoWhiskerLeft F e.counitIso ≪≫ F.rightUnitor) - (by aesop_cat)) -#align category_theory.equivalence.congr_right CategoryTheory.Equivalence.congrRight - -section CancellationLemmas - -variable (e : C ≌ D) - -/- We need special forms of `cancel_nat_iso_hom_right(_assoc)` and -`cancel_nat_iso_inv_right(_assoc)` for units and counits, because neither `simp` or `rw` will apply -those lemmas in this setting without providing `e.unit_iso` (or similar) as an explicit argument. -We also provide the lemmas for length four compositions, since they're occasionally useful. -(e.g. in proving that equivalences take monos to monos) -/ -@[simp] -theorem cancel_unit_right {X Y : C} (f f' : X ⟶ Y) : - f ≫ e.unit.app Y = f' ≫ e.unit.app Y ↔ f = f' := by simp only [cancel_mono] -#align category_theory.equivalence.cancel_unit_right CategoryTheory.Equivalence.cancel_unit_right - -@[simp] -theorem cancel_unit_inv_right {X Y : C} (f f' : X ⟶ e.inverse.obj (e.functor.obj Y)) : - f ≫ e.unitInv.app Y = f' ≫ e.unitInv.app Y ↔ f = f' := by simp only [cancel_mono] -#align - category_theory.equivalence.cancel_unit_inv_right CategoryTheory.Equivalence.cancel_unit_inv_right - -@[simp] -theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.functor.obj (e.inverse.obj Y)) : - f ≫ e.counit.app Y = f' ≫ e.counit.app Y ↔ f = f' := by simp only [cancel_mono] -#align - category_theory.equivalence.cancel_counit_right CategoryTheory.Equivalence.cancel_counit_right - -@[simp] -theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : - f ≫ e.counitInv.app Y = f' ≫ e.counitInv.app Y ↔ f = f' := by simp only [cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right CategoryTheory.Equivalence.cancel_counit_inv_right - -@[simp] -theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : - f ≫ g ≫ e.unit.app Y = f' ≫ g' ≫ e.unit.app Y ↔ f ≫ g = f' ≫ g' := by - simp only [← category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc - -@[simp] -theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') - (g' : X' ⟶ Y) : f ≫ g ≫ e.counitInv.app Y = f' ≫ g' ≫ e.counitInv.app Y ↔ f ≫ g = f' ≫ g' := by - simp only [← category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right_assoc CategoryTheory.Equivalence.cancel_counit_inv_right_assoc - -@[simp] -theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) - (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : - f ≫ g ≫ h ≫ e.unit.app Z = f' ≫ g' ≫ h' ≫ e.unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by - simp only [← category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' - -@[simp] -theorem cancel_counit_inv_right_assoc' {W X X' Y Y' Z : D} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) - (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : - f ≫ g ≫ h ≫ e.counitInv.app Z = f' ≫ g' ≫ h' ≫ e.counitInv.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := - by simp only [← category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right_assoc' CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' - -end CancellationLemmas - -section - --- There's of course a monoid structure on `C ≌ C`, --- but let's not encourage using it. --- The power structure is nevertheless useful. -/-- Natural number powers of an auto-equivalence. Use `(^)` instead. -/ -def powNat (e : C ≌ C) : ℕ → (C ≌ C) - | 0 => Equivalence.refl - | 1 => e - | n + 2 => e.trans (powNat e (n + 1)) -#align category_theory.equivalence.pow_nat CategoryTheory.Equivalence.powNat - -/-- Powers of an auto-equivalence. Use `(^)` instead. -/ -def pow (e : C ≌ C) : ℤ → (C ≌ C) - | Int.ofNat n => e.powNat n - | Int.negSucc n => e.symm.powNat (n + 1) -#align category_theory.equivalence.pow CategoryTheory.Equivalence.pow - -instance : Pow (C ≌ C) ℤ := - ⟨pow⟩ - -@[simp] -theorem pow_zero (e : C ≌ C) : e ^ (0 : ℤ) = Equivalence.refl := - rfl -#align category_theory.equivalence.pow_zero CategoryTheory.Equivalence.pow_zero - -@[simp] -theorem pow_one (e : C ≌ C) : e ^ (1 : ℤ) = e := - rfl -#align category_theory.equivalence.pow_one CategoryTheory.Equivalence.pow_one - -@[simp] -theorem pow_neg_one (e : C ≌ C) : e ^ (-1 : ℤ) = e.symm := - rfl -#align category_theory.equivalence.pow_neg_one CategoryTheory.Equivalence.pow_neg_one - --- TODO as necessary, add the natural isomorphisms `(e^a).trans e^b ≅ e^(a+b)`. --- At this point, we haven't even defined the category of equivalences. -end - -end Equivalence - -/-- A functor that is part of a (half) adjoint equivalence -/ -class IsEquivalence (F : C ⥤ D) where mk' :: - inverse : D ⥤ C - unitIso : 𝟭 C ≅ F ⋙ inverse - counitIso : inverse ⋙ F ≅ 𝟭 D - functor_unit_iso_comp' : - ∀ X : C, - F.map ((unitIso.hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counitIso.hom.app (F.obj X) = - 𝟙 (F.obj X) := by - aesop_cat -#align category_theory.is_equivalence CategoryTheory.IsEquivalence - -restate_axiom IsEquivalence.functor_unit_iso_comp' - -attribute [simp, reassoc.1] is_equivalence.functor_unit_iso_comp - -namespace IsEquivalence - -instance ofEquivalence (F : C ≌ D) : IsEquivalence F.functor := - { F with } -#align category_theory.is_equivalence.of_equivalence CategoryTheory.IsEquivalence.ofEquivalence - -instance ofEquivalenceInverse (F : C ≌ D) : IsEquivalence F.inverse := - IsEquivalence.ofEquivalence F.symm -#align - category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.IsEquivalence.ofEquivalenceInverse - -open Equivalence - -/-- To see that a functor is an equivalence, it suffices to provide an inverse functor `G` such that - `F ⋙ G` and `G ⋙ F` are naturally isomorphic to identity functors. -/ -protected def mk {F : C ⥤ D} (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : IsEquivalence F := - ⟨G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ -#align category_theory.is_equivalence.mk CategoryTheory.IsEquivalence.mk - -end IsEquivalence - -namespace Functor - -/-- Interpret a functor that is an equivalence as an equivalence. -/ -def asEquivalence (F : C ⥤ D) [IsEquivalence F] : C ≌ D := - ⟨F, IsEquivalence.inverse F, IsEquivalence.unitIso, IsEquivalence.counitIso, - IsEquivalence.functor_unit_iso_comp⟩ -#align category_theory.functor.as_equivalence CategoryTheory.Functor.asEquivalence - -instance isEquivalenceRefl : IsEquivalence (𝟭 C) := - IsEquivalence.ofEquivalence Equivalence.refl -#align category_theory.functor.is_equivalence_refl CategoryTheory.Functor.isEquivalenceRefl - -/-- The inverse functor of a functor that is an equivalence. -/ -def inv (F : C ⥤ D) [IsEquivalence F] : D ⥤ C := - IsEquivalence.inverse F -#align category_theory.functor.inv CategoryTheory.Functor.inv - -instance isEquivalenceInv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv := - IsEquivalence.ofEquivalence F.asEquivalence.symm -#align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalenceInv - -@[simp] -theorem as_equivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.functor = F := - rfl -#align category_theory.functor.as_equivalence_functor CategoryTheory.Functor.as_equivalence_functor - -@[simp] -theorem as_equivalence_inverse (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.inverse = inv F := - rfl -#align category_theory.functor.as_equivalence_inverse CategoryTheory.Functor.as_equivalence_inverse - -@[simp] -theorem as_equivalence_unit {F : C ⥤ D} [IsEquivalence F] : - F.asEquivalence.unitIso = IsEquivalence.unitIso := - rfl -#align category_theory.functor.as_equivalence_unit CategoryTheory.Functor.as_equivalence_unit - -@[simp] -theorem as_equivalence_counit {F : C ⥤ D} [IsEquivalence F] : - F.asEquivalence.counitIso = IsEquivalence.counitIso := - rfl -#align category_theory.functor.as_equivalence_counit CategoryTheory.Functor.as_equivalence_counit - -@[simp] -theorem inv_inv (F : C ⥤ D) [IsEquivalence F] : inv (inv F) = F := - rfl -#align category_theory.functor.inv_inv CategoryTheory.Functor.inv_inv - -variable {E : Type u₃} [Category.{v₃} E] - -instance isEquivalenceTrans (F : C ⥤ D) (G : D ⥤ E) [IsEquivalence F] [IsEquivalence G] : - IsEquivalence (F ⋙ G) := - IsEquivalence.ofEquivalence (Equivalence.trans (asEquivalence F) (asEquivalence G)) -#align category_theory.functor.is_equivalence_trans CategoryTheory.Functor.isEquivalenceTrans - -end Functor - -namespace Equivalence - -@[simp] -theorem functor_inv (E : C ≌ D) : E.functor.inv = E.inverse := - rfl -#align category_theory.equivalence.functor_inv CategoryTheory.Equivalence.functor_inv - -@[simp] -theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.functor := - rfl -#align category_theory.equivalence.inverse_inv CategoryTheory.Equivalence.inverse_inv - -@[simp] -theorem functor_as_equivalence (E : C ≌ D) : E.functor.asEquivalence = E := - by - cases E - congr -#align - category_theory.equivalence.functor_as_equivalence CategoryTheory.Equivalence.functor_as_equivalence - -@[simp] -theorem inverse_as_equivalence (E : C ≌ D) : E.inverse.asEquivalence = E.symm := - by - cases E - congr -#align - category_theory.equivalence.inverse_as_equivalence CategoryTheory.Equivalence.inverse_as_equivalence - -end Equivalence - -namespace IsEquivalence - -@[simp] -theorem fun_inv_map (F : C ⥤ D) [IsEquivalence F] (X Y : D) (f : X ⟶ Y) : - F.map (F.inv.map f) = F.asEquivalence.counit.app X ≫ f ≫ F.asEquivalence.counitInv.app Y := - by - erw [NatIso.naturality_2] - rfl -#align category_theory.is_equivalence.fun_inv_map CategoryTheory.IsEquivalence.fun_inv_map - -@[simp] -theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : - F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.unit.app Y := - by - erw [NatIso.naturality_1] - rfl -#align category_theory.is_equivalence.inv_fun_map CategoryTheory.IsEquivalence.inv_fun_map - -/-- When a functor `F` is an equivalence of categories, and `G` is isomorphic to `F`, then -`G` is also an equivalence of categories. -/ -@[simps] -def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G - where - inverse := hF.inverse - unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) - counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso - functor_unit_iso_comp' X := by - dsimp [NatIso.hcomp] - erw [id_comp, F.map_id, comp_id] - apply (cancel_epi (e.hom.app X)).mp - sliceLHS 1 2 => rw [← e.hom.naturality] - sliceLHS 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] - simp only [NatTrans.id_app, id_comp, comp_id, F.map_comp, assoc] - erw [hF.counitIso.hom.naturality] - sliceLHS 1 2 => rw [functor_unit_iso_comp] - simp only [functor.id_map, id_comp] -#align category_theory.is_equivalence.of_iso CategoryTheory.IsEquivalence.ofIso - -/-- Compatibility of `of_iso` with the composition of isomorphisms of functors -/ -theorem of_iso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEquivalence F) : - ofIso e' (ofIso e hF) = ofIso (e ≪≫ e') hF := - by - dsimp [ofIso] - congr 1 <;> ext X <;> dsimp [NatIso.hcomp] - · simp only [id_comp, assoc, Functor.map_comp] - · simp only [Functor.map_id, comp_id, id_comp, assoc] -#align category_theory.is_equivalence.of_iso_trans CategoryTheory.IsEquivalence.of_iso_trans - -/-- Compatibility of `of_iso` with identity isomorphisms of functors -/ -theorem of_iso_refl (F : C ⥤ D) (hF : IsEquivalence F) : ofIso (Iso.refl F) hF = hF := - by - rcases hF with ⟨Finv, Funit, Fcounit, Fcomp⟩ - dsimp [ofIso] - congr 1 <;> ext X <;> dsimp [NatIso.hcomp] - · simp only [comp_id, map_id] - · simp only [id_comp, map_id] -#align category_theory.is_equivalence.of_iso_refl CategoryTheory.IsEquivalence.of_iso_refl - -/-- When `F` and `G` are two isomorphic functors, then `F` is an equivalence iff `G` is. -/ -@[simps] -def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence G - where - toFun := ofIso e - invFun := ofIso e.symm - left_inv hF := by rw [of_iso_trans, Iso.self_symm_id, of_iso_refl] - right_inv hF := by rw [of_iso_trans, Iso.symm_self_id, of_iso_refl] -#align category_theory.is_equivalence.equiv_of_iso CategoryTheory.IsEquivalence.equivOfIso - -/-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ -@[simp] -def cancelCompRight {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : IsEquivalence G) - (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence F := - ofIso (Functor.associator F G G.inv ≪≫ NatIso.hcomp (Iso.refl F) hG.unitIso.symm ≪≫ rightUnitor F) - (Functor.isEquivalenceTrans (F ⋙ G) G.inv) -#align category_theory.is_equivalence.cancel_comp_right CategoryTheory.IsEquivalence.cancelCompRight - -/-- If `F` and `F ⋙ G` are equivalence of categories, then `G` is also an equivalence. -/ -@[simp] -def cancelCompLeft {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hF : IsEquivalence F) - (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence G := - ofIso - ((Functor.associator F.inv F G).symm ≪≫ NatIso.hcomp hF.counitIso (Iso.refl G) ≪≫ leftUnitor G) - (Functor.isEquivalenceTrans F.inv (F ⋙ G)) -#align category_theory.is_equivalence.cancel_comp_left CategoryTheory.IsEquivalence.cancelCompLeft - -end IsEquivalence - -namespace Equivalence - -/-- An equivalence is essentially surjective. - -See . --/ -theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := - ⟨fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩ -#align - category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.ess_surj_of_equivalence - --- see Note [lower instance priority] -/-- An equivalence is faithful. - -See . --/ -instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F where - map_injective X Y f g w := by - have p := congr_arg (@CategoryTheory.Functor.map _ _ _ _ F.inv _ _) w - simpa only [cancel_epi, cancel_mono, is_equivalence.inv_fun_map] using p -#align - category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithful_of_equivalence - --- see Note [lower instance priority] -/-- An equivalence is full. - -See . --/ -instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F - where - preimage X Y f := F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y - witness X Y f := - F.inv.map_injective <| by - simpa only [is_equivalence.inv_fun_map, assoc, iso.inv_hom_id_app_assoc, - iso.inv_hom_id_app] using comp_id _ -#align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.fullOfEquivalence - -@[simps] -private noncomputable def equivalenceInverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C - where - obj X := F.objPreimage X - map X Y f := F.preimage ((F.objObjPreimageIso X).Hom ≫ f ≫ (F.objObjPreimageIso Y).inv) - map_id X := by apply F.map_injective; aesop_cat - map_comp X Y Z f g := by apply F.map_injective <;> simp -#align - category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse - -/-- A functor which is full, faithful, and essentially surjective is an equivalence. - -See . --/ -noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : - IsEquivalence F := - IsEquivalence.mk (equivalenceInverse F) - (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) - fun X Y f => by - apply F.map_injective - aesop_cat) - (NatIso.ofComponents F.objObjPreimageIso (by aesop_cat)) -#align - category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj - -@[simp] -theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : - e.functor.map f = e.functor.map g ↔ f = g := - ⟨fun h => e.functor.map_injective h, fun h => h ▸ rfl⟩ -#align - category_theory.equivalence.functor_map_inj_iff CategoryTheory.Equivalence.functor_map_inj_iff - -@[simp] -theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : - e.inverse.map f = e.inverse.map g ↔ f = g := - functor_map_inj_iff e.symm f g -#align - category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff - -instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) - where mem_essImage Y := ⟨e.symm Y, by simp⟩ -#align - category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor - -noncomputable instance inducedFunctorOfEquiv {C' : Type _} (e : C' ≃ D) : - IsEquivalence (inducedFunctor e) := - Equivalence.ofFullyFaithfullyEssSurj _ -#align - category_theory.equivalence.induced_functor_of_equiv CategoryTheory.Equivalence.inducedFunctorOfEquiv - -noncomputable instance fullyFaithfulToEssImage (F : C ⥤ D) [Full F] [Faithful F] : - IsEquivalence F.toEssImage := - ofFullyFaithfullyEssSurj F.toEssImage -#align - category_theory.equivalence.fully_faithful_to_ess_image CategoryTheory.Equivalence.fullyFaithfulToEssImage - -end Equivalence - -end CategoryTheory - diff --git a/Mathlib/CategoryTheory/Functor/Basic.lean b/Mathlib/CategoryTheory/Functor/Basic.lean index fa994fdda6e07..03acea2012d70 100644 --- a/Mathlib/CategoryTheory/Functor/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/Basic.lean @@ -45,6 +45,7 @@ structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category. map_id : ∀ X : C, map (𝟙 X) = 𝟙 (obj X) := by aesop_cat /-- A functor preserves composition. -/ map_comp : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map (f ≫ g) = map f ≫ map g := by aesop_cat + #align category_theory.functor CategoryTheory.Functor #align category_theory.functor.map_comp CategoryTheory.Functor.map_comp #align category_theory.functor.map_id CategoryTheory.Functor.map_id diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 72fd2d376e88d..9a2648c1fb895 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -105,7 +105,6 @@ def rewriteTarget' (stx : Syntax) (symm : Bool) : TacticM Unit := do let mvarId' ← (← getMainGoal).replaceTargetEq r.eNew r.eqProof replaceMainGoal ([mvarId']) - /-- Run a tactic on all goals, and always succeeds. diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index b58b9a8df03b9..f020d245eeced 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -13,9 +13,7 @@ open Lean Parser.Tactic Elab Command Elab.Tactic Meta -- TODO someone might like to generalise this tactic to work with other associative structures. namespace Tactic -variable [Monad m] [MonadExceptOf Exception m] - -partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α) := do +partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do try let a ← tac let l ← iterateUntilFailureWithResults tac @@ -23,7 +21,7 @@ partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α catch _ => pure [] #align tactic.repeat_with_results Tactic.iterateUntilFailureWithResults -def iterateUntilFailureCount {α : Type} (tac : m α) : m ℕ := do +def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM ℕ := do let r ← iterateUntilFailureWithResults tac return r.length #align tactic.repeat_count Tactic.iterateUntilFailureCount @@ -37,17 +35,20 @@ open Tactic variable [Monad m] [MonadExceptOf Exception m] def evalSlice (a b : Nat) : TacticM Unit := do - iterateUntilFailure do - ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + let _ ← iterateUntilFailureWithResults do + -- ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + evalTactic (← `(conv| rw [Category.assoc])) iterateRange (a - 1) (a - 1) do evalTactic (← `(conv| congr)) evalTactic (← `(tactic| rotate_left)) let k ← iterateUntilFailureCount - <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) + <| evalTactic (← `(conv| rw [←Category.assoc])) + -- <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) let c := k+1+a-b iterateRange c c <| evalTactic (← `(conv| congr)) - iterateUntilFailure do - ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) + let _ ← iterateUntilFailureWithResults do + evalTactic (← `(conv| rw [Category.assoc])) + -- ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat From ef7ef69d790cee30f8d56ffdce215e9451058ae7 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 09:37:50 -0500 Subject: [PATCH 08/66] remove rewriteTarget'; change MonadExceptOf to MonadExcept --- Mathlib/Tactic/Core.lean | 12 +----------- 1 file changed, 1 insertion(+), 11 deletions(-) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 9a2648c1fb895..f7f7792b3928d 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -95,16 +95,6 @@ end Lean namespace Lean.Elab.Tactic -/-- -Variant on `rewriteTarget` that behaves more like the Lean 3 version --/ -def rewriteTarget' (stx : Syntax) (symm : Bool) : TacticM Unit := do - Term.withSynthesize <| withMainContext do - let e ← elabTerm stx none true - let r ← (← getMainGoal).rewrite (← getMainTarget) e symm (config := {}) - let mvarId' ← (← getMainGoal).replaceTargetEq r.eNew r.eqProof - replaceMainGoal ([mvarId']) - /-- Run a tactic on all goals, and always succeeds. @@ -135,7 +125,7 @@ def allGoals (tac : TacticM Unit) : TacticM Unit := do def andThenOnSubgoals (tac1 : TacticM Unit) (tac2 : TacticM Unit) : TacticM Unit := focus do tac1; allGoals tac2 -variable [Monad m] [MonadExceptOf Exception m] +variable [Monad m] [MonadExcept Exception m] /-- Repeats a tactic at most `n` times, stopping sooner if the tactic fails. Always succeeds. -/ From 84ddfcb900a73c679d770e47eebe2fdb7770e579 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:00:38 -0500 Subject: [PATCH 09/66] add documentation and clean up --- Mathlib/Tactic/Slice.lean | 36 +++++++++++++++++++++++++++++------- 1 file changed, 29 insertions(+), 7 deletions(-) diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index f020d245eeced..5b5f8ebbb7060 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -13,6 +13,11 @@ open Lean Parser.Tactic Elab Command Elab.Tactic Meta -- TODO someone might like to generalise this tactic to work with other associative structures. namespace Tactic +variable [Monad m] [MonadExcept Exception m] + +/-- `iterateUntilFailureWithResults` is a helper tactic which returns the results of `tac`'s +iterative application along the lines of `iterateUntilFailure` +-/ partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do try let a ← tac @@ -21,6 +26,9 @@ partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : Tact catch _ => pure [] #align tactic.repeat_with_results Tactic.iterateUntilFailureWithResults +/-- `iterateUntilFailureCount` is similiar to `iterateUntilFailure` except it counts +the number of successful calls to `tac` +-/ def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM ℕ := do let r ← iterateUntilFailureWithResults tac return r.length @@ -31,28 +39,41 @@ end Tactic namespace Conv open Tactic +open Parser.Tactic.Conv -variable [Monad m] [MonadExceptOf Exception m] - +/-- +`evalSlice` +- rewrites the target express using `Category.assoc`. +- uses `congr` to split off the first `a-1` terms and rotates to `a`-th (last) term +- it counts the number `k` of rewrites as it uses `←Category.assoc` to bring the target to + left associated form; from the first step this is the total number of remaining terms from `C` +- it now splits off `b-a` terms from target using `congr` leaving the desired subterm +- finally, it rewrites it once more using `Category.assoc` to bring it right associated + normal form +-/ def evalSlice (a b : Nat) : TacticM Unit := do let _ ← iterateUntilFailureWithResults do - -- ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) evalTactic (← `(conv| rw [Category.assoc])) iterateRange (a - 1) (a - 1) do evalTactic (← `(conv| congr)) evalTactic (← `(tactic| rotate_left)) let k ← iterateUntilFailureCount <| evalTactic (← `(conv| rw [←Category.assoc])) - -- <| ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := true) let c := k+1+a-b iterateRange c c <| evalTactic (← `(conv| congr)) let _ ← iterateUntilFailureWithResults do evalTactic (← `(conv| rw [Category.assoc])) - -- ``(Category.assoc) >>= fun e => rewriteTarget' e (symm := false) elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat -open Lean Parser.Tactic Parser.Tactic.Conv Elab.Tactic Meta +/-- +`sliceLHS a b => tac` is a conv tactic which enters the LHS of a target, +uses `slice` to extract the `a` through `b` terms, and then applies +`tac` to the result. + +`sliceRHS a b => tac` works on the RHS similarly +-/ + syntax (name := sliceLHS) "sliceLHS" num num " => " convSeq : tactic macro_rules | `(tactic| sliceLHS $a $b => $seq) => @@ -63,9 +84,10 @@ macro_rules | `(tactic| sliceRHS $a $b => $seq) => `(tactic| conv => rhs; slice $a $b; ($seq:convSeq)) +/- Porting note: update when `add_tactic_doc` is supported` -/ -- add_tactic_doc -- { Name := "slice" -- category := DocCategory.tactic --- declNames := [`tactic.interactive.slice_lhs, `tactic.interactive.slice_rhs] +-- declNames := [`tactic.interactive.sliceLHS, `tactic.interactive.sliceRHS] -- tags := ["category theory"] } -- From 3c369ce1bb7ec4a6689730f20ab566dc48c28129 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:05:43 -0500 Subject: [PATCH 10/66] move iteration tactics to core --- Mathlib/Tactic/Core.lean | 17 +++++++++++++++++ Mathlib/Tactic/Slice.lean | 26 +------------------------- 2 files changed, 18 insertions(+), 25 deletions(-) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index f7f7792b3928d..2feef91248963 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -155,6 +155,23 @@ def iterateRange : Nat → Nat → m Unit → m Unit partial def iterateUntilFailure (tac : m Unit) : m Unit := try tac; iterateUntilFailure tac catch _ => pure () +/-- `iterateUntilFailureWithResults` is a helper tactic which returns the results of `tac`'s +iterative application along the lines of `iterateUntilFailure` +-/ +partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do + try + let a ← tac + let l ← iterateUntilFailureWithResults tac + pure (a :: l) + catch _ => pure [] + +/-- `iterateUntilFailureCount` is similiar to `iterateUntilFailure` except it counts +the number of successful calls to `tac` +-/ +def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM Nat := do + let r ← iterateUntilFailureWithResults tac + return r.length + end Lean.Elab.Tactic namespace Mathlib diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index 5b5f8ebbb7060..afcb917cbf55f 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -11,32 +11,8 @@ open CategoryTheory open Lean Parser.Tactic Elab Command Elab.Tactic Meta -- TODO someone might like to generalise this tactic to work with other associative structures. -namespace Tactic -variable [Monad m] [MonadExcept Exception m] - -/-- `iterateUntilFailureWithResults` is a helper tactic which returns the results of `tac`'s -iterative application along the lines of `iterateUntilFailure` --/ -partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do - try - let a ← tac - let l ← iterateUntilFailureWithResults tac - pure (a :: l) - catch _ => pure [] -#align tactic.repeat_with_results Tactic.iterateUntilFailureWithResults - -/-- `iterateUntilFailureCount` is similiar to `iterateUntilFailure` except it counts -the number of successful calls to `tac` --/ -def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM ℕ := do - let r ← iterateUntilFailureWithResults tac - return r.length -#align tactic.repeat_count Tactic.iterateUntilFailureCount - -end Tactic - -namespace Conv +/- Porting note: moved `repeat_with_results` to `repeat_count` to `Mathlib.Tactic.Core` -/ open Tactic open Parser.Tactic.Conv From 69bbb9346f7e9bad98ef050eac82c1a63d55dcdc Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:08:38 -0500 Subject: [PATCH 11/66] add slice to global import file --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 6a446f2fa0329..e0a7fbf82f764 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -986,6 +986,7 @@ import Mathlib.Tactic.SimpIntro import Mathlib.Tactic.SimpRw import Mathlib.Tactic.Simps.Basic import Mathlib.Tactic.Simps.NotationClass +import Mathlib.Tactic.Slice import Mathlib.Tactic.SolveByElim import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Spread From 319fa30f96cbc9d4bdaeb654f12da65ee34c1c29 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:11:35 -0500 Subject: [PATCH 12/66] modify documentation lines --- Mathlib/Tactic/Core.lean | 4 ++-- Mathlib/Tactic/Slice.lean | 7 ++++--- 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 2feef91248963..4810ec4315a60 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -156,7 +156,7 @@ partial def iterateUntilFailure (tac : m Unit) : m Unit := try tac; iterateUntilFailure tac catch _ => pure () /-- `iterateUntilFailureWithResults` is a helper tactic which returns the results of `tac`'s -iterative application along the lines of `iterateUntilFailure` +iterative application along the lines of `iterateUntilFailure`. Always succeeds. -/ partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do try @@ -166,7 +166,7 @@ partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : Tact catch _ => pure [] /-- `iterateUntilFailureCount` is similiar to `iterateUntilFailure` except it counts -the number of successful calls to `tac` +the number of successful calls to `tac`. Always succeeds. -/ def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM Nat := do let r ← iterateUntilFailureWithResults tac diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index afcb917cbf55f..3a9eacd923b2f 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -40,21 +40,22 @@ def evalSlice (a b : Nat) : TacticM Unit := do let _ ← iterateUntilFailureWithResults do evalTactic (← `(conv| rw [Category.assoc])) +/-- `slice` is implemented by `evalSlice` -/ elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat /-- `sliceLHS a b => tac` is a conv tactic which enters the LHS of a target, uses `slice` to extract the `a` through `b` terms, and then applies `tac` to the result. - -`sliceRHS a b => tac` works on the RHS similarly -/ - syntax (name := sliceLHS) "sliceLHS" num num " => " convSeq : tactic macro_rules | `(tactic| sliceLHS $a $b => $seq) => `(tactic| conv => lhs; slice $a $b; ($seq:convSeq)) +/-- +`sliceRHS a b => tac` works as `sliceLHS` but on the RHS similarly +-/ syntax (name := sliceRHS) "sliceRHS" num num " => " convSeq : tactic macro_rules | `(tactic| sliceRHS $a $b => $seq) => From 720d3112e3d6761c8520455cad9a208bf4f14380 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:30:40 -0500 Subject: [PATCH 13/66] add module documentation(?) --- Mathlib/Tactic/Core.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 4810ec4315a60..446b9febbea02 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -6,6 +6,14 @@ Authors: Arthur Paulino, Aurélien Saue, Mario Carneiro import Std.Tactic.Simpa import Mathlib.Lean.Expr +/-! +# The `slice` tactic + +Applies a tactic to an interval of terms from a term obtained by repeated application +of `Category.comp`. + +-/ + open Lean.Elab.Tactic namespace Lean From 8256a85b5a1300254f6311e8dc6520d7be89787d Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:32:20 -0500 Subject: [PATCH 14/66] minor changes --- Mathlib/CategoryTheory/Equivalence.lean | 101 ++++++++++++------------ Mathlib/CategoryTheory/Iso.lean | 6 +- 2 files changed, 56 insertions(+), 51 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index ae41d1168b7ba..8e387a568c199 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -12,7 +12,8 @@ import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.FullSubcategory import Mathlib.CategoryTheory.Whiskering import Mathlib.CategoryTheory.EssentialImage --- TODO: import Mathlib.Tactic.Slice +import Mathlib.Tactic.Slice +import Mathlib.Tactic.Relation.Rfl -- TODO: remove set_option autoImplicit false /-! @@ -184,28 +185,28 @@ theorem counit_app_functor (e : C ≌ D) (X : C) : http://globular.science/1905.001 -/ @[simp] theorem unit_inverse_comp (e : C ≌ D) (Y : D) : - e.Unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := + e.unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := by - rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counit_inv_functor_comp, map_comp] + rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counitInv_functor_comp, map_comp] dsimp - rw [← iso.hom_inv_id_assoc (e.unit_iso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] - slice_lhs 2 3 => erw [e.unit.naturality] - slice_lhs 1 2 => erw [e.unit.naturality] - slice_lhs 4 4 => - rw [← iso.hom_inv_id_assoc (e.inverse.map_iso (e.counit_iso.app _)) (e.unit_inv.app _)] - slice_lhs 3 4 => + rw [← Iso.hom_inv_id_assoc (e.unitIso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] + sliceLHS 2 3 => erw [e.unit.naturality] + sliceLHS 1 2 => erw [e.unit.naturality] + sliceLHS 4 4 => + rw [← Iso.hom_inv_id_assoc (e.inverse.mapIso (e.counitIso.app _)) (e.unitInv.app _)] + sliceLHS 3 4 => erw [← map_comp e.inverse, e.counit.naturality] - erw [(e.counit_iso.app _).hom_inv_id, map_id] + erw [(e.counitIso.app _).hom_inv_id, map_id] erw [id_comp] - slice_lhs 2 3 => erw [← map_comp e.inverse, e.counit_iso.inv.naturality, map_comp] - slice_lhs 3 4 => erw [e.unit_inv.naturality] - slice_lhs 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unit_iso.app _).hom_inv_id, map_id] + sliceLHS 2 3 => erw [← map_comp e.inverse, e.counitIso.inv.naturality, map_comp] + sliceLHS 3 4 => erw [e.unitInv.naturality] + sliceLHS 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unitIso.app _).hom_inv_id, map_id] erw [id_comp] - slice_lhs 3 4 => erw [← e.unit_inv.naturality] - slice_lhs 2 3 => - erw [← map_comp e.inverse, ← e.counit_iso.inv.naturality, (e.counit_iso.app _).hom_inv_id, + sliceLHS 3 4 => erw [← e.unitInv.naturality] + sliceLHS 2 3 => + erw [← map_comp e.inverse, ← e.counitIso.inv.naturality, (e.counitIso.app _).hom_inv_id, map_id] - erw [id_comp, (e.unit_iso.app _).hom_inv_id]; rfl + erw [id_comp, (e.unitIso.app _).hom_inv_id]; rfl #align category_theory.equivalence.unit_inverse_comp CategoryTheory.Equivalence.unit_inverse_comp @[simp] @@ -219,14 +220,13 @@ theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp theorem unit_app_inverse (e : C ≌ D) (Y : D) : - e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := - by + e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := by erw [← Iso.comp_hom_eq_id (e.inverse.mapIso (e.counitIso.app Y)), unit_inverse_comp] + dsimp #align category_theory.equivalence.unit_app_inverse CategoryTheory.Equivalence.unit_app_inverse theorem unit_inv_app_inverse (e : C ≌ D) (Y : D) : - e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := - by + e.unitInv.app (e.inverse.obj Y) = e.inverse.map (e.counit.app Y) := by symm erw [← Iso.hom_comp_eq_id (e.unitIso.app _), unit_inverse_comp] rfl @@ -253,7 +253,7 @@ variable {F : C ⥤ D} {G : D ⥤ C} (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ /-- If `η : 𝟭 C ≅ F ⋙ G` is part of a (not necessarily half-adjoint) equivalence, we can upgrade it to a refined natural isomorphism `adjointify_η η : 𝟭 C ≅ F ⋙ G` which exhibits the properties required for a half-adjoint equivalence. See `equivalence.mk`. -/ -def adjointifyη : 𝟭 C ≅ F ⋙ G := +def adjointifyη : 𝟭 C ≅ F ⋙ G := by calc 𝟭 C ≅ F ⋙ G := η _ ≅ F ⋙ 𝟭 D ⋙ G := isoWhiskerLeft F (leftUnitor G).symm @@ -427,14 +427,14 @@ theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : @[simp] theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : f ≫ g ≫ e.unit.app Y = f' ≫ g' ≫ e.unit.app Y ↔ f ≫ g = f' ≫ g' := by - simp only [← category.assoc, cancel_mono] + simp only [← Category.assoc, cancel_mono] #align category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc @[simp] theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : f ≫ g ≫ e.counitInv.app Y = f' ≫ g' ≫ e.counitInv.app Y ↔ f ≫ g = f' ≫ g' := by - simp only [← category.assoc, cancel_mono] + simp only [← Category.assoc, cancel_mono] #align category_theory.equivalence.cancel_counit_inv_right_assoc CategoryTheory.Equivalence.cancel_counit_inv_right_assoc @@ -442,7 +442,7 @@ theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : f ≫ g ≫ h ≫ e.unit.app Z = f' ≫ g' ≫ h' ≫ e.unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by - simp only [← category.assoc, cancel_mono] + simp only [← Category.assoc, cancel_mono] #align category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' @@ -450,7 +450,7 @@ theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) theorem cancel_counit_inv_right_assoc' {W X X' Y Y' Z : D} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : f ≫ g ≫ h ≫ e.counitInv.app Z = f' ≫ g' ≫ h' ≫ e.counitInv.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := - by simp only [← category.assoc, cancel_mono] + by simp only [← Category.assoc, cancel_mono] #align category_theory.equivalence.cancel_counit_inv_right_assoc' CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' @@ -512,7 +512,7 @@ class IsEquivalence (F : C ⥤ D) where mk' :: restate_axiom IsEquivalence.functor_unit_iso_comp' -attribute [simp, reassoc.1] is_equivalence.functor_unit_iso_comp +attribute [reassoc (attr := simp)] IsEquivalence.functor_unit_iso_comp namespace IsEquivalence @@ -649,15 +649,15 @@ def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso functor_unit_iso_comp' X := by - dsimp [nat_iso.hcomp] + dsimp [NatIso.hcomp] erw [id_comp, F.map_id, comp_id] apply (cancel_epi (e.hom.app X)).mp - slice_lhs 1 2 => rw [← e.hom.naturality] - slice_lhs 2 3 => rw [← nat_trans.vcomp_app', e.hom_inv_id] - simp only [nat_trans.id_app, id_comp, comp_id, F.map_comp, assoc] - erw [hF.counit_iso.hom.naturality] - slice_lhs 1 2 => rw [functor_unit_iso_comp] - simp only [functor.id_map, id_comp] + sliceLHS 1 2 => rw [← e.hom.naturality] + sliceLHS 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] + simp only [NatTrans.id_app, id_comp, comp_id, F.map_comp, assoc] + erw [hF.counitIso.hom.naturality] + sliceLHS 1 2 => rw [functor_unit_iso_comp] + simp only [Functor.id_map, id_comp] #align category_theory.is_equivalence.of_iso CategoryTheory.IsEquivalence.ofIso /-- Compatibility of `of_iso` with the composition of isomorphisms of functors -/ @@ -693,7 +693,7 @@ def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence /-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ @[simp] def cancelCompRight {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : IsEquivalence G) - (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence F := + (_ : IsEquivalence (F ⋙ G)) : IsEquivalence F := ofIso (Functor.associator F G G.inv ≪≫ NatIso.hcomp (Iso.refl F) hG.unitIso.symm ≪≫ rightUnitor F) (Functor.isEquivalenceTrans (F ⋙ G) G.inv) #align category_theory.is_equivalence.cancel_comp_right CategoryTheory.IsEquivalence.cancelCompRight @@ -701,7 +701,7 @@ def cancelCompRight {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : /-- If `F` and `F ⋙ G` are equivalence of categories, then `G` is also an equivalence. -/ @[simp] def cancelCompLeft {E : Type _} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hF : IsEquivalence F) - (hGF : IsEquivalence (F ⋙ G)) : IsEquivalence G := + (_ : IsEquivalence (F ⋙ G)) : IsEquivalence G := ofIso ((Functor.associator F.inv F G).symm ≪≫ NatIso.hcomp hF.counitIso (Iso.refl G) ≪≫ leftUnitor G) (Functor.isEquivalenceTrans F.inv (F ⋙ G)) @@ -726,9 +726,9 @@ theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := See . -/ instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence F] : Faithful F where - map_injective X Y f g w := by - have p := congr_arg (@CategoryTheory.Functor.map _ _ _ _ F.inv _ _) w - simpa only [cancel_epi, cancel_mono, is_equivalence.inv_fun_map] using p + map_injective := @fun X Y f g h => by + have p : F.inv.map (F.map f) = F.inv.map (F.map g) := congrArg F.inv.map h + simpa only [cancel_epi, cancel_mono, IsEquivalence.inv_fun_map] using p #align category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithful_of_equivalence @@ -739,23 +739,25 @@ See . -/ instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F where - preimage X Y f := F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y - witness X Y f := + preimage := @fun X Y f => F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y + witness := @fun X Y f => F.inv.map_injective <| by - simpa only [is_equivalence.inv_fun_map, assoc, iso.inv_hom_id_app_assoc, - iso.inv_hom_id_app] using comp_id _ + simpa only [IsEquivalence.inv_fun_map, assoc, Iso.inv_hom_id_app_assoc, + Iso.inv_hom_id_app] using comp_id _ #align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.fullOfEquivalence @[simps] private noncomputable def equivalenceInverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C where obj X := F.objPreimage X - map X Y f := F.preimage ((F.objObjPreimageIso X).Hom ≫ f ≫ (F.objObjPreimageIso Y).inv) + map := @fun X Y f => F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv) map_id X := by apply F.map_injective; aesop_cat - map_comp X Y Z f g := by apply F.map_injective <;> simp -#align - category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse + map_comp := @fun X Y Z f g => by apply F.map_injective; simp +-- #align +-- category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse +/- Porting note: this is a private def in mathlib -/ +-- /-- A functor which is full, faithful, and essentially surjective is an equivalence. See . @@ -764,7 +766,7 @@ noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [ IsEquivalence F := IsEquivalence.mk (equivalenceInverse F) (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) - fun X Y f => by + @fun X Y f => by apply F.map_injective aesop_cat) (NatIso.ofComponents F.objObjPreimageIso (by aesop_cat)) @@ -786,7 +788,8 @@ theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) - where mem_essImage Y := ⟨e.symm Y, by simp⟩ + where mem_essImage Y := + ⟨e.symm Y, by simp ⟩ #align category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 1d599f2db23c9..02b6f08af893d 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -140,6 +140,8 @@ def refl (X : C) : X ≅ X where instance : Inhabited (X ≅ X) := ⟨Iso.refl X⟩ +theorem nonempty_iso_refl (X : C) : Nonempty (X ≅ X) := ⟨default⟩ + @[simp] theorem refl_symm (X : C) : (Iso.refl X).symm = Iso.refl X := rfl #align category_theory.iso.refl_symm CategoryTheory.Iso.refl_symm @@ -153,8 +155,8 @@ def trans (α : X ≅ Y) (β : Y ≅ Z) : X ≅ Z where #align category_theory.iso.trans_hom CategoryTheory.Iso.trans_hom #align category_theory.iso.trans_inv CategoryTheory.Iso.trans_inv -instance : Trans (Iso : C → C → Type v) (Iso : C → C → Type v) (Iso : C → C → Type v) where - trans := Iso.trans +instance : Trans (· ≅ ·) (· ≅ ·) ((·:C) ≅ ·) where + trans := trans /-- Notation for composition of isomorphisms. -/ infixr:80 " ≪≫ " => Iso.trans -- type as `\ll \gg`. From 388aaa03c4b09f991b25b230ffa4086bb9ffafc1 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 10:51:12 -0500 Subject: [PATCH 15/66] fix module docs --- Mathlib/Tactic/Core.lean | 5 ++--- Mathlib/Tactic/Slice.lean | 9 +++++++++ 2 files changed, 11 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 446b9febbea02..f145a06d42f5e 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -7,10 +7,9 @@ import Std.Tactic.Simpa import Mathlib.Lean.Expr /-! -# The `slice` tactic +# -Applies a tactic to an interval of terms from a term obtained by repeated application -of `Category.comp`. +Generally useful tactics. -/ diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index 3a9eacd923b2f..c36c1e1f6a681 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -7,6 +7,15 @@ Authors: Scott Morrison import Mathlib.CategoryTheory.Category.Basic import Mathlib.Tactic.Conv +/-! +# The `slice` tactic + +Applies a tactic to an interval of terms from a term obtained by repeated application +of `Category.comp`. + +-/ + + open CategoryTheory open Lean Parser.Tactic Elab Command Elab.Tactic Meta From 6bae31a152336995e8aa9bc953fc9aa7ae266df0 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 12:05:58 -0500 Subject: [PATCH 16/66] fix test/slice.lean --- test/slice.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/test/slice.lean b/test/slice.lean index e1d8162dbfed6..c8c8fac9a649e 100644 --- a/test/slice.lean +++ b/test/slice.lean @@ -13,4 +13,3 @@ example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ lhs slice 1 1 rw [h₁] - rfl From 7140d1c4c844f59c3f16ba2e2a3681cec306643c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 12:48:18 -0500 Subject: [PATCH 17/66] fix all but refl error --- Mathlib/CategoryTheory/Equivalence.lean | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 8e387a568c199..5d3e6dc2db975 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -14,6 +14,7 @@ import Mathlib.CategoryTheory.Whiskering import Mathlib.CategoryTheory.EssentialImage import Mathlib.Tactic.Slice import Mathlib.Tactic.Relation.Rfl +import Mathlib.Tactic.Relation.Trans -- TODO: remove set_option autoImplicit false /-! @@ -268,12 +269,13 @@ def adjointifyη : 𝟭 C ≅ F ⋙ G := by theorem adjointify_η_ε (X : C) : F.map ((adjointifyη η ε).hom.app X) ≫ ε.hom.app (F.obj X) = 𝟙 (F.obj X) := by - dsimp [adjointifyη]; simp + dsimp [adjointifyη,Trans.trans] + simp have := ε.hom.naturality (F.map (η.inv.app X)); dsimp at this; rw [this]; clear this rw [← assoc _ _ (F.map _)] have := ε.hom.naturality (ε.inv.app <| F.obj X); dsimp at this; rw [this]; clear this have := (ε.app <| F.obj X).hom_inv_id; dsimp at this; rw [this]; clear this - rw [id_comp]; have := (F.map_iso <| η.app X).hom_inv_id; dsimp at this; rw [this] + rw [id_comp]; have := (F.mapIso <| η.app X).hom_inv_id; dsimp at this; rw [this] #align category_theory.equivalence.adjointify_η_ε CategoryTheory.Equivalence.adjointify_η_ε end @@ -373,7 +375,7 @@ theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app /-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ -@[simps functor inverse unitIso counitIso] +@[simps! functor inverse unitIso counitIso] def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := Equivalence.mk ((whiskeringLeft _ _ _).obj e.inverse) ((whiskeringLeft _ _ _).obj e.functor) (NatIso.ofComponents (fun F => (e.funInvIdAssoc F).symm) (by aesop_cat)) @@ -381,7 +383,7 @@ def congrLeft (e : C ≌ D) : C ⥤ E ≌ D ⥤ E := #align category_theory.equivalence.congr_left CategoryTheory.Equivalence.congrLeft /-- If `C` is equivalent to `D`, then `E ⥤ C` is equivalent to `E ⥤ D`. -/ -@[simps functor inverse unitIso counitIso] +@[simps! functor inverse unitIso counitIso] def congrRight (e : C ≌ D) : E ⥤ C ≌ E ⥤ D := Equivalence.mk ((whiskeringRight _ _ _).obj e.functor) ((whiskeringRight _ _ _).obj e.inverse) (NatIso.ofComponents @@ -789,7 +791,7 @@ theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) where mem_essImage Y := - ⟨e.symm Y, by simp ⟩ + ⟨e.symm Y, by simp only [inducedFunctor_obj,Equiv.apply_symm_apply]; exact ⟨default⟩⟩ #align category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor From 2c8c8a74a92cf0e19914fe47d679e40adde0564d Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 14:15:47 -0500 Subject: [PATCH 18/66] fix refl error and lint errors --- Mathlib/CategoryTheory/Equivalence.lean | 135 ++++++++++++------------ 1 file changed, 69 insertions(+), 66 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 5d3e6dc2db975..263b7ccad920d 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -13,10 +13,6 @@ import Mathlib.CategoryTheory.FullSubcategory import Mathlib.CategoryTheory.Whiskering import Mathlib.CategoryTheory.EssentialImage import Mathlib.Tactic.Slice -import Mathlib.Tactic.Relation.Rfl -import Mathlib.Tactic.Relation.Trans --- TODO: remove -set_option autoImplicit false /-! # Equivalence of categories @@ -83,19 +79,24 @@ universe v₁ v₂ v₃ u₁ u₂ u₃ See -/ -structure Equivalence (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] where mk' :: +structure Equivalence (C : Type u₁) (D : Type u₂) [Category.{v₁} C] [Category.{v₂} D] where mk' :: + /-- A functor in one direction -/ functor : C ⥤ D + /-- A functor in the other direction -/ inverse : D ⥤ C + /-- The composition `functor ⋙ inverse` is isomorphic to the identity -/ unitIso : 𝟭 C ≅ functor ⋙ inverse + /-- The composition `inverse ⋙ functor` is also isomorphic to the identity -/ counitIso : inverse ⋙ functor ≅ 𝟭 D - functor_unit_iso_comp' : ∀ X : C, functor.map (unitIso.hom.app X) ≫ counitIso.hom.app (functor.obj X) = 𝟙 (functor.obj X) := by aesop_cat + /-- The natural isomorphism compose to the identity -/ + functor_unit_iso_comp : + ∀ X : C, functor.map (unitIso.hom.app X) ≫ counitIso.hom.app (functor.obj X) = + 𝟙 (functor.obj X) := by aesop_cat #align category_theory.equivalence CategoryTheory.Equivalence -restate_axiom Equivalence.functor_unit_iso_comp' - --- mathport name: «expr ≌ » +/-- We infix the usual notation for an equivalence -/ infixr:10 " ≌ " => Equivalence - + variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] namespace Equivalence @@ -140,15 +141,15 @@ theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : theorem equivalence_mk'_unitInv (functor inverse unit_iso counit_iso f) : (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).unitInv = unit_iso.inv := rfl -#align - category_theory.equivalence.equivalence_mk'_unit_inv CategoryTheory.Equivalence.equivalence_mk'_unitInv +#align category_theory.equivalence.equivalence_mk'_unit_inv + CategoryTheory.Equivalence.equivalence_mk'_unitInv @[simp] theorem equivalence_mk'_counitInv (functor inverse unit_iso counit_iso f) : (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counitInv = counit_iso.inv := rfl -#align - category_theory.equivalence.equivalence_mk'_counit_inv CategoryTheory.Equivalence.equivalence_mk'_counitInv +#align category_theory.equivalence.equivalence_mk'_counit_inv + CategoryTheory.Equivalence.equivalence_mk'_counitInv @[simp] theorem functor_unit_comp (e : C ≌ D) (X : C) : @@ -163,8 +164,8 @@ theorem counitInv_functor_comp (e : C ≌ D) (X : C) : erw [Iso.inv_eq_inv (e.functor.mapIso (e.unitIso.app X) ≪≫ e.counitIso.app (e.functor.obj X)) (Iso.refl _)] exact e.functor_unit_comp X -#align - category_theory.equivalence.counit_inv_functor_comp CategoryTheory.Equivalence.counitInv_functor_comp +#align category_theory.equivalence.counit_inv_functor_comp + CategoryTheory.Equivalence.counitInv_functor_comp theorem counit_inv_app_functor (e : C ≌ D) (X : C) : e.counitInv.app (e.functor.obj X) = e.functor.map (e.unit.app X) := @@ -172,15 +173,16 @@ theorem counit_inv_app_functor (e : C ≌ D) (X : C) : symm erw [← Iso.comp_hom_eq_id (e.counitIso.app _), functor_unit_comp] rfl -#align - category_theory.equivalence.counit_inv_app_functor CategoryTheory.Equivalence.counit_inv_app_functor +#align category_theory.equivalence.counit_inv_app_functor + CategoryTheory.Equivalence.counit_inv_app_functor theorem counit_app_functor (e : C ≌ D) (X : C) : e.counit.app (e.functor.obj X) = e.functor.map (e.unitInv.app X) := by erw [← Iso.hom_comp_eq_id (e.functor.mapIso (e.unitIso.app X)), functor_unit_comp] rfl -#align category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor +#align + category_theory.equivalence.counit_app_functor CategoryTheory.Equivalence.counit_app_functor /-- The other triangle equality. The proof follows the following proof in Globular: http://globular.science/1905.001 -/ @@ -217,8 +219,8 @@ theorem inverse_counit_inv_comp (e : C ≌ D) (Y : D) : erw [Iso.inv_eq_inv (e.unitIso.app (e.inverse.obj Y) ≪≫ e.inverse.mapIso (e.counitIso.app Y)) (Iso.refl _)] exact e.unit_inverse_comp Y -#align - category_theory.equivalence.inverse_counit_inv_comp CategoryTheory.Equivalence.inverse_counit_inv_comp +#align category_theory.equivalence.inverse_counit_inv_comp + CategoryTheory.Equivalence.inverse_counit_inv_comp theorem unit_app_inverse (e : C ≌ D) (Y : D) : e.unit.app (e.inverse.obj Y) = e.inverse.map (e.counitInv.app Y) := by @@ -289,8 +291,8 @@ protected def mk (F : C ⥤ D) (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G /-- Equivalence of categories is reflexive. -/ @[refl, simps] -def refl : C ≌ C := - ⟨𝟭 C, 𝟭 C, Iso.refl _, Iso.refl _, fun X => Category.id_comp _⟩ +def refl : C ≌ C := + ⟨𝟭 C, 𝟭 C, Iso.refl _, Iso.refl _, fun _ => Category.id_comp _⟩ #align category_theory.equivalence.refl CategoryTheory.Equivalence.refl instance : Inhabited (C ≌ C) := @@ -319,7 +321,7 @@ def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E -- We wouldn't have needed to give this proof if we'd used `equivalence.mk`, -- but we choose to avoid using that here, for the sake of good structure projection `simp` -- lemmas. - functor_unit_iso_comp' X := by + functor_unit_iso_comp X := by dsimp rw [← f.functor.map_comp_assoc, e.functor.map_comp, ← counit_inv_app_functor, fun_inv_map, Iso.inv_hom_id_app_assoc, assoc, Iso.inv_hom_id_app, counit_app_functor, ← Functor.map_comp] @@ -338,8 +340,8 @@ theorem fun_inv_id_assoc_hom_app (e : C ≌ D) (F : C ⥤ E) (X : C) : by dsimp [funInvIdAssoc] aesop_cat -#align - category_theory.equivalence.fun_inv_id_assoc_hom_app CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app +#align category_theory.equivalence.fun_inv_id_assoc_hom_app + CategoryTheory.Equivalence.fun_inv_id_assoc_hom_app @[simp] theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : @@ -347,8 +349,8 @@ theorem fun_inv_id_assoc_inv_app (e : C ≌ D) (F : C ⥤ E) (X : C) : by dsimp [funInvIdAssoc] aesop_cat -#align - category_theory.equivalence.fun_inv_id_assoc_inv_app CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app +#align category_theory.equivalence.fun_inv_id_assoc_inv_app + CategoryTheory.Equivalence.fun_inv_id_assoc_inv_app /-- Composing a functor with both functors of an equivalence yields a naturally isomorphic functor. -/ @@ -362,8 +364,8 @@ theorem inv_fun_id_assoc_hom_app (e : C ≌ D) (F : D ⥤ E) (X : D) : by dsimp [invFunIdAssoc] aesop_cat -#align - category_theory.equivalence.inv_fun_id_assoc_hom_app CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app +#align category_theory.equivalence.inv_fun_id_assoc_hom_app + CategoryTheory.Equivalence.inv_fun_id_assoc_hom_app @[simp] theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : @@ -371,8 +373,8 @@ theorem inv_fun_id_assoc_inv_app (e : C ≌ D) (F : D ⥤ E) (X : D) : by dsimp [invFunIdAssoc] aesop_cat -#align - category_theory.equivalence.inv_fun_id_assoc_inv_app CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app +#align category_theory.equivalence.inv_fun_id_assoc_inv_app + CategoryTheory.Equivalence.inv_fun_id_assoc_inv_app /-- If `C` is equivalent to `D`, then `C ⥤ E` is equivalent to `D ⥤ E`. -/ @[simps! functor inverse unitIso counitIso] @@ -423,38 +425,38 @@ theorem cancel_counit_right {X Y : D} (f f' : X ⟶ e.functor.obj (e.inverse.obj @[simp] theorem cancel_counit_inv_right {X Y : D} (f f' : X ⟶ Y) : f ≫ e.counitInv.app Y = f' ≫ e.counitInv.app Y ↔ f = f' := by simp only [cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right CategoryTheory.Equivalence.cancel_counit_inv_right +#align category_theory.equivalence.cancel_counit_inv_right + CategoryTheory.Equivalence.cancel_counit_inv_right @[simp] theorem cancel_unit_right_assoc {W X X' Y : C} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : f ≫ g ≫ e.unit.app Y = f' ≫ g' ≫ e.unit.app Y ↔ f ≫ g = f' ≫ g' := by simp only [← Category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_unit_right_assoc CategoryTheory.Equivalence.cancel_unit_right_assoc +#align category_theory.equivalence.cancel_unit_right_assoc + CategoryTheory.Equivalence.cancel_unit_right_assoc @[simp] theorem cancel_counit_inv_right_assoc {W X X' Y : D} (f : W ⟶ X) (g : X ⟶ Y) (f' : W ⟶ X') (g' : X' ⟶ Y) : f ≫ g ≫ e.counitInv.app Y = f' ≫ g' ≫ e.counitInv.app Y ↔ f ≫ g = f' ≫ g' := by simp only [← Category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right_assoc CategoryTheory.Equivalence.cancel_counit_inv_right_assoc +#align category_theory.equivalence.cancel_counit_inv_right_assoc + CategoryTheory.Equivalence.cancel_counit_inv_right_assoc @[simp] theorem cancel_unit_right_assoc' {W X X' Y Y' Z : C} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : f ≫ g ≫ h ≫ e.unit.app Z = f' ≫ g' ≫ h' ≫ e.unit.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by simp only [← Category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_unit_right_assoc' CategoryTheory.Equivalence.cancel_unit_right_assoc' +#align category_theory.equivalence.cancel_unit_right_assoc' + CategoryTheory.Equivalence.cancel_unit_right_assoc' @[simp] theorem cancel_counit_inv_right_assoc' {W X X' Y Y' Z : D} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z) (f' : W ⟶ X') (g' : X' ⟶ Y') (h' : Y' ⟶ Z) : f ≫ g ≫ h ≫ e.counitInv.app Z = f' ≫ g' ≫ h' ≫ e.counitInv.app Z ↔ f ≫ g ≫ h = f' ≫ g' ≫ h' := by simp only [← Category.assoc, cancel_mono] -#align - category_theory.equivalence.cancel_counit_inv_right_assoc' CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' +#align category_theory.equivalence.cancel_counit_inv_right_assoc' + CategoryTheory.Equivalence.cancel_counit_inv_right_assoc' end CancellationLemmas @@ -502,18 +504,20 @@ end Equivalence /-- A functor that is part of a (half) adjoint equivalence -/ class IsEquivalence (F : C ⥤ D) where mk' :: + /-- The inverse functor to `F` -/ inverse : D ⥤ C + /-- Composition `F ⋙ inverse` is isomorphic to the identity. -/ unitIso : 𝟭 C ≅ F ⋙ inverse + /-- Composition `inverse ⋙ F` is isomorphic to the identity. =-/ counitIso : inverse ⋙ F ≅ 𝟭 D - functor_unit_iso_comp' : + /-- We natural isomorphisms are inverse -/ + functor_unit_iso_comp : ∀ X : C, F.map ((unitIso.hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counitIso.hom.app (F.obj X) = 𝟙 (F.obj X) := by aesop_cat #align category_theory.is_equivalence CategoryTheory.IsEquivalence -restate_axiom IsEquivalence.functor_unit_iso_comp' - attribute [reassoc (attr := simp)] IsEquivalence.functor_unit_iso_comp namespace IsEquivalence @@ -524,8 +528,8 @@ instance ofEquivalence (F : C ≌ D) : IsEquivalence F.functor := instance ofEquivalenceInverse (F : C ≌ D) : IsEquivalence F.inverse := IsEquivalence.ofEquivalence F.symm -#align - category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.IsEquivalence.ofEquivalenceInverse +#align category_theory.is_equivalence.of_equivalence_inverse + CategoryTheory.IsEquivalence.ofEquivalenceInverse open Equivalence @@ -611,16 +615,16 @@ theorem functor_as_equivalence (E : C ≌ D) : E.functor.asEquivalence = E := by cases E congr -#align - category_theory.equivalence.functor_as_equivalence CategoryTheory.Equivalence.functor_as_equivalence +#align category_theory.equivalence.functor_as_equivalence + CategoryTheory.Equivalence.functor_as_equivalence @[simp] theorem inverse_as_equivalence (E : C ≌ D) : E.inverse.asEquivalence = E.symm := by cases E congr -#align - category_theory.equivalence.inverse_as_equivalence CategoryTheory.Equivalence.inverse_as_equivalence +#align category_theory.equivalence.inverse_as_equivalence + CategoryTheory.Equivalence.inverse_as_equivalence end Equivalence @@ -644,13 +648,13 @@ theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : /-- When a functor `F` is an equivalence of categories, and `G` is isomorphic to `F`, then `G` is also an equivalence of categories. -/ -@[simps] +@[simps!] def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G where inverse := hF.inverse unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso - functor_unit_iso_comp' X := by + functor_unit_iso_comp X := by dsimp [NatIso.hcomp] erw [id_comp, F.map_id, comp_id] apply (cancel_epi (e.hom.app X)).mp @@ -719,8 +723,8 @@ See . -/ theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F := ⟨fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩ -#align - category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.ess_surj_of_equivalence +#align category_theory.equivalence.ess_surj_of_equivalence + CategoryTheory.Equivalence.ess_surj_of_equivalence -- see Note [lower instance priority] /-- An equivalence is faithful. @@ -731,8 +735,8 @@ instance (priority := 100) faithful_of_equivalence (F : C ⥤ D) [IsEquivalence map_injective := @fun X Y f g h => by have p : F.inv.map (F.map f) = F.inv.map (F.map g) := congrArg F.inv.map h simpa only [cancel_epi, cancel_mono, IsEquivalence.inv_fun_map] using p -#align - category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithful_of_equivalence +#align category_theory.equivalence.faithful_of_equivalence + CategoryTheory.Equivalence.faithful_of_equivalence -- see Note [lower instance priority] /-- An equivalence is full. @@ -772,8 +776,8 @@ noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [ apply F.map_injective aesop_cat) (NatIso.ofComponents F.objObjPreimageIso (by aesop_cat)) -#align - category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj +#align category_theory.equivalence.of_fully_faithfully_ess_surj + CategoryTheory.Equivalence.ofFullyFaithfullyEssSurj @[simp] theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : @@ -792,22 +796,21 @@ theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : instance ess_surj_induced_functor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) where mem_essImage Y := ⟨e.symm Y, by simp only [inducedFunctor_obj,Equiv.apply_symm_apply]; exact ⟨default⟩⟩ -#align - category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.ess_surj_induced_functor +#align category_theory.equivalence.ess_surj_induced_functor + CategoryTheory.Equivalence.ess_surj_induced_functor noncomputable instance inducedFunctorOfEquiv {C' : Type _} (e : C' ≃ D) : IsEquivalence (inducedFunctor e) := Equivalence.ofFullyFaithfullyEssSurj _ -#align - category_theory.equivalence.induced_functor_of_equiv CategoryTheory.Equivalence.inducedFunctorOfEquiv +#align category_theory.equivalence.induced_functor_of_equiv + CategoryTheory.Equivalence.inducedFunctorOfEquiv noncomputable instance fullyFaithfulToEssImage (F : C ⥤ D) [Full F] [Faithful F] : IsEquivalence F.toEssImage := ofFullyFaithfullyEssSurj F.toEssImage -#align - category_theory.equivalence.fully_faithful_to_ess_image CategoryTheory.Equivalence.fullyFaithfulToEssImage +#align category_theory.equivalence.fully_faithful_to_ess_image + CategoryTheory.Equivalence.fullyFaithfulToEssImage end Equivalence end CategoryTheory - From 0c247c701c0a6f511f9483a88d9c38e7966af99a Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 14:19:58 -0500 Subject: [PATCH 19/66] fix final long line --- Mathlib/CategoryTheory/Equivalence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 263b7ccad920d..e1e6274b05bfe 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -134,8 +134,8 @@ theorem equivalence_mk'_unit (functor inverse unit_iso counit_iso f) : theorem equivalence_mk'_counit (functor inverse unit_iso counit_iso f) : (⟨functor, inverse, unit_iso, counit_iso, f⟩ : C ≌ D).counit = counit_iso.hom := rfl -#align - category_theory.equivalence.equivalence_mk'_counit CategoryTheory.Equivalence.equivalence_mk'_counit +#align category_theory.equivalence.equivalence_mk'_counit + CategoryTheory.Equivalence.equivalence_mk'_counit @[simp] theorem equivalence_mk'_unitInv (functor inverse unit_iso counit_iso f) : From 2c56111eafac3bdc6298b3b4bfec8bc209fc8c92 Mon Sep 17 00:00:00 2001 From: thorimur <68410468+thorimur@users.noreply.github.com> Date: Wed, 8 Feb 2023 14:57:53 -0500 Subject: [PATCH 20/66] use `Mathport` syntax * use existing docs * fix docs typos --- Mathlib/Tactic/Slice.lean | 44 ++++++++++++++++++++++----------------- 1 file changed, 25 insertions(+), 19 deletions(-) diff --git a/Mathlib/Tactic/Slice.lean b/Mathlib/Tactic/Slice.lean index c36c1e1f6a681..3c960a9eb5f6b 100644 --- a/Mathlib/Tactic/Slice.lean +++ b/Mathlib/Tactic/Slice.lean @@ -10,12 +10,11 @@ import Mathlib.Tactic.Conv /-! # The `slice` tactic -Applies a tactic to an interval of terms from a term obtained by repeated application -of `Category.comp`. +Applies a tactic to an interval of terms from a term obtained by repeated application +of `Category.comp`. -/ - open CategoryTheory open Lean Parser.Tactic Elab Command Elab.Tactic Meta @@ -27,13 +26,20 @@ open Tactic open Parser.Tactic.Conv /-- -`evalSlice` -- rewrites the target express using `Category.assoc`. +`slice` is a conv tactic; if the current focus is a composition of several morphisms, +`slice a b` reassociates as needed, and zooms in on the `a`-th through `b`-th morphisms. +Thus if the current focus is `(a ≫ b) ≫ ((c ≫ d) ≫ e)`, then `slice 2 3` zooms to `b ≫ c`. + -/ +syntax (name := slice) "slice " num num : conv + +/-- +`evalSlice` +- rewrites the target expression using `Category.assoc`. - uses `congr` to split off the first `a-1` terms and rotates to `a`-th (last) term -- it counts the number `k` of rewrites as it uses `←Category.assoc` to bring the target to +- counts the number `k` of rewrites as it uses `←Category.assoc` to bring the target to left associated form; from the first step this is the total number of remaining terms from `C` -- it now splits off `b-a` terms from target using `congr` leaving the desired subterm -- finally, it rewrites it once more using `Category.assoc` to bring it right associated +- it now splits off `b-a` terms from target using `congr` leaving the desired subterm +- finally, it rewrites it once more using `Category.assoc` to bring it to right-associated normal form -/ def evalSlice (a b : Nat) : TacticM Unit := do @@ -49,25 +55,25 @@ def evalSlice (a b : Nat) : TacticM Unit := do let _ ← iterateUntilFailureWithResults do evalTactic (← `(conv| rw [Category.assoc])) -/-- `slice` is implemented by `evalSlice` -/ -elab "slice" a:num b:num : conv => evalSlice a.getNat b.getNat +/-- `slice` is implemented by `evalSlice`. -/ +elab "slice " a:num b:num : conv => evalSlice a.getNat b.getNat -/-- -`sliceLHS a b => tac` is a conv tactic which enters the LHS of a target, -uses `slice` to extract the `a` through `b` terms, and then applies -`tac` to the result. +/-- +`slice_lhs a b => tac` zooms to the left hand side, uses associativity for categorical +composition as needed, zooms in on the `a`-th through `b`-th morphisms, and invokes `tac`. -/ -syntax (name := sliceLHS) "sliceLHS" num num " => " convSeq : tactic +syntax (name := sliceLHS) "slice_lhs " num num " => " convSeq : tactic macro_rules - | `(tactic| sliceLHS $a $b => $seq) => + | `(tactic| slice_lhs $a $b => $seq) => `(tactic| conv => lhs; slice $a $b; ($seq:convSeq)) /-- -`sliceRHS a b => tac` works as `sliceLHS` but on the RHS similarly +`slice_rhs a b => tac` zooms to the right hand side, uses associativity for categorical +composition as needed, zooms in on the `a`-th through `b`-th morphisms, and invokes `tac`. -/ -syntax (name := sliceRHS) "sliceRHS" num num " => " convSeq : tactic +syntax (name := sliceRHS) "slice_rhs " num num " => " convSeq : tactic macro_rules - | `(tactic| sliceRHS $a $b => $seq) => + | `(tactic| slice_rhs $a $b => $seq) => `(tactic| conv => rhs; slice $a $b; ($seq:convSeq)) /- Porting note: update when `add_tactic_doc` is supported` -/ From 80596f2643b17886668466a6c58802db2692817c Mon Sep 17 00:00:00 2001 From: thorimur <68410468+thorimur@users.noreply.github.com> Date: Wed, 8 Feb 2023 14:58:19 -0500 Subject: [PATCH 21/66] test: add simple `rhs`/`lhs` tests --- test/slice.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/test/slice.lean b/test/slice.lean index c8c8fac9a649e..818c82d6eb420 100644 --- a/test/slice.lean +++ b/test/slice.lean @@ -13,3 +13,9 @@ example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ lhs slice 1 1 rw [h₁] + +example (h₁ : f₁ = f₂) : f₁ ≫ g ≫ h ≫ l = ((f₂ ≫ g) ≫ h) ≫ l := by + slice_lhs 1 1 => rw [h₁] + +example (h₁ : f₁ = f₂) : ((f₂ ≫ g) ≫ h) ≫ l = f₁ ≫ g ≫ h ≫ l := by + slice_rhs 1 1 => rw [h₁] From ce383e4b8827543f5159043d7336daf84397233f Mon Sep 17 00:00:00 2001 From: thorimur <68410468+thorimur@users.noreply.github.com> Date: Wed, 8 Feb 2023 15:14:47 -0500 Subject: [PATCH 22/66] minor changes to `Tactic.Core` * use `m` instead of `TacticM` now that we use `MonadExcept` * simplify code for `iterateRange` * minor docs tweaks --- Mathlib/Tactic/Core.lean | 25 +++++++++++-------------- 1 file changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index f145a06d42f5e..0aa9edc66aae9 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -7,7 +7,7 @@ import Std.Tactic.Simpa import Mathlib.Lean.Expr /-! -# +# Generally useful tactics. @@ -142,40 +142,37 @@ def iterateAtMost : Nat → m Unit → m Unit /-- `iterateExactly' n t` executes `t` `n` times. If any iteration fails, the whole tactic fails. -/ -def iterateExactly' : Nat → m Unit → m Unit -| 0, _ => pure () +def iterateExactly' : Nat → m Unit → m Unit +| 0, _ => pure () | n+1, tac => tac *> iterateExactly' n tac /-- `iterateRange m n t`: Repeat the given tactic at least `m` times and at most `n` times or until `t` fails. Fails if `t` does not run at least `m` times. -/ -def iterateRange : Nat → Nat → m Unit → m Unit +def iterateRange : Nat → Nat → m Unit → m Unit | 0, 0, _ => pure () | 0, b, tac => iterateAtMost b tac -| (a+1), n, tac => do - let _ ← tac - let _ ← iterateRange a (n-1) tac - pure () +| (a+1), n, tac => do tac; iterateRange a (n-1) tac /-- Repeats a tactic until it fails. Always succeeds. -/ partial def iterateUntilFailure (tac : m Unit) : m Unit := try tac; iterateUntilFailure tac catch _ => pure () -/-- `iterateUntilFailureWithResults` is a helper tactic which returns the results of `tac`'s -iterative application along the lines of `iterateUntilFailure`. Always succeeds. +/-- `iterateUntilFailureWithResults` is a helper tactic which accumulates the list of results +obtained from iterating `tac` until it fails. Always succeeds. -/ -partial def iterateUntilFailureWithResults {α : Type} (tac : TacticM α) : TacticM (List α) := do +partial def iterateUntilFailureWithResults {α : Type} (tac : m α) : m (List α) := do try let a ← tac let l ← iterateUntilFailureWithResults tac pure (a :: l) catch _ => pure [] -/-- `iterateUntilFailureCount` is similiar to `iterateUntilFailure` except it counts -the number of successful calls to `tac`. Always succeeds. +/-- `iterateUntilFailureCount` is similar to `iterateUntilFailure` except it counts +the number of successful calls to `tac`. Always succeeds. -/ -def iterateUntilFailureCount {α : Type} (tac : TacticM α) : TacticM Nat := do +def iterateUntilFailureCount {α : Type} (tac : m α) : m Nat := do let r ← iterateUntilFailureWithResults tac return r.length From e6ec92461ccdb71348ed8a5cc8970ccdd0bd3fe7 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 8 Feb 2023 15:26:24 -0500 Subject: [PATCH 23/66] update slice naming --- Mathlib/CategoryTheory/Equivalence.lean | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index e1e6274b05bfe..3db834162b42d 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -193,20 +193,20 @@ theorem unit_inverse_comp (e : C ≌ D) (Y : D) : rw [← id_comp (e.inverse.map _), ← map_id e.inverse, ← counitInv_functor_comp, map_comp] dsimp rw [← Iso.hom_inv_id_assoc (e.unitIso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv] - sliceLHS 2 3 => erw [e.unit.naturality] - sliceLHS 1 2 => erw [e.unit.naturality] - sliceLHS 4 4 => + slice_lhs 2 3 => erw [e.unit.naturality] + slice_lhs 1 2 => erw [e.unit.naturality] + slice_lhs 4 4 => rw [← Iso.hom_inv_id_assoc (e.inverse.mapIso (e.counitIso.app _)) (e.unitInv.app _)] - sliceLHS 3 4 => + slice_lhs 3 4 => erw [← map_comp e.inverse, e.counit.naturality] erw [(e.counitIso.app _).hom_inv_id, map_id] erw [id_comp] - sliceLHS 2 3 => erw [← map_comp e.inverse, e.counitIso.inv.naturality, map_comp] - sliceLHS 3 4 => erw [e.unitInv.naturality] - sliceLHS 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unitIso.app _).hom_inv_id, map_id] + slice_lhs 2 3 => erw [← map_comp e.inverse, e.counitIso.inv.naturality, map_comp] + slice_lhs 3 4 => erw [e.unitInv.naturality] + slice_lhs 4 5 => erw [← map_comp (e.functor ⋙ e.inverse), (e.unitIso.app _).hom_inv_id, map_id] erw [id_comp] - sliceLHS 3 4 => erw [← e.unitInv.naturality] - sliceLHS 2 3 => + slice_lhs 3 4 => erw [← e.unitInv.naturality] + slice_lhs 2 3 => erw [← map_comp e.inverse, ← e.counitIso.inv.naturality, (e.counitIso.app _).hom_inv_id, map_id] erw [id_comp, (e.unitIso.app _).hom_inv_id]; rfl @@ -658,11 +658,11 @@ def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G dsimp [NatIso.hcomp] erw [id_comp, F.map_id, comp_id] apply (cancel_epi (e.hom.app X)).mp - sliceLHS 1 2 => rw [← e.hom.naturality] - sliceLHS 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] + slice_lhs 1 2 => rw [← e.hom.naturality] + slice_lhs 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] simp only [NatTrans.id_app, id_comp, comp_id, F.map_comp, assoc] erw [hF.counitIso.hom.naturality] - sliceLHS 1 2 => rw [functor_unit_iso_comp] + slice_lhs 1 2 => rw [functor_unit_iso_comp] simp only [Functor.id_map, id_comp] #align category_theory.is_equivalence.of_iso CategoryTheory.IsEquivalence.ofIso From bcc402f2fcb6d43de80c6eaac7fcfc1f038186b7 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Thu, 9 Feb 2023 13:49:56 -0500 Subject: [PATCH 24/66] feat: port CategoryTheory.Adjunction.Basic From 5bde22616c0eb101abbbb494ca351efbaa58e2a1 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Thu, 9 Feb 2023 13:49:57 -0500 Subject: [PATCH 25/66] Initial file copy from mathport --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 616 +++++++++++++++++++ 1 file changed, 616 insertions(+) create mode 100644 Mathlib/CategoryTheory/Adjunction/Basic.lean diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean new file mode 100644 index 0000000000000..57a6d3985bbd8 --- /dev/null +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -0,0 +1,616 @@ +/- +Copyright (c) 2019 Reid Barton. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Reid Barton, Johan Commelin, Bhavik Mehta + +! This file was ported from Lean 3 source module category_theory.adjunction.basic +! leanprover-community/mathlib commit d101e93197bb5f6ea89bd7ba386b7f7dff1f3903 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Equivalence + +/-! +# Adjunctions between functors + +`F ⊣ G` represents the data of an adjunction between two functors +`F : C ⥤ D` and `G : D ⥤ C`. `F` is the left adjoint and `G` is the right adjoint. + +We provide various useful constructors: +* `mk_of_hom_equiv` +* `mk_of_unit_counit` +* `left_adjoint_of_equiv` / `right_adjoint_of equiv` + construct a left/right adjoint of a given functor given the action on objects and + the relevant equivalence of morphism spaces. +* `adjunction_of_equiv_left` / `adjunction_of_equiv_right` witness that these constructions + give adjunctions. + +There are also typeclasses `is_left_adjoint` / `is_right_adjoint`, carrying data witnessing +that a given functor is a left or right adjoint. +Given `[is_left_adjoint F]`, a right adjoint of `F` can be constructed as `right_adjoint F`. + +`adjunction.comp` composes adjunctions. + +`to_equivalence` upgrades an adjunction to an equivalence, +given witnesses that the unit and counit are pointwise isomorphisms. +Conversely `equivalence.to_adjunction` recovers the underlying adjunction from an equivalence. +-/ + + +namespace CategoryTheory + +open Category + +-- declare the `v`'s first; see `category_theory.category` for an explanation +universe v₁ v₂ v₃ u₁ u₂ u₃ + +attribute [local elab_without_expected_type] whisker_left whisker_right + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +/-- `F ⊣ G` represents the data of an adjunction between two functors +`F : C ⥤ D` and `G : D ⥤ C`. `F` is the left adjoint and `G` is the right adjoint. + +To construct an `adjunction` between two functors, it's often easier to instead use the +constructors `mk_of_hom_equiv` or `mk_of_unit_counit`. To construct a left adjoint, +there are also constructors `left_adjoint_of_equiv` and `adjunction_of_equiv_left` (as +well as their duals) which can be simpler in practice. + +Uniqueness of adjoints is shown in `category_theory.adjunction.opposites`. + +See . +-/ +structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where + homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) + Unit : 𝟭 C ⟶ F.comp G + counit : G.comp F ⟶ 𝟭 D + homEquiv_unit' : ∀ {X Y f}, (hom_equiv X Y) f = (Unit : _ ⟶ _).app X ≫ G.map f := by obviously + homEquiv_counit' : ∀ {X Y g}, (hom_equiv X Y).symm g = F.map g ≫ counit.app Y := by obviously +#align category_theory.adjunction CategoryTheory.Adjunction + +-- mathport name: «expr ⊣ » +infixl:15 " ⊣ " => Adjunction + +/-- A class giving a chosen right adjoint to the functor `left`. -/ +class IsLeftAdjoint (left : C ⥤ D) where + right : D ⥤ C + adj : left ⊣ right +#align category_theory.is_left_adjoint CategoryTheory.IsLeftAdjoint + +/-- A class giving a chosen left adjoint to the functor `right`. -/ +class IsRightAdjoint (right : D ⥤ C) where + left : C ⥤ D + adj : left ⊣ right +#align category_theory.is_right_adjoint CategoryTheory.IsRightAdjoint + +/-- Extract the left adjoint from the instance giving the chosen adjoint. -/ +def leftAdjoint (R : D ⥤ C) [IsRightAdjoint R] : C ⥤ D := + IsRightAdjoint.left R +#align category_theory.left_adjoint CategoryTheory.leftAdjoint + +/-- Extract the right adjoint from the instance giving the chosen adjoint. -/ +def rightAdjoint (L : C ⥤ D) [IsLeftAdjoint L] : D ⥤ C := + IsLeftAdjoint.right L +#align category_theory.right_adjoint CategoryTheory.rightAdjoint + +/-- The adjunction associated to a functor known to be a left adjoint. -/ +def Adjunction.ofLeftAdjoint (left : C ⥤ D) [IsLeftAdjoint left] : + Adjunction left (rightAdjoint left) := + IsLeftAdjoint.adj +#align category_theory.adjunction.of_left_adjoint CategoryTheory.Adjunction.ofLeftAdjoint + +/-- The adjunction associated to a functor known to be a right adjoint. -/ +def Adjunction.ofRightAdjoint (right : C ⥤ D) [IsRightAdjoint right] : + Adjunction (leftAdjoint right) right := + IsRightAdjoint.adj +#align category_theory.adjunction.of_right_adjoint CategoryTheory.Adjunction.ofRightAdjoint + +namespace Adjunction + +restate_axiom hom_equiv_unit' + +restate_axiom hom_equiv_counit' + +attribute [simp] hom_equiv_unit hom_equiv_counit + +section + +variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) {X' X : C} {Y Y' : D} + +theorem homEquiv_id (X : C) : adj.homEquiv X _ (𝟙 _) = adj.Unit.app X := by simp +#align category_theory.adjunction.hom_equiv_id CategoryTheory.Adjunction.homEquiv_id + +theorem homEquiv_symm_id (X : D) : (adj.homEquiv _ X).symm (𝟙 _) = adj.counit.app X := by simp +#align category_theory.adjunction.hom_equiv_symm_id CategoryTheory.Adjunction.homEquiv_symm_id + +@[simp] +theorem homEquiv_naturality_left_symm (f : X' ⟶ X) (g : X ⟶ G.obj Y) : + (adj.homEquiv X' Y).symm (f ≫ g) = F.map f ≫ (adj.homEquiv X Y).symm g := by + rw [hom_equiv_counit, F.map_comp, assoc, adj.hom_equiv_counit.symm] +#align category_theory.adjunction.hom_equiv_naturality_left_symm CategoryTheory.Adjunction.homEquiv_naturality_left_symm + +@[simp] +theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : + (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by + rw [← Equiv.eq_symm_apply] <;> simp [-hom_equiv_unit] +#align category_theory.adjunction.hom_equiv_naturality_left CategoryTheory.Adjunction.homEquiv_naturality_left + +@[simp] +theorem homEquiv_naturality_right (f : F.obj X ⟶ Y) (g : Y ⟶ Y') : + (adj.homEquiv X Y') (f ≫ g) = (adj.homEquiv X Y) f ≫ G.map g := by + rw [hom_equiv_unit, G.map_comp, ← assoc, ← hom_equiv_unit] +#align category_theory.adjunction.hom_equiv_naturality_right CategoryTheory.Adjunction.homEquiv_naturality_right + +@[simp] +theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : + (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by + rw [Equiv.symm_apply_eq] <;> simp [-hom_equiv_counit] +#align category_theory.adjunction.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.homEquiv_naturality_right_symm + +@[simp] +theorem left_triangle : whiskerRight adj.Unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := + by + ext; dsimp + erw [← adj.hom_equiv_counit, Equiv.symm_apply_eq, adj.hom_equiv_unit] + simp +#align category_theory.adjunction.left_triangle CategoryTheory.Adjunction.left_triangle + +@[simp] +theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = NatTrans.id _ := + by + ext; dsimp + erw [← adj.hom_equiv_unit, ← Equiv.eq_symm_apply, adj.hom_equiv_counit] + simp +#align category_theory.adjunction.right_triangle CategoryTheory.Adjunction.right_triangle + +@[simp, reassoc.1] +theorem left_triangle_components : + F.map (adj.Unit.app X) ≫ adj.counit.app (F.obj X) = 𝟙 (F.obj X) := + congr_arg (fun t : NatTrans _ (𝟭 C ⋙ F) => t.app X) adj.left_triangle +#align category_theory.adjunction.left_triangle_components CategoryTheory.Adjunction.left_triangle_components + +@[simp, reassoc.1] +theorem right_triangle_components {Y : D} : + adj.Unit.app (G.obj Y) ≫ G.map (adj.counit.app Y) = 𝟙 (G.obj Y) := + congr_arg (fun t : NatTrans _ (G ⋙ 𝟭 C) => t.app Y) adj.right_triangle +#align category_theory.adjunction.right_triangle_components CategoryTheory.Adjunction.right_triangle_components + +@[simp, reassoc.1] +theorem counit_naturality {X Y : D} (f : X ⟶ Y) : + F.map (G.map f) ≫ adj.counit.app Y = adj.counit.app X ≫ f := + adj.counit.naturality f +#align category_theory.adjunction.counit_naturality CategoryTheory.Adjunction.counit_naturality + +@[simp, reassoc.1] +theorem unit_naturality {X Y : C} (f : X ⟶ Y) : + adj.Unit.app X ≫ G.map (F.map f) = f ≫ adj.Unit.app Y := + (adj.Unit.naturality f).symm +#align category_theory.adjunction.unit_naturality CategoryTheory.Adjunction.unit_naturality + +theorem homEquiv_apply_eq {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B) : + adj.homEquiv A B f = g ↔ f = (adj.homEquiv A B).symm g := + ⟨fun h => by + cases h + simp, fun h => by + cases h + simp⟩ +#align category_theory.adjunction.hom_equiv_apply_eq CategoryTheory.Adjunction.homEquiv_apply_eq + +theorem eq_homEquiv_apply {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B) : + g = adj.homEquiv A B f ↔ (adj.homEquiv A B).symm g = f := + ⟨fun h => by + cases h + simp, fun h => by + cases h + simp⟩ +#align category_theory.adjunction.eq_hom_equiv_apply CategoryTheory.Adjunction.eq_homEquiv_apply + +end + +end Adjunction + +namespace Adjunction + +/-- This is an auxiliary data structure useful for constructing adjunctions. +See `adjunction.mk_of_hom_equiv`. +This structure won't typically be used anywhere else. +-/ +@[nolint has_nonempty_instance] +structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where + homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) + homEquiv_naturality_left_symm' : + ∀ {X' X Y} (f : X' ⟶ X) (g : X ⟶ G.obj Y), + (hom_equiv X' Y).symm (f ≫ g) = F.map f ≫ (hom_equiv X Y).symm g := by + obviously + homEquiv_naturality_right' : + ∀ {X Y Y'} (f : F.obj X ⟶ Y) (g : Y ⟶ Y'), + (hom_equiv X Y') (f ≫ g) = (hom_equiv X Y) f ≫ G.map g := by + obviously +#align category_theory.adjunction.core_hom_equiv CategoryTheory.Adjunction.CoreHomEquiv + +namespace CoreHomEquiv + +restate_axiom hom_equiv_naturality_left_symm' + +restate_axiom hom_equiv_naturality_right' + +attribute [simp] hom_equiv_naturality_left_symm hom_equiv_naturality_right + +variable {F : C ⥤ D} {G : D ⥤ C} (adj : CoreHomEquiv F G) {X' X : C} {Y Y' : D} + +@[simp] +theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : + (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by + rw [← Equiv.eq_symm_apply] <;> simp +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_left CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_left + +@[simp] +theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : + (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by + rw [Equiv.symm_apply_eq] <;> simp +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_right_symm + +end CoreHomEquiv + +/-- This is an auxiliary data structure useful for constructing adjunctions. +See `adjunction.mk_of_unit_counit`. +This structure won't typically be used anywhere else. +-/ +@[nolint has_nonempty_instance] +structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where + Unit : 𝟭 C ⟶ F.comp G + counit : G.comp F ⟶ 𝟭 D + left_triangle' : + whiskerRight Unit F ≫ (Functor.associator F G F).Hom ≫ whiskerLeft F counit = + NatTrans.id (𝟭 C ⋙ F) := by + obviously + right_triangle' : + whiskerLeft G Unit ≫ (Functor.associator G F G).inv ≫ whiskerRight counit G = + NatTrans.id (G ⋙ 𝟭 C) := by + obviously +#align category_theory.adjunction.core_unit_counit CategoryTheory.Adjunction.CoreUnitCounit + +namespace CoreUnitCounit + +restate_axiom left_triangle' + +restate_axiom right_triangle' + +attribute [simp] left_triangle right_triangle + +end CoreUnitCounit + +variable {F : C ⥤ D} {G : D ⥤ C} + +/-- Construct an adjunction between `F` and `G` out of a natural bijection between each +`F.obj X ⟶ Y` and `X ⟶ G.obj Y`. -/ +@[simps] +def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := + {-- See note [dsimp, simp]. + adj with + Unit := + { app := fun X => (adj.homEquiv X (F.obj X)) (𝟙 (F.obj X)) + naturality' := by + intros + erw [← adj.hom_equiv_naturality_left, ← adj.hom_equiv_naturality_right] + dsimp; simp } + counit := + { app := fun Y => (adj.homEquiv _ _).invFun (𝟙 (G.obj Y)) + naturality' := by + intros + erw [← adj.hom_equiv_naturality_left_symm, ← adj.hom_equiv_naturality_right_symm] + dsimp; simp } + homEquiv_unit' := fun X Y f => by erw [← adj.hom_equiv_naturality_right] <;> simp + homEquiv_counit' := fun X Y f => by erw [← adj.hom_equiv_naturality_left_symm] <;> simp } +#align category_theory.adjunction.mk_of_hom_equiv CategoryTheory.Adjunction.mkOfHomEquiv + +/-- Construct an adjunction between functors `F` and `G` given a unit and counit for the adjunction +satisfying the triangle identities. -/ +@[simps] +def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := + { adj with + homEquiv := fun X Y => + { toFun := fun f => adj.Unit.app X ≫ G.map f + invFun := fun g => F.map g ≫ adj.counit.app Y + left_inv := fun f => by + change F.map (_ ≫ _) ≫ _ = _ + rw [F.map_comp, assoc, ← functor.comp_map, adj.counit.naturality, ← assoc] + convert id_comp f + have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.left_triangle + dsimp at t + simp only [id_comp] at t + exact t + right_inv := fun g => by + change _ ≫ G.map (_ ≫ _) = _ + rw [G.map_comp, ← assoc, ← functor.comp_map, ← adj.unit.naturality, assoc] + convert comp_id g + have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.right_triangle + dsimp at t + simp only [id_comp] at t + exact t } } +#align category_theory.adjunction.mk_of_unit_counit CategoryTheory.Adjunction.mkOfUnitCounit + +/-- The adjunction between the identity functor on a category and itself. -/ +def id : 𝟭 C ⊣ 𝟭 C where + homEquiv X Y := Equiv.refl _ + Unit := 𝟙 _ + counit := 𝟙 _ +#align category_theory.adjunction.id CategoryTheory.Adjunction.id + +-- Satisfy the inhabited linter. +instance : Inhabited (Adjunction (𝟭 C) (𝟭 C)) := + ⟨id⟩ + +/-- If F and G are naturally isomorphic functors, establish an equivalence of hom-sets. -/ +@[simps] +def equivHomsetLeftOfNatIso {F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} : + (F.obj X ⟶ Y) ≃ (F'.obj X ⟶ Y) + where + toFun f := iso.inv.app _ ≫ f + invFun g := iso.Hom.app _ ≫ g + left_inv f := by simp + right_inv g := by simp +#align category_theory.adjunction.equiv_homset_left_of_nat_iso CategoryTheory.Adjunction.equivHomsetLeftOfNatIso + +/-- If G and H are naturally isomorphic functors, establish an equivalence of hom-sets. -/ +@[simps] +def equivHomsetRightOfNatIso {G G' : D ⥤ C} (iso : G ≅ G') {X : C} {Y : D} : + (X ⟶ G.obj Y) ≃ (X ⟶ G'.obj Y) + where + toFun f := f ≫ iso.Hom.app _ + invFun g := g ≫ iso.inv.app _ + left_inv f := by simp + right_inv g := by simp +#align category_theory.adjunction.equiv_homset_right_of_nat_iso CategoryTheory.Adjunction.equivHomsetRightOfNatIso + +/-- Transport an adjunction along an natural isomorphism on the left. -/ +def ofNatIsoLeft {F G : C ⥤ D} {H : D ⥤ C} (adj : F ⊣ H) (iso : F ≅ G) : G ⊣ H := + Adjunction.mkOfHomEquiv + { homEquiv := fun X Y => (equivHomsetLeftOfNatIso iso.symm).trans (adj.homEquiv X Y) } +#align category_theory.adjunction.of_nat_iso_left CategoryTheory.Adjunction.ofNatIsoLeft + +/-- Transport an adjunction along an natural isomorphism on the right. -/ +def ofNatIsoRight {F : C ⥤ D} {G H : D ⥤ C} (adj : F ⊣ G) (iso : G ≅ H) : F ⊣ H := + Adjunction.mkOfHomEquiv + { homEquiv := fun X Y => (adj.homEquiv X Y).trans (equivHomsetRightOfNatIso iso) } +#align category_theory.adjunction.of_nat_iso_right CategoryTheory.Adjunction.ofNatIsoRight + +/-- Transport being a right adjoint along a natural isomorphism. -/ +def rightAdjointOfNatIso {F G : C ⥤ D} (h : F ≅ G) [r : IsRightAdjoint F] : IsRightAdjoint G + where + left := r.left + adj := ofNatIsoRight r.adj h +#align category_theory.adjunction.right_adjoint_of_nat_iso CategoryTheory.Adjunction.rightAdjointOfNatIso + +/-- Transport being a left adjoint along a natural isomorphism. -/ +def leftAdjointOfNatIso {F G : C ⥤ D} (h : F ≅ G) [r : IsLeftAdjoint F] : IsLeftAdjoint G + where + right := r.right + adj := ofNatIsoLeft r.adj h +#align category_theory.adjunction.left_adjoint_of_nat_iso CategoryTheory.Adjunction.leftAdjointOfNatIso + +section + +variable {E : Type u₃} [ℰ : Category.{v₃} E] {H : D ⥤ E} {I : E ⥤ D} + +/-- Composition of adjunctions. + +See . +-/ +def comp (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) : F ⋙ H ⊣ I ⋙ G + where + homEquiv X Z := Equiv.trans (adj₂.homEquiv _ _) (adj₁.homEquiv _ _) + Unit := adj₁.Unit ≫ (whiskerLeft F <| whiskerRight adj₂.Unit G) ≫ (Functor.associator _ _ _).inv + counit := + (Functor.associator _ _ _).Hom ≫ (whiskerLeft I <| whiskerRight adj₁.counit H) ≫ adj₂.counit +#align category_theory.adjunction.comp CategoryTheory.Adjunction.comp + +/-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/ +instance leftAdjointOfComp {E : Type u₃} [ℰ : Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) + [Fl : IsLeftAdjoint F] [Gl : IsLeftAdjoint G] : IsLeftAdjoint (F ⋙ G) + where + right := Gl.right ⋙ Fl.right + adj := Fl.adj.comp Gl.adj +#align category_theory.adjunction.left_adjoint_of_comp CategoryTheory.Adjunction.leftAdjointOfComp + +/-- If `F` and `G` are right adjoints then `F ⋙ G` is a right adjoint too. -/ +instance rightAdjointOfComp {E : Type u₃} [ℰ : Category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E} + [Fr : IsRightAdjoint F] [Gr : IsRightAdjoint G] : IsRightAdjoint (F ⋙ G) + where + left := Gr.left ⋙ Fr.left + adj := Gr.adj.comp Fr.adj +#align category_theory.adjunction.right_adjoint_of_comp CategoryTheory.Adjunction.rightAdjointOfComp + +end + +section ConstructLeft + +-- Construction of a left adjoint. In order to construct a left +-- adjoint to a functor G : D → C, it suffices to give the object part +-- of a functor F : C → D together with isomorphisms Hom(FX, Y) ≃ +-- Hom(X, GY) natural in Y. The action of F on morphisms can be +-- constructed from this data. +variable {F_obj : C → D} {G} + +variable (e : ∀ X Y, (F_obj X ⟶ Y) ≃ (X ⟶ G.obj Y)) + +variable (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) + +include he + +private theorem he' {X Y Y'} (f g) : (e X Y').symm (f ≫ G.map g) = (e X Y).symm f ≫ g := by + intros <;> rw [Equiv.symm_apply_eq, he] <;> simp +#align category_theory.adjunction.he' category_theory.adjunction.he' + +/-- Construct a left adjoint functor to `G`, given the functor's value on objects `F_obj` and +a bijection `e` between `F_obj X ⟶ Y` and `X ⟶ G.obj Y` satisfying a naturality law +`he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g`. +Dual to `right_adjoint_of_equiv`. -/ +@[simps] +def leftAdjointOfEquiv : C ⥤ D where + obj := F_obj + map X X' f := (e X (F_obj X')).symm (f ≫ e X' (F_obj X') (𝟙 _)) + map_comp' X X' X'' f f' := + by + rw [Equiv.symm_apply_eq, he, Equiv.apply_symm_apply] + conv => + rhs + rw [assoc, ← he, id_comp, Equiv.apply_symm_apply] + simp +#align category_theory.adjunction.left_adjoint_of_equiv CategoryTheory.Adjunction.leftAdjointOfEquiv + +/-- Show that the functor given by `left_adjoint_of_equiv` is indeed left adjoint to `G`. Dual +to `adjunction_of_equiv_right`. -/ +@[simps] +def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := + mkOfHomEquiv + { homEquiv := e + homEquiv_naturality_left_symm' := by + intros + erw [← he' e he, ← Equiv.apply_eq_iff_eq] + simp [(he _ _ _ _ _).symm] } +#align category_theory.adjunction.adjunction_of_equiv_left CategoryTheory.Adjunction.adjunctionOfEquivLeft + +end ConstructLeft + +section ConstructRight + +-- Construction of a right adjoint, analogous to the above. +variable {F} {G_obj : D → C} + +variable (e : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G_obj Y)) + +variable (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) + +include he + +private theorem he' {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm (f ≫ g) := by + intros <;> rw [Equiv.eq_symm_apply, he] <;> simp +#align category_theory.adjunction.he' category_theory.adjunction.he' + +/-- Construct a right adjoint functor to `F`, given the functor's value on objects `G_obj` and +a bijection `e` between `F.obj X ⟶ Y` and `X ⟶ G_obj Y` satisfying a naturality law +`he : ∀ X Y Y' g h, e X' Y (F.map f ≫ g) = f ≫ e X Y g`. +Dual to `left_adjoint_of_equiv`. -/ +@[simps] +def rightAdjointOfEquiv : D ⥤ C where + obj := G_obj + map Y Y' g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) + map_comp' Y Y' Y'' g g' := + by + rw [← Equiv.eq_symm_apply, ← he' e he, Equiv.symm_apply_apply] + conv => + rhs + rw [← assoc, he' e he, comp_id, Equiv.symm_apply_apply] + simp +#align category_theory.adjunction.right_adjoint_of_equiv CategoryTheory.Adjunction.rightAdjointOfEquiv + +/-- Show that the functor given by `right_adjoint_of_equiv` is indeed right adjoint to `F`. Dual +to `adjunction_of_equiv_left`. -/ +@[simps] +def adjunctionOfEquivRight : F ⊣ rightAdjointOfEquiv e he := + mkOfHomEquiv + { homEquiv := e + homEquiv_naturality_left_symm' := by intros <;> rw [Equiv.symm_apply_eq, he] <;> simp + homEquiv_naturality_right' := by + intro X Y Y' g h + erw [← he, Equiv.apply_eq_iff_eq, ← assoc, he' e he, comp_id, Equiv.symm_apply_apply] } +#align category_theory.adjunction.adjunction_of_equiv_right CategoryTheory.Adjunction.adjunctionOfEquivRight + +end ConstructRight + +/-- +If the unit and counit of a given adjunction are (pointwise) isomorphisms, then we can upgrade the +adjunction to an equivalence. +-/ +@[simps] +noncomputable def toEquivalence (adj : F ⊣ G) [∀ X, IsIso (adj.Unit.app X)] + [∀ Y, IsIso (adj.counit.app Y)] : C ≌ D + where + Functor := F + inverse := G + unitIso := NatIso.ofComponents (fun X => asIso (adj.Unit.app X)) (by simp) + counitIso := NatIso.ofComponents (fun Y => asIso (adj.counit.app Y)) (by simp) +#align category_theory.adjunction.to_equivalence CategoryTheory.Adjunction.toEquivalence + +/-- +If the unit and counit for the adjunction corresponding to a right adjoint functor are (pointwise) +isomorphisms, then the functor is an equivalence of categories. +-/ +@[simps] +noncomputable def isRightAdjointToIsEquivalence [IsRightAdjoint G] + [∀ X, IsIso ((Adjunction.ofRightAdjoint G).Unit.app X)] + [∀ Y, IsIso ((Adjunction.ofRightAdjoint G).counit.app Y)] : IsEquivalence G := + IsEquivalence.ofEquivalenceInverse (Adjunction.ofRightAdjoint G).toEquivalence +#align category_theory.adjunction.is_right_adjoint_to_is_equivalence CategoryTheory.Adjunction.isRightAdjointToIsEquivalence + +end Adjunction + +open Adjunction + +namespace Equivalence + +/-- The adjunction given by an equivalence of categories. (To obtain the opposite adjunction, +simply use `e.symm.to_adjunction`. -/ +def toAdjunction (e : C ≌ D) : e.Functor ⊣ e.inverse := + mkOfUnitCounit + ⟨e.Unit, e.counit, by + ext + dsimp + simp only [id_comp] + exact e.functor_unit_comp _, by + ext + dsimp + simp only [id_comp] + exact e.unit_inverse_comp _⟩ +#align category_theory.equivalence.to_adjunction CategoryTheory.Equivalence.toAdjunction + +@[simp] +theorem asEquivalence_toAdjunction_unit {e : C ≌ D} : + e.Functor.asEquivalence.toAdjunction.Unit = e.Unit := + rfl +#align category_theory.equivalence.as_equivalence_to_adjunction_unit CategoryTheory.Equivalence.asEquivalence_toAdjunction_unit + +@[simp] +theorem asEquivalence_toAdjunction_counit {e : C ≌ D} : + e.Functor.asEquivalence.toAdjunction.counit = e.counit := + rfl +#align category_theory.equivalence.as_equivalence_to_adjunction_counit CategoryTheory.Equivalence.asEquivalence_toAdjunction_counit + +end Equivalence + +namespace Functor + +/-- An equivalence `E` is left adjoint to its inverse. -/ +def adjunction (E : C ⥤ D) [IsEquivalence E] : E ⊣ E.inv := + E.asEquivalence.toAdjunction +#align category_theory.functor.adjunction CategoryTheory.Functor.adjunction + +/-- If `F` is an equivalence, it's a left adjoint. -/ +instance (priority := 10) leftAdjointOfEquivalence {F : C ⥤ D} [IsEquivalence F] : IsLeftAdjoint F + where + right := _ + adj := Functor.adjunction F +#align category_theory.functor.left_adjoint_of_equivalence CategoryTheory.Functor.leftAdjointOfEquivalence + +@[simp] +theorem rightAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : rightAdjoint F = inv F := + rfl +#align category_theory.functor.right_adjoint_of_is_equivalence CategoryTheory.Functor.rightAdjoint_of_isEquivalence + +/-- If `F` is an equivalence, it's a right adjoint. -/ +instance (priority := 10) rightAdjointOfEquivalence {F : C ⥤ D} [IsEquivalence F] : IsRightAdjoint F + where + left := _ + adj := Functor.adjunction F.inv +#align category_theory.functor.right_adjoint_of_equivalence CategoryTheory.Functor.rightAdjointOfEquivalence + +@[simp] +theorem leftAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : leftAdjoint F = inv F := + rfl +#align category_theory.functor.left_adjoint_of_is_equivalence CategoryTheory.Functor.leftAdjoint_of_isEquivalence + +end Functor + +end CategoryTheory + From e0de19809e25a0da91f9b856757b72ef74403e29 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Thu, 9 Feb 2023 13:49:59 -0500 Subject: [PATCH 26/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Adjunction/Basic.lean | 8 +++----- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index e0a7fbf82f764..ae08d42e274e3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -212,6 +212,7 @@ import Mathlib.Algebra.Support import Mathlib.Algebra.Tropical.Basic import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice +import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Bicategory.Basic import Mathlib.CategoryTheory.Category.Basic import Mathlib.CategoryTheory.Category.KleisliCat diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 57a6d3985bbd8..aa084ed165080 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -8,7 +8,7 @@ Authors: Reid Barton, Johan Commelin, Bhavik Mehta ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Equivalence +import Mathlib.CategoryTheory.Equivalence /-! # Adjunctions between functors @@ -148,16 +148,14 @@ theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : #align category_theory.adjunction.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.homEquiv_naturality_right_symm @[simp] -theorem left_triangle : whiskerRight adj.Unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := - by +theorem left_triangle : whiskerRight adj.Unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := by ext; dsimp erw [← adj.hom_equiv_counit, Equiv.symm_apply_eq, adj.hom_equiv_unit] simp #align category_theory.adjunction.left_triangle CategoryTheory.Adjunction.left_triangle @[simp] -theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = NatTrans.id _ := - by +theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = NatTrans.id _ := by ext; dsimp erw [← adj.hom_equiv_unit, ← Equiv.eq_symm_apply, adj.hom_equiv_counit] simp From cfe5dc1f7ceede7309c956343bdca17cbe0de5c7 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Fri, 10 Feb 2023 20:43:41 -0500 Subject: [PATCH 27/66] some progress --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 90 ++++++++++---------- 1 file changed, 46 insertions(+), 44 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index aa084ed165080..566254f9c5e1f 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -44,7 +44,8 @@ open Category -- declare the `v`'s first; see `category_theory.category` for an explanation universe v₁ v₂ v₃ u₁ u₂ u₃ -attribute [local elab_without_expected_type] whisker_left whisker_right +-- Porting Note: `elab_without_expected_type` cannot be a local attribute +-- attribute [local elab_without_expected_type] whiskerLeft whiskerRight variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] @@ -64,8 +65,8 @@ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) Unit : 𝟭 C ⟶ F.comp G counit : G.comp F ⟶ 𝟭 D - homEquiv_unit' : ∀ {X Y f}, (hom_equiv X Y) f = (Unit : _ ⟶ _).app X ≫ G.map f := by obviously - homEquiv_counit' : ∀ {X Y g}, (hom_equiv X Y).symm g = F.map g ≫ counit.app Y := by obviously + homEquiv_unit' : ∀ {X Y f}, (homEquiv X Y) f = (Unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat + homEquiv_counit' : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat #align category_theory.adjunction CategoryTheory.Adjunction -- mathport name: «expr ⊣ » @@ -107,11 +108,11 @@ def Adjunction.ofRightAdjoint (right : C ⥤ D) [IsRightAdjoint right] : namespace Adjunction -restate_axiom hom_equiv_unit' +restate_axiom homEquiv_unit' -restate_axiom hom_equiv_counit' +restate_axiom homEquiv_counit' -attribute [simp] hom_equiv_unit hom_equiv_counit +attribute [simp] homEquiv_unit homEquiv_counit section @@ -126,60 +127,60 @@ theorem homEquiv_symm_id (X : D) : (adj.homEquiv _ X).symm (𝟙 _) = adj.counit @[simp] theorem homEquiv_naturality_left_symm (f : X' ⟶ X) (g : X ⟶ G.obj Y) : (adj.homEquiv X' Y).symm (f ≫ g) = F.map f ≫ (adj.homEquiv X Y).symm g := by - rw [hom_equiv_counit, F.map_comp, assoc, adj.hom_equiv_counit.symm] + rw [homEquiv_counit, F.map_comp, assoc, adj.homEquiv_counit.symm] #align category_theory.adjunction.hom_equiv_naturality_left_symm CategoryTheory.Adjunction.homEquiv_naturality_left_symm @[simp] theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by - rw [← Equiv.eq_symm_apply] <;> simp [-hom_equiv_unit] + --rw [← Equiv.eq_symm_apply] ; simp [-homEquiv_unit] #align category_theory.adjunction.hom_equiv_naturality_left CategoryTheory.Adjunction.homEquiv_naturality_left @[simp] theorem homEquiv_naturality_right (f : F.obj X ⟶ Y) (g : Y ⟶ Y') : (adj.homEquiv X Y') (f ≫ g) = (adj.homEquiv X Y) f ≫ G.map g := by - rw [hom_equiv_unit, G.map_comp, ← assoc, ← hom_equiv_unit] + rw [homEquiv_unit, G.map_comp, ← assoc, ← homEquiv_unit] #align category_theory.adjunction.hom_equiv_naturality_right CategoryTheory.Adjunction.homEquiv_naturality_right @[simp] theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by - rw [Equiv.symm_apply_eq] <;> simp [-hom_equiv_counit] + rw [Equiv.symm_apply_eq] ; simp [-homEquiv_counit] #align category_theory.adjunction.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.homEquiv_naturality_right_symm @[simp] theorem left_triangle : whiskerRight adj.Unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := by ext; dsimp - erw [← adj.hom_equiv_counit, Equiv.symm_apply_eq, adj.hom_equiv_unit] + erw [← adj.homEquiv_counit, Equiv.symm_apply_eq, adj.homEquiv_unit] simp #align category_theory.adjunction.left_triangle CategoryTheory.Adjunction.left_triangle @[simp] theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = NatTrans.id _ := by ext; dsimp - erw [← adj.hom_equiv_unit, ← Equiv.eq_symm_apply, adj.hom_equiv_counit] + erw [← adj.homEquiv_unit, ← Equiv.eq_symm_apply, adj.homEquiv_counit] simp #align category_theory.adjunction.right_triangle CategoryTheory.Adjunction.right_triangle -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem left_triangle_components : F.map (adj.Unit.app X) ≫ adj.counit.app (F.obj X) = 𝟙 (F.obj X) := congr_arg (fun t : NatTrans _ (𝟭 C ⋙ F) => t.app X) adj.left_triangle #align category_theory.adjunction.left_triangle_components CategoryTheory.Adjunction.left_triangle_components -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem right_triangle_components {Y : D} : adj.Unit.app (G.obj Y) ≫ G.map (adj.counit.app Y) = 𝟙 (G.obj Y) := congr_arg (fun t : NatTrans _ (G ⋙ 𝟭 C) => t.app Y) adj.right_triangle #align category_theory.adjunction.right_triangle_components CategoryTheory.Adjunction.right_triangle_components -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem counit_naturality {X Y : D} (f : X ⟶ Y) : F.map (G.map f) ≫ adj.counit.app Y = adj.counit.app X ≫ f := adj.counit.naturality f #align category_theory.adjunction.counit_naturality CategoryTheory.Adjunction.counit_naturality -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem unit_naturality {X Y : C} (f : X ⟶ Y) : adj.Unit.app X ≫ G.map (F.map f) = f ≫ adj.Unit.app Y := (adj.Unit.naturality f).symm @@ -213,39 +214,40 @@ namespace Adjunction See `adjunction.mk_of_hom_equiv`. This structure won't typically be used anywhere else. -/ -@[nolint has_nonempty_instance] +-- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) +-- @[nolint has_nonempty_instance] structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) homEquiv_naturality_left_symm' : ∀ {X' X Y} (f : X' ⟶ X) (g : X ⟶ G.obj Y), - (hom_equiv X' Y).symm (f ≫ g) = F.map f ≫ (hom_equiv X Y).symm g := by - obviously + (homEquiv X' Y).symm (f ≫ g) = F.map f ≫ (homEquiv X Y).symm g := by + aesop_cat homEquiv_naturality_right' : ∀ {X Y Y'} (f : F.obj X ⟶ Y) (g : Y ⟶ Y'), - (hom_equiv X Y') (f ≫ g) = (hom_equiv X Y) f ≫ G.map g := by - obviously + (homEquiv X Y') (f ≫ g) = (homEquiv X Y) f ≫ G.map g := by + aesop_cat #align category_theory.adjunction.core_hom_equiv CategoryTheory.Adjunction.CoreHomEquiv namespace CoreHomEquiv -restate_axiom hom_equiv_naturality_left_symm' +restate_axiom homEquiv_naturality_left_symm' -restate_axiom hom_equiv_naturality_right' +restate_axiom homEquiv_naturality_right' -attribute [simp] hom_equiv_naturality_left_symm hom_equiv_naturality_right +attribute [simp] homEquiv_naturality_left_symm homEquiv_naturality_right variable {F : C ⥤ D} {G : D ⥤ C} (adj : CoreHomEquiv F G) {X' X : C} {Y Y' : D} @[simp] theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by - rw [← Equiv.eq_symm_apply] <;> simp + rw [← Equiv.eq_symm_apply] ; simp #align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_left CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_left @[simp] theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by - rw [Equiv.symm_apply_eq] <;> simp + rw [Equiv.symm_apply_eq] ; simp #align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_right_symm end CoreHomEquiv @@ -254,18 +256,19 @@ end CoreHomEquiv See `adjunction.mk_of_unit_counit`. This structure won't typically be used anywhere else. -/ -@[nolint has_nonempty_instance] +-- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) +-- @[nolint has_nonempty_instance] structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where Unit : 𝟭 C ⟶ F.comp G counit : G.comp F ⟶ 𝟭 D left_triangle' : - whiskerRight Unit F ≫ (Functor.associator F G F).Hom ≫ whiskerLeft F counit = + whiskerRight Unit F ≫ (Functor.associator F G F).hom ≫ whiskerLeft F counit = NatTrans.id (𝟭 C ⋙ F) := by - obviously + aesop_cat right_triangle' : whiskerLeft G Unit ≫ (Functor.associator G F G).inv ≫ whiskerRight counit G = NatTrans.id (G ⋙ 𝟭 C) := by - obviously + aesop_cat #align category_theory.adjunction.core_unit_counit CategoryTheory.Adjunction.CoreUnitCounit namespace CoreUnitCounit @@ -288,18 +291,20 @@ def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := adj with Unit := { app := fun X => (adj.homEquiv X (F.obj X)) (𝟙 (F.obj X)) - naturality' := by + naturality := by intros - erw [← adj.hom_equiv_naturality_left, ← adj.hom_equiv_naturality_right] + erw [← adj.homEquiv_naturality_left, ← adj.homEquiv_naturality_right] dsimp; simp } counit := { app := fun Y => (adj.homEquiv _ _).invFun (𝟙 (G.obj Y)) - naturality' := by + naturality := by intros - erw [← adj.hom_equiv_naturality_left_symm, ← adj.hom_equiv_naturality_right_symm] + erw [← adj.homEquiv_naturality_left_symm, ← adj.homEquiv_naturality_right_symm] dsimp; simp } - homEquiv_unit' := fun X Y f => by erw [← adj.hom_equiv_naturality_right] <;> simp - homEquiv_counit' := fun X Y f => by erw [← adj.hom_equiv_naturality_left_symm] <;> simp } + -- homEquiv_unit' := by sorry -- erw [← adj.homEquiv_naturality_right] ; simp + -- homEquiv_counit' := by sorry + -- homEquiv_counit' := fun X Y f => by erw [← adj.hom_equiv_naturality_left_symm] <;> simp } + } #align category_theory.adjunction.mk_of_hom_equiv CategoryTheory.Adjunction.mkOfHomEquiv /-- Construct an adjunction between functors `F` and `G` given a unit and counit for the adjunction @@ -312,7 +317,7 @@ def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := invFun := fun g => F.map g ≫ adj.counit.app Y left_inv := fun f => by change F.map (_ ≫ _) ≫ _ = _ - rw [F.map_comp, assoc, ← functor.comp_map, adj.counit.naturality, ← assoc] + rw [F.map_comp, assoc, ← Functor.comp_map, adj.counit.naturality, ← assoc] convert id_comp f have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.left_triangle dsimp at t @@ -320,7 +325,7 @@ def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := exact t right_inv := fun g => by change _ ≫ G.map (_ ≫ _) = _ - rw [G.map_comp, ← assoc, ← functor.comp_map, ← adj.unit.naturality, assoc] + rw [G.map_comp, ← assoc, ← Functor.comp_map, ← adj.unit.naturality, assoc] convert comp_id g have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.right_triangle dsimp at t @@ -404,7 +409,7 @@ def comp (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) : F ⋙ H ⊣ I ⋙ G #align category_theory.adjunction.comp CategoryTheory.Adjunction.comp /-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/ -instance leftAdjointOfComp {E : Type u₃} [ℰ : Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) +instance leftAdjointOfComp {E : Type u₃} [Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) [Fl : IsLeftAdjoint F] [Gl : IsLeftAdjoint G] : IsLeftAdjoint (F ⋙ G) where right := Gl.right ⋙ Fl.right @@ -412,7 +417,7 @@ instance leftAdjointOfComp {E : Type u₃} [ℰ : Category.{v₃} E] (F : C ⥤ #align category_theory.adjunction.left_adjoint_of_comp CategoryTheory.Adjunction.leftAdjointOfComp /-- If `F` and `G` are right adjoints then `F ⋙ G` is a right adjoint too. -/ -instance rightAdjointOfComp {E : Type u₃} [ℰ : Category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E} +instance rightAdjointOfComp {E : Type u₃} [Category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E} [Fr : IsRightAdjoint F] [Gr : IsRightAdjoint G] : IsRightAdjoint (F ⋙ G) where left := Gr.left ⋙ Fr.left @@ -428,14 +433,12 @@ section ConstructLeft -- of a functor F : C → D together with isomorphisms Hom(FX, Y) ≃ -- Hom(X, GY) natural in Y. The action of F on morphisms can be -- constructed from this data. -variable {F_obj : C → D} {G} +variable {F_obj : C → D} variable (e : ∀ X Y, (F_obj X ⟶ Y) ≃ (X ⟶ G.obj Y)) variable (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) -include he - private theorem he' {X Y Y'} (f g) : (e X Y').symm (f ≫ G.map g) = (e X Y).symm f ≫ g := by intros <;> rw [Equiv.symm_apply_eq, he] <;> simp #align category_theory.adjunction.he' category_theory.adjunction.he' @@ -611,4 +614,3 @@ theorem leftAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : leftAdjoi end Functor end CategoryTheory - From 3278d7dfbbe23ee4a65f304d478c30b1733ab5ae Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Sat, 11 Feb 2023 06:49:21 -0500 Subject: [PATCH 28/66] only one error left --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 138 +++++++++---------- 1 file changed, 68 insertions(+), 70 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 566254f9c5e1f..0681424718a59 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -63,9 +63,9 @@ See . -/ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) - Unit : 𝟭 C ⟶ F.comp G + unit : 𝟭 C ⟶ F.comp G counit : G.comp F ⟶ 𝟭 D - homEquiv_unit' : ∀ {X Y f}, (homEquiv X Y) f = (Unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat + homEquiv_unit' : ∀ {X Y f}, (homEquiv X Y) f = (unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat homEquiv_counit' : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat #align category_theory.adjunction CategoryTheory.Adjunction @@ -118,7 +118,7 @@ section variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) {X' X : C} {Y Y' : D} -theorem homEquiv_id (X : C) : adj.homEquiv X _ (𝟙 _) = adj.Unit.app X := by simp +theorem homEquiv_id (X : C) : adj.homEquiv X _ (𝟙 _) = adj.unit.app X := by simp #align category_theory.adjunction.hom_equiv_id CategoryTheory.Adjunction.homEquiv_id theorem homEquiv_symm_id (X : D) : (adj.homEquiv _ X).symm (𝟙 _) = adj.counit.app X := by simp @@ -133,7 +133,7 @@ theorem homEquiv_naturality_left_symm (f : X' ⟶ X) (g : X ⟶ G.obj Y) : @[simp] theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by - --rw [← Equiv.eq_symm_apply] ; simp [-homEquiv_unit] + rw [← Equiv.eq_symm_apply] ; simp only [Equiv.symm_apply_apply,eq_self_iff_true,homEquiv_naturality_left_symm] #align category_theory.adjunction.hom_equiv_naturality_left CategoryTheory.Adjunction.homEquiv_naturality_left @[simp] @@ -145,18 +145,18 @@ theorem homEquiv_naturality_right (f : F.obj X ⟶ Y) (g : Y ⟶ Y') : @[simp] theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by - rw [Equiv.symm_apply_eq] ; simp [-homEquiv_counit] + rw [Equiv.symm_apply_eq] ; simp only [homEquiv_naturality_right,eq_self_iff_true,Equiv.apply_symm_apply] #align category_theory.adjunction.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.homEquiv_naturality_right_symm @[simp] -theorem left_triangle : whiskerRight adj.Unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := by +theorem left_triangle : whiskerRight adj.unit F ≫ whiskerLeft F adj.counit = NatTrans.id _ := by ext; dsimp erw [← adj.homEquiv_counit, Equiv.symm_apply_eq, adj.homEquiv_unit] simp #align category_theory.adjunction.left_triangle CategoryTheory.Adjunction.left_triangle @[simp] -theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = NatTrans.id _ := by +theorem right_triangle : whiskerLeft G adj.unit ≫ whiskerRight adj.counit G = NatTrans.id _ := by ext; dsimp erw [← adj.homEquiv_unit, ← Equiv.eq_symm_apply, adj.homEquiv_counit] simp @@ -164,13 +164,13 @@ theorem right_triangle : whiskerLeft G adj.Unit ≫ whiskerRight adj.counit G = @[reassoc (attr := simp)] theorem left_triangle_components : - F.map (adj.Unit.app X) ≫ adj.counit.app (F.obj X) = 𝟙 (F.obj X) := + F.map (adj.unit.app X) ≫ adj.counit.app (F.obj X) = 𝟙 (F.obj X) := congr_arg (fun t : NatTrans _ (𝟭 C ⋙ F) => t.app X) adj.left_triangle #align category_theory.adjunction.left_triangle_components CategoryTheory.Adjunction.left_triangle_components @[reassoc (attr := simp)] theorem right_triangle_components {Y : D} : - adj.Unit.app (G.obj Y) ≫ G.map (adj.counit.app Y) = 𝟙 (G.obj Y) := + adj.unit.app (G.obj Y) ≫ G.map (adj.counit.app Y) = 𝟙 (G.obj Y) := congr_arg (fun t : NatTrans _ (G ⋙ 𝟭 C) => t.app Y) adj.right_triangle #align category_theory.adjunction.right_triangle_components CategoryTheory.Adjunction.right_triangle_components @@ -182,8 +182,8 @@ theorem counit_naturality {X Y : D} (f : X ⟶ Y) : @[reassoc (attr := simp)] theorem unit_naturality {X Y : C} (f : X ⟶ Y) : - adj.Unit.app X ≫ G.map (F.map f) = f ≫ adj.Unit.app Y := - (adj.Unit.naturality f).symm + adj.unit.app X ≫ G.map (F.map f) = f ≫ adj.unit.app Y := + (adj.unit.naturality f).symm #align category_theory.adjunction.unit_naturality CategoryTheory.Adjunction.unit_naturality theorem homEquiv_apply_eq {A : C} {B : D} (f : F.obj A ⟶ B) (g : A ⟶ G.obj B) : @@ -259,24 +259,20 @@ This structure won't typically be used anywhere else. -- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) -- @[nolint has_nonempty_instance] structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where - Unit : 𝟭 C ⟶ F.comp G + unit : 𝟭 C ⟶ F.comp G counit : G.comp F ⟶ 𝟭 D - left_triangle' : - whiskerRight Unit F ≫ (Functor.associator F G F).hom ≫ whiskerLeft F counit = + left_triangle : + whiskerRight unit F ≫ (Functor.associator F G F).hom ≫ whiskerLeft F counit = NatTrans.id (𝟭 C ⋙ F) := by aesop_cat - right_triangle' : - whiskerLeft G Unit ≫ (Functor.associator G F G).inv ≫ whiskerRight counit G = + right_triangle : + whiskerLeft G unit ≫ (Functor.associator G F G).inv ≫ whiskerRight counit G = NatTrans.id (G ⋙ 𝟭 C) := by aesop_cat #align category_theory.adjunction.core_unit_counit CategoryTheory.Adjunction.CoreUnitCounit namespace CoreUnitCounit -restate_axiom left_triangle' - -restate_axiom right_triangle' - attribute [simp] left_triangle right_triangle end CoreUnitCounit @@ -287,9 +283,9 @@ variable {F : C ⥤ D} {G : D ⥤ C} `F.obj X ⟶ Y` and `X ⟶ G.obj Y`. -/ @[simps] def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := - {-- See note [dsimp, simp]. - adj with - Unit := + -- See note [dsimp, simp]. + { adj with + unit := { app := fun X => (adj.homEquiv X (F.obj X)) (𝟙 (F.obj X)) naturality := by intros @@ -301,25 +297,24 @@ def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := intros erw [← adj.homEquiv_naturality_left_symm, ← adj.homEquiv_naturality_right_symm] dsimp; simp } - -- homEquiv_unit' := by sorry -- erw [← adj.homEquiv_naturality_right] ; simp - -- homEquiv_counit' := by sorry - -- homEquiv_counit' := fun X Y f => by erw [← adj.hom_equiv_naturality_left_symm] <;> simp } + homEquiv_unit' := @fun X Y f => by erw [← adj.homEquiv_naturality_right]; simp + homEquiv_counit' := @fun X Y f => by erw [← adj.homEquiv_naturality_left_symm]; simp } #align category_theory.adjunction.mk_of_hom_equiv CategoryTheory.Adjunction.mkOfHomEquiv /-- Construct an adjunction between functors `F` and `G` given a unit and counit for the adjunction satisfying the triangle identities. -/ -@[simps] +@[simps!] def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := { adj with homEquiv := fun X Y => - { toFun := fun f => adj.Unit.app X ≫ G.map f + { toFun := fun f => adj.unit.app X ≫ G.map f invFun := fun g => F.map g ≫ adj.counit.app Y left_inv := fun f => by change F.map (_ ≫ _) ≫ _ = _ rw [F.map_comp, assoc, ← Functor.comp_map, adj.counit.naturality, ← assoc] convert id_comp f - have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.left_triangle + have t := congrArg (fun (s : NatTrans (𝟭 C ⋙ F) (F ⋙ 𝟭 D)) => s.app X) adj.left_triangle dsimp at t simp only [id_comp] at t exact t @@ -327,7 +322,7 @@ def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := change _ ≫ G.map (_ ≫ _) = _ rw [G.map_comp, ← assoc, ← Functor.comp_map, ← adj.unit.naturality, assoc] convert comp_id g - have t := congr_arg (fun t : nat_trans _ _ => t.app _) adj.right_triangle + have t := congrArg (fun t : NatTrans (G ⋙ 𝟭 C) (𝟭 D ⋙ G) => t.app Y) adj.right_triangle dsimp at t simp only [id_comp] at t exact t } } @@ -336,7 +331,7 @@ def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := /-- The adjunction between the identity functor on a category and itself. -/ def id : 𝟭 C ⊣ 𝟭 C where homEquiv X Y := Equiv.refl _ - Unit := 𝟙 _ + unit := 𝟙 _ counit := 𝟙 _ #align category_theory.adjunction.id CategoryTheory.Adjunction.id @@ -350,7 +345,7 @@ def equivHomsetLeftOfNatIso {F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} : (F.obj X ⟶ Y) ≃ (F'.obj X ⟶ Y) where toFun f := iso.inv.app _ ≫ f - invFun g := iso.Hom.app _ ≫ g + invFun g := iso.hom.app _ ≫ g left_inv f := by simp right_inv g := by simp #align category_theory.adjunction.equiv_homset_left_of_nat_iso CategoryTheory.Adjunction.equivHomsetLeftOfNatIso @@ -360,7 +355,7 @@ def equivHomsetLeftOfNatIso {F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} : def equivHomsetRightOfNatIso {G G' : D ⥤ C} (iso : G ≅ G') {X : C} {Y : D} : (X ⟶ G.obj Y) ≃ (X ⟶ G'.obj Y) where - toFun f := f ≫ iso.Hom.app _ + toFun f := f ≫ iso.hom.app _ invFun g := g ≫ iso.inv.app _ left_inv f := by simp right_inv g := by simp @@ -403,9 +398,9 @@ See . def comp (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) : F ⋙ H ⊣ I ⋙ G where homEquiv X Z := Equiv.trans (adj₂.homEquiv _ _) (adj₁.homEquiv _ _) - Unit := adj₁.Unit ≫ (whiskerLeft F <| whiskerRight adj₂.Unit G) ≫ (Functor.associator _ _ _).inv + unit := adj₁.unit ≫ (whiskerLeft F <| whiskerRight adj₂.unit G) ≫ (Functor.associator _ _ _).inv counit := - (Functor.associator _ _ _).Hom ≫ (whiskerLeft I <| whiskerRight adj₁.counit H) ≫ adj₂.counit + (Functor.associator _ _ _).hom ≫ (whiskerLeft I <| whiskerRight adj₁.counit H) ≫ adj₂.counit #align category_theory.adjunction.comp CategoryTheory.Adjunction.comp /-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/ @@ -440,18 +435,18 @@ variable (e : ∀ X Y, (F_obj X ⟶ Y) ≃ (X ⟶ G.obj Y)) variable (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) private theorem he' {X Y Y'} (f g) : (e X Y').symm (f ≫ G.map g) = (e X Y).symm f ≫ g := by - intros <;> rw [Equiv.symm_apply_eq, he] <;> simp -#align category_theory.adjunction.he' category_theory.adjunction.he' + intros ; rw [Equiv.symm_apply_eq, he] ; simp +-- #align category_theory.adjunction.he' category_theory.adjunction.he' /-- Construct a left adjoint functor to `G`, given the functor's value on objects `F_obj` and a bijection `e` between `F_obj X ⟶ Y` and `X ⟶ G.obj Y` satisfying a naturality law `he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g`. Dual to `right_adjoint_of_equiv`. -/ -@[simps] +@[simps!] def leftAdjointOfEquiv : C ⥤ D where obj := F_obj - map X X' f := (e X (F_obj X')).symm (f ≫ e X' (F_obj X') (𝟙 _)) - map_comp' X X' X'' f f' := + map {X} {X'} f := (e X (F_obj X')).symm (f ≫ e X' (F_obj X') (𝟙 _)) + map_comp := fun f f' => by rw [Equiv.symm_apply_eq, he, Equiv.apply_symm_apply] conv => @@ -462,14 +457,19 @@ def leftAdjointOfEquiv : C ⥤ D where /-- Show that the functor given by `left_adjoint_of_equiv` is indeed left adjoint to `G`. Dual to `adjunction_of_equiv_right`. -/ -@[simps] +@[simps!] def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := by - intros - erw [← he' e he, ← Equiv.apply_eq_iff_eq] - simp [(he _ _ _ _ _).symm] } + homEquiv_naturality_left_symm' := fun {X'} {X} {Y} f g => by + erw [← (he' e he), ← Equiv.apply_eq_iff_eq] + rw [Equiv.apply_symm_apply] + + + -- simp only [(he _ _ _ _ _).symm] + -- simp + -- simp only [id_comp,eq_self_iff_true,Equiv.apply_symm_apply,assoc] + } #align category_theory.adjunction.adjunction_of_equiv_left CategoryTheory.Adjunction.adjunctionOfEquivLeft end ConstructLeft @@ -477,45 +477,43 @@ end ConstructLeft section ConstructRight -- Construction of a right adjoint, analogous to the above. -variable {F} {G_obj : D → C} +variable {G_obj : D → C} variable (e : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G_obj Y)) variable (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) -include he - -private theorem he' {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm (f ≫ g) := by - intros <;> rw [Equiv.eq_symm_apply, he] <;> simp -#align category_theory.adjunction.he' category_theory.adjunction.he' +private theorem he'' {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm (f ≫ g) := by + intros ; rw [Equiv.eq_symm_apply, he] ; simp +-- #align category_theory.adjunction.he' category_theory.adjunction.he' /-- Construct a right adjoint functor to `F`, given the functor's value on objects `G_obj` and a bijection `e` between `F.obj X ⟶ Y` and `X ⟶ G_obj Y` satisfying a naturality law `he : ∀ X Y Y' g h, e X' Y (F.map f ≫ g) = f ≫ e X Y g`. Dual to `left_adjoint_of_equiv`. -/ -@[simps] +@[simps!] def rightAdjointOfEquiv : D ⥤ C where obj := G_obj - map Y Y' g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) - map_comp' Y Y' Y'' g g' := + map {Y} {Y'} g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) + map_comp := fun {Y} {Y'} {Y''} g g' => by - rw [← Equiv.eq_symm_apply, ← he' e he, Equiv.symm_apply_apply] + rw [← Equiv.eq_symm_apply, ← he'' e he, Equiv.symm_apply_apply] conv => rhs - rw [← assoc, he' e he, comp_id, Equiv.symm_apply_apply] + rw [← assoc, he'' e he, comp_id, Equiv.symm_apply_apply] simp #align category_theory.adjunction.right_adjoint_of_equiv CategoryTheory.Adjunction.rightAdjointOfEquiv /-- Show that the functor given by `right_adjoint_of_equiv` is indeed right adjoint to `F`. Dual to `adjunction_of_equiv_left`. -/ -@[simps] -def adjunctionOfEquivRight : F ⊣ rightAdjointOfEquiv e he := +@[simps!] +def adjunctionOfEquivRight : F ⊣ (rightAdjointOfEquiv e he) := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := by intros <;> rw [Equiv.symm_apply_eq, he] <;> simp + homEquiv_naturality_left_symm' := by intro X X' Y f g; rw [Equiv.symm_apply_eq]; dsimp; rw [he]; simp homEquiv_naturality_right' := by intro X Y Y' g h - erw [← he, Equiv.apply_eq_iff_eq, ← assoc, he' e he, comp_id, Equiv.symm_apply_apply] } + erw [← he, Equiv.apply_eq_iff_eq, ← assoc, he'' e he, comp_id, Equiv.symm_apply_apply] } #align category_theory.adjunction.adjunction_of_equiv_right CategoryTheory.Adjunction.adjunctionOfEquivRight end ConstructRight @@ -524,13 +522,13 @@ end ConstructRight If the unit and counit of a given adjunction are (pointwise) isomorphisms, then we can upgrade the adjunction to an equivalence. -/ -@[simps] -noncomputable def toEquivalence (adj : F ⊣ G) [∀ X, IsIso (adj.Unit.app X)] +@[simps!] +noncomputable def toEquivalence (adj : F ⊣ G) [∀ X, IsIso (adj.unit.app X)] [∀ Y, IsIso (adj.counit.app Y)] : C ≌ D where - Functor := F + functor := F inverse := G - unitIso := NatIso.ofComponents (fun X => asIso (adj.Unit.app X)) (by simp) + unitIso := NatIso.ofComponents (fun X => asIso (adj.unit.app X)) (by simp) counitIso := NatIso.ofComponents (fun Y => asIso (adj.counit.app Y)) (by simp) #align category_theory.adjunction.to_equivalence CategoryTheory.Adjunction.toEquivalence @@ -538,9 +536,9 @@ noncomputable def toEquivalence (adj : F ⊣ G) [∀ X, IsIso (adj.Unit.app X)] If the unit and counit for the adjunction corresponding to a right adjoint functor are (pointwise) isomorphisms, then the functor is an equivalence of categories. -/ -@[simps] +@[simps!] noncomputable def isRightAdjointToIsEquivalence [IsRightAdjoint G] - [∀ X, IsIso ((Adjunction.ofRightAdjoint G).Unit.app X)] + [∀ X, IsIso ((Adjunction.ofRightAdjoint G).unit.app X)] [∀ Y, IsIso ((Adjunction.ofRightAdjoint G).counit.app Y)] : IsEquivalence G := IsEquivalence.ofEquivalenceInverse (Adjunction.ofRightAdjoint G).toEquivalence #align category_theory.adjunction.is_right_adjoint_to_is_equivalence CategoryTheory.Adjunction.isRightAdjointToIsEquivalence @@ -553,9 +551,9 @@ namespace Equivalence /-- The adjunction given by an equivalence of categories. (To obtain the opposite adjunction, simply use `e.symm.to_adjunction`. -/ -def toAdjunction (e : C ≌ D) : e.Functor ⊣ e.inverse := +def toAdjunction (e : C ≌ D) : e.functor ⊣ e.inverse := mkOfUnitCounit - ⟨e.Unit, e.counit, by + ⟨e.unit, e.counit, by ext dsimp simp only [id_comp] @@ -568,13 +566,13 @@ def toAdjunction (e : C ≌ D) : e.Functor ⊣ e.inverse := @[simp] theorem asEquivalence_toAdjunction_unit {e : C ≌ D} : - e.Functor.asEquivalence.toAdjunction.Unit = e.Unit := + e.functor.asEquivalence.toAdjunction.unit = e.unit := rfl #align category_theory.equivalence.as_equivalence_to_adjunction_unit CategoryTheory.Equivalence.asEquivalence_toAdjunction_unit @[simp] theorem asEquivalence_toAdjunction_counit {e : C ≌ D} : - e.Functor.asEquivalence.toAdjunction.counit = e.counit := + e.functor.asEquivalence.toAdjunction.counit = e.counit := rfl #align category_theory.equivalence.as_equivalence_to_adjunction_counit CategoryTheory.Equivalence.asEquivalence_toAdjunction_counit From 1f56a41b344cb16296a6c6fc4260aa26fdd0f809 Mon Sep 17 00:00:00 2001 From: mpenciak Date: Mon, 13 Feb 2023 01:14:36 -0500 Subject: [PATCH 29/66] filled in last sorry --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 17 ++++++++--------- 1 file changed, 8 insertions(+), 9 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 0681424718a59..b2d0c575f72b8 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -461,14 +461,13 @@ to `adjunction_of_equiv_right`. -/ def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := fun {X'} {X} {Y} f g => by - erw [← (he' e he), ← Equiv.apply_eq_iff_eq] - rw [Equiv.apply_symm_apply] - - - -- simp only [(he _ _ _ _ _).symm] - -- simp - -- simp only [id_comp,eq_self_iff_true,Equiv.apply_symm_apply,assoc] + homEquiv_naturality_left_symm' := fun {X'} {X} {Y} f g => by + have := @he' C _ D _ G F_obj e he + erw [← this, ← Equiv.apply_eq_iff_eq (e X' Y)] + simp [(he X' (F_obj X) Y (e X Y |>.symm g) (leftAdjointOfEquiv e he |>.map f)).symm] + congr + rw [← he] + simp } #align category_theory.adjunction.adjunction_of_equiv_left CategoryTheory.Adjunction.adjunctionOfEquivLeft @@ -495,7 +494,7 @@ Dual to `left_adjoint_of_equiv`. -/ def rightAdjointOfEquiv : D ⥤ C where obj := G_obj map {Y} {Y'} g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) - map_comp := fun {Y} {Y'} {Y''} g g' => + map_comp := fun {Y} {Y'} {Y''} g g' => by rw [← Equiv.eq_symm_apply, ← he'' e he, Equiv.symm_apply_apply] conv => From 0914b9d0a55cc46a3204805ae989da5192f9bbbb Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Mon, 13 Feb 2023 11:31:31 -0500 Subject: [PATCH 30/66] break long lines --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 14 +++++++++----- 1 file changed, 9 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index b2d0c575f72b8..63a2243033f1c 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -133,7 +133,8 @@ theorem homEquiv_naturality_left_symm (f : X' ⟶ X) (g : X ⟶ G.obj Y) : @[simp] theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by - rw [← Equiv.eq_symm_apply] ; simp only [Equiv.symm_apply_apply,eq_self_iff_true,homEquiv_naturality_left_symm] + rw [← Equiv.eq_symm_apply] + simp only [Equiv.symm_apply_apply,eq_self_iff_true,homEquiv_naturality_left_symm] #align category_theory.adjunction.hom_equiv_naturality_left CategoryTheory.Adjunction.homEquiv_naturality_left @[simp] @@ -145,7 +146,8 @@ theorem homEquiv_naturality_right (f : F.obj X ⟶ Y) (g : Y ⟶ Y') : @[simp] theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by - rw [Equiv.symm_apply_eq] ; simp only [homEquiv_naturality_right,eq_self_iff_true,Equiv.apply_symm_apply] + rw [Equiv.symm_apply_eq] + simp only [homEquiv_naturality_right,eq_self_iff_true,Equiv.apply_symm_apply] #align category_theory.adjunction.hom_equiv_naturality_right_symm CategoryTheory.Adjunction.homEquiv_naturality_right_symm @[simp] @@ -494,8 +496,7 @@ Dual to `left_adjoint_of_equiv`. -/ def rightAdjointOfEquiv : D ⥤ C where obj := G_obj map {Y} {Y'} g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) - map_comp := fun {Y} {Y'} {Y''} g g' => - by + map_comp := fun {Y} {Y'} {Y''} g g' => by rw [← Equiv.eq_symm_apply, ← he'' e he, Equiv.symm_apply_apply] conv => rhs @@ -509,7 +510,8 @@ to `adjunction_of_equiv_left`. -/ def adjunctionOfEquivRight : F ⊣ (rightAdjointOfEquiv e he) := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := by intro X X' Y f g; rw [Equiv.symm_apply_eq]; dsimp; rw [he]; simp + homEquiv_naturality_left_symm' := by + intro X X' Y f g; rw [Equiv.symm_apply_eq]; dsimp; rw [he]; simp homEquiv_naturality_right' := by intro X Y Y' g h erw [← he, Equiv.apply_eq_iff_eq, ← assoc, he'' e he, comp_id, Equiv.symm_apply_apply] } @@ -611,3 +613,5 @@ theorem leftAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : leftAdjoi end Functor end CategoryTheory + +#lint From 671edf63e9b9de977d878dcc2bd4481dd330aba4 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Mon, 13 Feb 2023 11:32:00 -0500 Subject: [PATCH 31/66] delete linter command --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 63a2243033f1c..3f77c028fe46d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -613,5 +613,3 @@ theorem leftAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : leftAdjoi end Functor end CategoryTheory - -#lint From 0e48e0e5e009a59784fe9f418ec5d0bb502c6c93 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Mon, 13 Feb 2023 12:30:08 -0500 Subject: [PATCH 32/66] fix comments --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 44 ++++++++++---------- 1 file changed, 22 insertions(+), 22 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 3f77c028fe46d..3133c45bc6664 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -17,23 +17,23 @@ import Mathlib.CategoryTheory.Equivalence `F : C ⥤ D` and `G : D ⥤ C`. `F` is the left adjoint and `G` is the right adjoint. We provide various useful constructors: -* `mk_of_hom_equiv` -* `mk_of_unit_counit` -* `left_adjoint_of_equiv` / `right_adjoint_of equiv` +* `mkOfHomEquiv` +* `mkOfUnitCounit` +* `leftAdjointOfEquiv` / `rightAdjointOfEquiv` construct a left/right adjoint of a given functor given the action on objects and the relevant equivalence of morphism spaces. -* `adjunction_of_equiv_left` / `adjunction_of_equiv_right` witness that these constructions +* `adjunctionOfEquivLeft` / `adjunctionOfEquivRight` witness that these constructions give adjunctions. -There are also typeclasses `is_left_adjoint` / `is_right_adjoint`, carrying data witnessing +There are also typeclasses `IsLeftAdjoint` / `IsRightAdjoint`, carrying data witnessing that a given functor is a left or right adjoint. -Given `[is_left_adjoint F]`, a right adjoint of `F` can be constructed as `right_adjoint F`. +Given `[IsLeftAdjoint F]`, a right adjoint of `F` can be constructed as `rightAdjoint F`. -`adjunction.comp` composes adjunctions. +`Adjunction.comp` composes adjunctions. -`to_equivalence` upgrades an adjunction to an equivalence, +`toEquivalence` upgrades an adjunction to an equivalence, given witnesses that the unit and counit are pointwise isomorphisms. -Conversely `equivalence.to_adjunction` recovers the underlying adjunction from an equivalence. +Conversely `Equivalence.toAdjunction` recovers the underlying adjunction from an equivalence. -/ @@ -41,7 +41,7 @@ namespace CategoryTheory open Category --- declare the `v`'s first; see `category_theory.category` for an explanation +-- declare the `v`'s first; see `CategoryTheory.Category` for an explanation universe v₁ v₂ v₃ u₁ u₂ u₃ -- Porting Note: `elab_without_expected_type` cannot be a local attribute @@ -53,11 +53,11 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] `F : C ⥤ D` and `G : D ⥤ C`. `F` is the left adjoint and `G` is the right adjoint. To construct an `adjunction` between two functors, it's often easier to instead use the -constructors `mk_of_hom_equiv` or `mk_of_unit_counit`. To construct a left adjoint, -there are also constructors `left_adjoint_of_equiv` and `adjunction_of_equiv_left` (as +constructors `mkOfHomEquiv` or `mkOfUnitCounit`. To construct a left adjoint, +there are also constructors `leftAdjointOfEquiv` and `adjunctionOfEquivLeft` (as well as their duals) which can be simpler in practice. -Uniqueness of adjoints is shown in `category_theory.adjunction.opposites`. +Uniqueness of adjoints is shown in `CategoryTheory.Adjunction.Opposites`. See . -/ @@ -213,7 +213,7 @@ end Adjunction namespace Adjunction /-- This is an auxiliary data structure useful for constructing adjunctions. -See `adjunction.mk_of_hom_equiv`. +See `Adjunction.mkOfHomEquiv`. This structure won't typically be used anywhere else. -/ -- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) @@ -255,7 +255,7 @@ theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : end CoreHomEquiv /-- This is an auxiliary data structure useful for constructing adjunctions. -See `adjunction.mk_of_unit_counit`. +See `Adjunction.mkOfUnitCounit`. This structure won't typically be used anywhere else. -/ -- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) @@ -443,7 +443,7 @@ private theorem he' {X Y Y'} (f g) : (e X Y').symm (f ≫ G.map g) = (e X Y).sym /-- Construct a left adjoint functor to `G`, given the functor's value on objects `F_obj` and a bijection `e` between `F_obj X ⟶ Y` and `X ⟶ G.obj Y` satisfying a naturality law `he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g`. -Dual to `right_adjoint_of_equiv`. -/ +Dual to `rightAdjointOfEquiv`. -/ @[simps!] def leftAdjointOfEquiv : C ⥤ D where obj := F_obj @@ -457,8 +457,8 @@ def leftAdjointOfEquiv : C ⥤ D where simp #align category_theory.adjunction.left_adjoint_of_equiv CategoryTheory.Adjunction.leftAdjointOfEquiv -/-- Show that the functor given by `left_adjoint_of_equiv` is indeed left adjoint to `G`. Dual -to `adjunction_of_equiv_right`. -/ +/-- Show that the functor given by `leftAdjointOfEquiv` is indeed left adjoint to `G`. Dual +to `adjunctionOfRightEquiv`. -/ @[simps!] def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := mkOfHomEquiv @@ -491,7 +491,7 @@ private theorem he'' {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm /-- Construct a right adjoint functor to `F`, given the functor's value on objects `G_obj` and a bijection `e` between `F.obj X ⟶ Y` and `X ⟶ G_obj Y` satisfying a naturality law `he : ∀ X Y Y' g h, e X' Y (F.map f ≫ g) = f ≫ e X Y g`. -Dual to `left_adjoint_of_equiv`. -/ +Dual to `leftAdjointOfEquiv`. -/ @[simps!] def rightAdjointOfEquiv : D ⥤ C where obj := G_obj @@ -504,8 +504,8 @@ def rightAdjointOfEquiv : D ⥤ C where simp #align category_theory.adjunction.right_adjoint_of_equiv CategoryTheory.Adjunction.rightAdjointOfEquiv -/-- Show that the functor given by `right_adjoint_of_equiv` is indeed right adjoint to `F`. Dual -to `adjunction_of_equiv_left`. -/ +/-- Show that the functor given by `rightAdjointOfEquiv` is indeed right adjoint to `F`. Dual +to `adjunctionOfEquivRight`. -/ @[simps!] def adjunctionOfEquivRight : F ⊣ (rightAdjointOfEquiv e he) := mkOfHomEquiv @@ -551,7 +551,7 @@ open Adjunction namespace Equivalence /-- The adjunction given by an equivalence of categories. (To obtain the opposite adjunction, -simply use `e.symm.to_adjunction`. -/ +simply use `e.symm.toAdjunction`. -/ def toAdjunction (e : C ≌ D) : e.functor ⊣ e.inverse := mkOfUnitCounit ⟨e.unit, e.counit, by From 4d3bf86b12c08c68db4784ed37f47b96aceef3ee Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 14 Feb 2023 07:42:47 +0100 Subject: [PATCH 33/66] fix two simpNF lints --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 3133c45bc6664..f85a5a8179f6e 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -241,12 +241,22 @@ attribute [simp] homEquiv_naturality_left_symm homEquiv_naturality_right variable {F : C ⥤ D} {G : D ⥤ C} (adj : CoreHomEquiv F G) {X' X : C} {Y Y' : D} @[simp] +theorem homEquiv_naturality_left_aux (f : X' ⟶ X) (g : F.obj X ⟶ Y) : + (adj.homEquiv X' (F.obj X)) (F.map f) ≫ G.map g = f ≫ (adj.homEquiv X Y) g := by + rw [← homEquiv_naturality_right, ← Equiv.eq_symm_apply] ; simp + +-- @[simp] -- Porting note: LHS simplifies, added aux lemma above theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by rw [← Equiv.eq_symm_apply] ; simp #align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_left CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_left @[simp] +theorem homEquiv_naturality_right_symm_aux (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : + F.map f ≫ (adj.homEquiv (G.obj Y) Y').symm (G.map g) = (adj.homEquiv X Y).symm f ≫ g := by + rw [← homEquiv_naturality_left_symm, Equiv.symm_apply_eq] ; simp + +-- @[simp] -- Porting note: LHS simplifies, added aux lemma above theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by rw [Equiv.symm_apply_eq] ; simp From 5b11f3d9912b56e6f758fc83ca9cfa78f48770f3 Mon Sep 17 00:00:00 2001 From: mpenciak Date: Wed, 15 Feb 2023 14:29:42 -0500 Subject: [PATCH 34/66] fill in some docstrings --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index f85a5a8179f6e..d0aae081d84aa 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -62,25 +62,33 @@ Uniqueness of adjoints is shown in `CategoryTheory.Adjunction.Opposites`. See . -/ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where + /-- The equivalence between `Hom (F X) Y` and `Hom X (G Y)` coming from an adjunction -/ homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) + /-- The unit of an adjunction -/ unit : 𝟭 C ⟶ F.comp G + /-- The counit of an adjunction -/ counit : G.comp F ⟶ 𝟭 D homEquiv_unit' : ∀ {X Y f}, (homEquiv X Y) f = (unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat homEquiv_counit' : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat #align category_theory.adjunction CategoryTheory.Adjunction -- mathport name: «expr ⊣ » +/-- The notation `F ⊣ G` stands for `Adjunction F G` representing that `F` is left adjoint to `G` -/ infixl:15 " ⊣ " => Adjunction /-- A class giving a chosen right adjoint to the functor `left`. -/ class IsLeftAdjoint (left : C ⥤ D) where + /-- The right adjoint to `left` -/ right : D ⥤ C + /-- The adjunction between `left` and `right` -/ adj : left ⊣ right #align category_theory.is_left_adjoint CategoryTheory.IsLeftAdjoint /-- A class giving a chosen left adjoint to the functor `right`. -/ class IsRightAdjoint (right : D ⥤ C) where + /-- The left adjoint to `right` -/ left : C ⥤ D + /-- The adjunction between `left` and `right` -/ adj : left ⊣ right #align category_theory.is_right_adjoint CategoryTheory.IsRightAdjoint @@ -219,6 +227,7 @@ This structure won't typically be used anywhere else. -- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) -- @[nolint has_nonempty_instance] structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where + /-- The equivalence between `Hom (F X) Y` and `Hom X (G Y)` -/ homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) homEquiv_naturality_left_symm' : ∀ {X' X Y} (f : X' ⟶ X) (g : X ⟶ G.obj Y), @@ -271,7 +280,9 @@ This structure won't typically be used anywhere else. -- Porting comment: `has_nonempty_instance` linter doesn't exist (yet?) -- @[nolint has_nonempty_instance] structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where + /-- The unit of an adjunction between `F` and `G` -/ unit : 𝟭 C ⟶ F.comp G + /-- The counit of an adjunction between `F` and `G`s -/ counit : G.comp F ⟶ 𝟭 D left_triangle : whiskerRight unit F ≫ (Functor.associator F G F).hom ≫ whiskerLeft F counit = From ccb1de541ebe73e857f02ff9b674fe4509be6ffd Mon Sep 17 00:00:00 2001 From: mpenciak Date: Wed, 15 Feb 2023 17:26:37 -0500 Subject: [PATCH 35/66] fix most linter issues --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 56 +++++++++++++------- 1 file changed, 38 insertions(+), 18 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index d0aae081d84aa..d5371a22d05bf 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -68,8 +68,12 @@ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where unit : 𝟭 C ⟶ F.comp G /-- The counit of an adjunction -/ counit : G.comp F ⟶ 𝟭 D - homEquiv_unit' : ∀ {X Y f}, (homEquiv X Y) f = (unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat - homEquiv_counit' : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat + -- Porting note: It's strange that this `Prop` is being flagged by the `docBlame` linter + /-- Naturality of the unit of an adjunction -/ + homEquiv_unit : ∀ {X Y f}, (homEquiv X Y) f = (unit : _ ⟶ _).app X ≫ G.map f := by aesop_cat + -- Porting note: It's strange that this `Prop` is being flagged by the `docBlame` linter + /-- Naturality of the counit of an adjunction -/ + homEquiv_counit : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat #align category_theory.adjunction CategoryTheory.Adjunction -- mathport name: «expr ⊣ » @@ -116,9 +120,10 @@ def Adjunction.ofRightAdjoint (right : C ⥤ D) [IsRightAdjoint right] : namespace Adjunction -restate_axiom homEquiv_unit' +-- porting note: Workaround not needed in Lean 4 +-- restate_axiom homEquiv_unit' -restate_axiom homEquiv_counit' +-- restate_axiom homEquiv_counit' attribute [simp] homEquiv_unit homEquiv_counit @@ -132,26 +137,34 @@ theorem homEquiv_id (X : C) : adj.homEquiv X _ (𝟙 _) = adj.unit.app X := by s theorem homEquiv_symm_id (X : D) : (adj.homEquiv _ X).symm (𝟙 _) = adj.counit.app X := by simp #align category_theory.adjunction.hom_equiv_symm_id CategoryTheory.Adjunction.homEquiv_symm_id -@[simp] +/- +Porting note: `nolint simpNF` as the linter was complaining that this was provable using `simp` +but it is in fact not. Also the `docBlame` linter expects a docstring even though this is `Prop` +valued +-/ +@[simp, nolint simpNF] theorem homEquiv_naturality_left_symm (f : X' ⟶ X) (g : X ⟶ G.obj Y) : (adj.homEquiv X' Y).symm (f ≫ g) = F.map f ≫ (adj.homEquiv X Y).symm g := by rw [homEquiv_counit, F.map_comp, assoc, adj.homEquiv_counit.symm] #align category_theory.adjunction.hom_equiv_naturality_left_symm CategoryTheory.Adjunction.homEquiv_naturality_left_symm -@[simp] +-- Porting note: Same as above +@[simp, nolint simpNF] theorem homEquiv_naturality_left (f : X' ⟶ X) (g : F.obj X ⟶ Y) : (adj.homEquiv X' Y) (F.map f ≫ g) = f ≫ (adj.homEquiv X Y) g := by rw [← Equiv.eq_symm_apply] simp only [Equiv.symm_apply_apply,eq_self_iff_true,homEquiv_naturality_left_symm] #align category_theory.adjunction.hom_equiv_naturality_left CategoryTheory.Adjunction.homEquiv_naturality_left -@[simp] +-- Porting note: Same as above +@[simp, nolint simpNF] theorem homEquiv_naturality_right (f : F.obj X ⟶ Y) (g : Y ⟶ Y') : (adj.homEquiv X Y') (f ≫ g) = (adj.homEquiv X Y) f ≫ G.map g := by rw [homEquiv_unit, G.map_comp, ← assoc, ← homEquiv_unit] #align category_theory.adjunction.hom_equiv_naturality_right CategoryTheory.Adjunction.homEquiv_naturality_right -@[simp] +-- Porting note: Same as above +@[simp, nolint simpNF] theorem homEquiv_naturality_right_symm (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : (adj.homEquiv X Y').symm (f ≫ G.map g) = (adj.homEquiv X Y).symm f ≫ g := by rw [Equiv.symm_apply_eq] @@ -229,11 +242,13 @@ This structure won't typically be used anywhere else. structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where /-- The equivalence between `Hom (F X) Y` and `Hom X (G Y)` -/ homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) - homEquiv_naturality_left_symm' : + /-- The property that describes how `homEquiv.symm` transforms compositions `X' ⟶ X ⟶ G Y` -/ + homEquiv_naturality_left_symm : ∀ {X' X Y} (f : X' ⟶ X) (g : X ⟶ G.obj Y), (homEquiv X' Y).symm (f ≫ g) = F.map f ≫ (homEquiv X Y).symm g := by aesop_cat - homEquiv_naturality_right' : + /-- The property that describes how `homEquiv` transforms compositions `F X ⟶ Y ⟶ Y'` -/ + homEquiv_naturality_right : ∀ {X Y Y'} (f : F.obj X ⟶ Y) (g : Y ⟶ Y'), (homEquiv X Y') (f ≫ g) = (homEquiv X Y) f ≫ G.map g := by aesop_cat @@ -241,9 +256,10 @@ structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where namespace CoreHomEquiv -restate_axiom homEquiv_naturality_left_symm' +-- Porting note: Workaround not needed in Lean 4. +-- restate_axiom homEquiv_naturality_left_symm' -restate_axiom homEquiv_naturality_right' +-- restate_axiom homEquiv_naturality_right' attribute [simp] homEquiv_naturality_left_symm homEquiv_naturality_right @@ -284,10 +300,14 @@ structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where unit : 𝟭 C ⟶ F.comp G /-- The counit of an adjunction between `F` and `G`s -/ counit : G.comp F ⟶ 𝟭 D + /-- Equality of the composition of the unit, associator, and counit with the identity + `F ⟶ (F G) F ⟶ F (G F) ⟶ F = NatTrans.id F` -/ left_triangle : whiskerRight unit F ≫ (Functor.associator F G F).hom ≫ whiskerLeft F counit = NatTrans.id (𝟭 C ⋙ F) := by aesop_cat + /-- Equality of the composition of the unit, associator, and counit with the identity + `G ⟶ G (F G) ⟶ (F G) F ⟶ G = NatTrans.id G` -/ right_triangle : whiskerLeft G unit ≫ (Functor.associator G F G).inv ≫ whiskerRight counit G = NatTrans.id (G ⋙ 𝟭 C) := by @@ -320,14 +340,14 @@ def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := intros erw [← adj.homEquiv_naturality_left_symm, ← adj.homEquiv_naturality_right_symm] dsimp; simp } - homEquiv_unit' := @fun X Y f => by erw [← adj.homEquiv_naturality_right]; simp - homEquiv_counit' := @fun X Y f => by erw [← adj.homEquiv_naturality_left_symm]; simp + homEquiv_unit := @fun X Y f => by erw [← adj.homEquiv_naturality_right]; simp + homEquiv_counit := @fun X Y f => by erw [← adj.homEquiv_naturality_left_symm]; simp } #align category_theory.adjunction.mk_of_hom_equiv CategoryTheory.Adjunction.mkOfHomEquiv /-- Construct an adjunction between functors `F` and `G` given a unit and counit for the adjunction satisfying the triangle identities. -/ -@[simps!] +@[simps!, nolint simpNF] def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := { adj with homEquiv := fun X Y => @@ -484,7 +504,7 @@ to `adjunctionOfRightEquiv`. -/ def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := fun {X'} {X} {Y} f g => by + homEquiv_naturality_left_symm := fun {X'} {X} {Y} f g => by have := @he' C _ D _ G F_obj e he erw [← this, ← Equiv.apply_eq_iff_eq (e X' Y)] simp [(he X' (F_obj X) Y (e X Y |>.symm g) (leftAdjointOfEquiv e he |>.map f)).symm] @@ -531,9 +551,9 @@ to `adjunctionOfEquivRight`. -/ def adjunctionOfEquivRight : F ⊣ (rightAdjointOfEquiv e he) := mkOfHomEquiv { homEquiv := e - homEquiv_naturality_left_symm' := by + homEquiv_naturality_left_symm := by intro X X' Y f g; rw [Equiv.symm_apply_eq]; dsimp; rw [he]; simp - homEquiv_naturality_right' := by + homEquiv_naturality_right := by intro X Y Y' g h erw [← he, Equiv.apply_eq_iff_eq, ← assoc, he'' e he, comp_id, Equiv.symm_apply_apply] } #align category_theory.adjunction.adjunction_of_equiv_right CategoryTheory.Adjunction.adjunctionOfEquivRight From 3603497c98fe27899475c05bedd76d299af605b4 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 14:25:11 -0500 Subject: [PATCH 36/66] add missing aligns for fields; lint --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 23 +++++++++++++++++++- 1 file changed, 22 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index d5371a22d05bf..a5b65f05aaa50 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -75,6 +75,11 @@ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where /-- Naturality of the counit of an adjunction -/ homEquiv_counit : ∀ {X Y g}, (homEquiv X Y).symm g = F.map g ≫ counit.app Y := by aesop_cat #align category_theory.adjunction CategoryTheory.Adjunction +#align category_theory.adjunction.hom_equiv CategoryTheory.Adjunction.homEquiv +#align category_theory.adjunction.hom_equiv_unit CategoryTheory.Adjunction.homEquiv_unit +#align category_theory.adjunction.hom_equiv_unit' CategoryTheory.Adjunction.homEquiv_unit +#align category_theory.adjunction.hom_equiv_counit CategoryTheory.Adjunction.homEquiv_counit +#align category_theory.adjunction.hom_equiv_counit' CategoryTheory.Adjunction.homEquiv_counit -- mathport name: «expr ⊣ » /-- The notation `F ⊣ G` stands for `Adjunction F G` representing that `F` is left adjoint to `G` -/ @@ -253,6 +258,12 @@ structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where (homEquiv X Y') (f ≫ g) = (homEquiv X Y) f ≫ G.map g := by aesop_cat #align category_theory.adjunction.core_hom_equiv CategoryTheory.Adjunction.CoreHomEquiv +#align category_theory.adjunction.core_hom_equiv.hom_equiv CategoryTheory.Adjunction.CoreHomEquiv.homEquiv +#align category_theory.adjunction.core_hom_equiv.hom_equiv' CategoryTheory.Adjunction.CoreHomEquiv.homEquiv +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_right CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_right +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_right' CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_right +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_left_symm CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_left_symm +#align category_theory.adjunction.core_hom_equiv.hom_equiv_naturality_left_symm' CategoryTheory.Adjunction.CoreHomEquiv.homEquiv_naturality_left_symm namespace CoreHomEquiv @@ -313,6 +324,10 @@ structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where NatTrans.id (G ⋙ 𝟭 C) := by aesop_cat #align category_theory.adjunction.core_unit_counit CategoryTheory.Adjunction.CoreUnitCounit +#align category_theory.adjunction.core_unit_counit.left_triangle' CategoryTheory.Adjunction.CoreUnitCounit.left_triangle +#align category_theory.adjunction.core_unit_counit.left_triangle CategoryTheory.Adjunction.CoreUnitCounit.left_triangle +#align category_theory.adjunction.core_unit_counit.right_triangle' CategoryTheory.Adjunction.CoreUnitCounit.right_triangle +#align category_theory.adjunction.core_unit_counit.right_triangle CategoryTheory.Adjunction.CoreUnitCounit.right_triangle namespace CoreUnitCounit @@ -347,7 +362,8 @@ def mkOfHomEquiv (adj : CoreHomEquiv F G) : F ⊣ G := /-- Construct an adjunction between functors `F` and `G` given a unit and counit for the adjunction satisfying the triangle identities. -/ -@[simps!, nolint simpNF] + +@[simps!] def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := { adj with homEquiv := fun X Y => @@ -371,6 +387,11 @@ def mkOfUnitCounit (adj : CoreUnitCounit F G) : F ⊣ G := exact t } } #align category_theory.adjunction.mk_of_unit_counit CategoryTheory.Adjunction.mkOfUnitCounit +/- Porting note: simpNF linter claims these are solved by simp but that +is not true -/ +attribute [nolint simpNF] CategoryTheory.Adjunction.mkOfUnitCounit_homEquiv_symm_apply +attribute [nolint simpNF] CategoryTheory.Adjunction.mkOfUnitCounit_homEquiv_apply + /-- The adjunction between the identity functor on a category and itself. -/ def id : 𝟭 C ⊣ 𝟭 C where homEquiv X Y := Equiv.refl _ From 0e077343138b6fee647400c24e7070429f8f3577 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 14:38:55 -0500 Subject: [PATCH 37/66] restore lost import line --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 32883cab7d4bb..ff35b174ad072 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -218,6 +218,7 @@ import Mathlib.Algebra.Support import Mathlib.Algebra.Tropical.Basic import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice +import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Arrow import Mathlib.CategoryTheory.Balanced import Mathlib.CategoryTheory.Bicategory.Basic From 5fedaef75b579d43b6c673749ca5a6ad605172e8 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 20:42:37 +0100 Subject: [PATCH 38/66] feat: port CategoryTheory.CommSq From b108a9db4790757011118df697a33081815cfb64 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 20:42:37 +0100 Subject: [PATCH 39/66] Initial file copy from mathport --- Mathlib/CategoryTheory/CommSq.lean | 228 +++++++++++++++++++++++++++++ 1 file changed, 228 insertions(+) create mode 100644 Mathlib/CategoryTheory/CommSq.lean diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean new file mode 100644 index 0000000000000..1d1332b2c2b4c --- /dev/null +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Joël Riou + +! This file was ported from Lean 3 source module category_theory.comm_sq +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Arrow + +/-! +# Commutative squares + +This file provide an API for commutative squares in categories. +If `top`, `left`, `right` and `bottom` are four morphisms which are the edges +of a square, `comm_sq top left right bottom` is the predicate that this +square is commutative. + +The structure `comm_sq` is extended in `category_theory/shapes/limits/comm_sq.lean` +as `is_pullback` and `is_pushout` in order to define pullback and pushout squares. + +## Future work + +Refactor `lift_struct` from `arrow.lean` and lifting properties using `comm_sq.lean`. + +-/ + + +namespace CategoryTheory + +variable {C : Type _} [Category C] + +/-- The proposition that a square +``` + W ---f---> X + | | + g h + | | + v v + Y ---i---> Z + +``` +is a commuting square. +-/ +structure CommSq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y ⟶ Z) : Prop where + w : f ≫ h = g ≫ i +#align category_theory.comm_sq CategoryTheory.CommSq + +attribute [reassoc.1] comm_sq.w + +namespace CommSq + +variable {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + +theorem flip (p : CommSq f g h i) : CommSq g f i h := + ⟨p.w.symm⟩ +#align category_theory.comm_sq.flip CategoryTheory.CommSq.flip + +theorem of_arrow {f g : Arrow C} (h : f ⟶ g) : CommSq f.Hom h.left h.right g.Hom := + ⟨h.w.symm⟩ +#align category_theory.comm_sq.of_arrow CategoryTheory.CommSq.of_arrow + +/-- The commutative square in the opposite category associated to a commutative square. -/ +theorem op (p : CommSq f g h i) : CommSq i.op h.op g.op f.op := + ⟨by simp only [← op_comp, p.w]⟩ +#align category_theory.comm_sq.op CategoryTheory.CommSq.op + +/-- The commutative square associated to a commutative square in the opposite category. -/ +theorem unop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} (p : CommSq f g h i) : + CommSq i.unop h.unop g.unop f.unop := + ⟨by simp only [← unop_comp, p.w]⟩ +#align category_theory.comm_sq.unop CategoryTheory.CommSq.unop + +end CommSq + +namespace Functor + +variable {D : Type _} [Category D] + +variable (F : C ⥤ D) {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + +theorem map_commSq (s : CommSq f g h i) : CommSq (F.map f) (F.map g) (F.map h) (F.map i) := + ⟨by simpa using congr_arg (fun k : W ⟶ Z => F.map k) s.w⟩ +#align category_theory.functor.map_comm_sq CategoryTheory.Functor.map_commSq + +end Functor + +alias functor.map_comm_sq ← comm_sq.map +#align category_theory.comm_sq.map CategoryTheory.CommSq.map + +namespace CommSq + +variable {A B X Y : C} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} + +/-- The datum of a lift in a commutative square, i.e. a up-right-diagonal +morphism which makes both triangles commute. -/ +@[ext, nolint has_nonempty_instance] +structure LiftStruct (sq : CommSq f i p g) where + l : B ⟶ X + fac_left' : i ≫ l = f + fac_right' : l ≫ p = g +#align category_theory.comm_sq.lift_struct CategoryTheory.CommSq.LiftStruct + +namespace LiftStruct + +restate_axiom fac_left' + +restate_axiom fac_right' + +/-- A `lift_struct` for a commutative square gives a `lift_struct` for the +corresponding square in the opposite category. -/ +@[simps] +def op {sq : CommSq f i p g} (l : LiftStruct sq) : LiftStruct sq.op + where + l := l.l.op + fac_left' := by rw [← op_comp, l.fac_right] + fac_right' := by rw [← op_comp, l.fac_left] +#align category_theory.comm_sq.lift_struct.op CategoryTheory.CommSq.LiftStruct.op + +/-- A `lift_struct` for a commutative square in the opposite category +gives a `lift_struct` for the corresponding square in the original category. -/ +@[simps] +def unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} {sq : CommSq f i p g} + (l : LiftStruct sq) : LiftStruct sq.unop + where + l := l.l.unop + fac_left' := by rw [← unop_comp, l.fac_right] + fac_right' := by rw [← unop_comp, l.fac_left] +#align category_theory.comm_sq.lift_struct.unop CategoryTheory.CommSq.LiftStruct.unop + +/-- Equivalences of `lift_struct` for a square and the corresponding square +in the opposite category. -/ +@[simps] +def opEquiv (sq : CommSq f i p g) : LiftStruct sq ≃ LiftStruct sq.op + where + toFun := op + invFun := unop + left_inv := by tidy + right_inv := by tidy +#align category_theory.comm_sq.lift_struct.op_equiv CategoryTheory.CommSq.LiftStruct.opEquiv + +/-- Equivalences of `lift_struct` for a square in the oppositive category and +the corresponding square in the original category. -/ +def unopEquiv {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} + (sq : CommSq f i p g) : LiftStruct sq ≃ LiftStruct sq.unop + where + toFun := unop + invFun := op + left_inv := by tidy + right_inv := by tidy +#align category_theory.comm_sq.lift_struct.unop_equiv CategoryTheory.CommSq.LiftStruct.unopEquiv + +end LiftStruct + +instance subsingleton_liftStruct_of_epi (sq : CommSq f i p g) [Epi i] : + Subsingleton (LiftStruct sq) := + ⟨fun l₁ l₂ => by + ext + simp only [← cancel_epi i, lift_struct.fac_left]⟩ +#align category_theory.comm_sq.subsingleton_lift_struct_of_epi CategoryTheory.CommSq.subsingleton_liftStruct_of_epi + +instance subsingleton_liftStruct_of_mono (sq : CommSq f i p g) [Mono p] : + Subsingleton (LiftStruct sq) := + ⟨fun l₁ l₂ => by + ext + simp only [← cancel_mono p, lift_struct.fac_right]⟩ +#align category_theory.comm_sq.subsingleton_lift_struct_of_mono CategoryTheory.CommSq.subsingleton_liftStruct_of_mono + +variable (sq : CommSq f i p g) + +/-- The assertion that a square has a `lift_struct`. -/ +class HasLift : Prop where + exists_lift : Nonempty sq.LiftStruct +#align category_theory.comm_sq.has_lift CategoryTheory.CommSq.HasLift + +namespace HasLift + +variable {sq} + +theorem mk' (l : sq.LiftStruct) : HasLift sq := + ⟨Nonempty.intro l⟩ +#align category_theory.comm_sq.has_lift.mk' CategoryTheory.CommSq.HasLift.mk' + +variable (sq) + +theorem iff : HasLift sq ↔ Nonempty sq.LiftStruct := + by + constructor + exacts[fun h => h.exists_lift, fun h => mk h] +#align category_theory.comm_sq.has_lift.iff CategoryTheory.CommSq.HasLift.iff + +theorem iff_op : HasLift sq ↔ HasLift sq.op := + by + rw [Iff, Iff] + exact Nonempty.congr (lift_struct.op_equiv sq).toFun (lift_struct.op_equiv sq).invFun +#align category_theory.comm_sq.has_lift.iff_op CategoryTheory.CommSq.HasLift.iff_op + +theorem iff_unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} + (sq : CommSq f i p g) : HasLift sq ↔ HasLift sq.unop := + by + rw [Iff, Iff] + exact Nonempty.congr (lift_struct.unop_equiv sq).toFun (lift_struct.unop_equiv sq).invFun +#align category_theory.comm_sq.has_lift.iff_unop CategoryTheory.CommSq.HasLift.iff_unop + +end HasLift + +/-- A choice of a diagonal morphism that is part of a `lift_struct` when +the square has a lift. -/ +noncomputable def lift [hsq : HasLift sq] : B ⟶ X := + hsq.exists_lift.some.l +#align category_theory.comm_sq.lift CategoryTheory.CommSq.lift + +@[simp, reassoc.1] +theorem fac_left [hsq : HasLift sq] : i ≫ sq.lift = f := + hsq.exists_lift.some.fac_left +#align category_theory.comm_sq.fac_left CategoryTheory.CommSq.fac_left + +@[simp, reassoc.1] +theorem fac_right [hsq : HasLift sq] : sq.lift ≫ p = g := + hsq.exists_lift.some.fac_right +#align category_theory.comm_sq.fac_right CategoryTheory.CommSq.fac_right + +end CommSq + +end CategoryTheory + From 8e67aa7ef360a9dfb3d0b148287069f625b5eff0 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 20:42:37 +0100 Subject: [PATCH 40/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/CommSq.lean | 11 ++++------- 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 32883cab7d4bb..65f546eeccaff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -226,6 +226,7 @@ import Mathlib.CategoryTheory.Category.Basic import Mathlib.CategoryTheory.Category.KleisliCat import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.Category.RelCat +import Mathlib.CategoryTheory.CommSq import Mathlib.CategoryTheory.Comma import Mathlib.CategoryTheory.ConcreteCategory.Bundled import Mathlib.CategoryTheory.Endomorphism diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index 1d1332b2c2b4c..b0e9f9499d38c 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -8,7 +8,7 @@ Authors: Scott Morrison, Joël Riou ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Arrow +import Mathlib.CategoryTheory.Arrow /-! # Commutative squares @@ -185,21 +185,18 @@ theorem mk' (l : sq.LiftStruct) : HasLift sq := variable (sq) -theorem iff : HasLift sq ↔ Nonempty sq.LiftStruct := - by +theorem iff : HasLift sq ↔ Nonempty sq.LiftStruct := by constructor exacts[fun h => h.exists_lift, fun h => mk h] #align category_theory.comm_sq.has_lift.iff CategoryTheory.CommSq.HasLift.iff -theorem iff_op : HasLift sq ↔ HasLift sq.op := - by +theorem iff_op : HasLift sq ↔ HasLift sq.op := by rw [Iff, Iff] exact Nonempty.congr (lift_struct.op_equiv sq).toFun (lift_struct.op_equiv sq).invFun #align category_theory.comm_sq.has_lift.iff_op CategoryTheory.CommSq.HasLift.iff_op theorem iff_unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} - (sq : CommSq f i p g) : HasLift sq ↔ HasLift sq.unop := - by + (sq : CommSq f i p g) : HasLift sq ↔ HasLift sq.unop := by rw [Iff, Iff] exact Nonempty.congr (lift_struct.unop_equiv sq).toFun (lift_struct.unop_equiv sq).invFun #align category_theory.comm_sq.has_lift.iff_unop CategoryTheory.CommSq.HasLift.iff_unop From ca10bfb783814ba6eae8969b2efac1f518f912d7 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 21:02:54 +0100 Subject: [PATCH 41/66] first pass --- Mathlib/CategoryTheory/CommSq.lean | 37 ++++++++++++++++-------------- 1 file changed, 20 insertions(+), 17 deletions(-) diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index b0e9f9499d38c..b4990b4b9a476 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -48,7 +48,7 @@ structure CommSq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y w : f ≫ h = g ≫ i #align category_theory.comm_sq CategoryTheory.CommSq -attribute [reassoc.1] comm_sq.w +attribute [reassoc] CommSq.w namespace CommSq @@ -58,7 +58,7 @@ theorem flip (p : CommSq f g h i) : CommSq g f i h := ⟨p.w.symm⟩ #align category_theory.comm_sq.flip CategoryTheory.CommSq.flip -theorem of_arrow {f g : Arrow C} (h : f ⟶ g) : CommSq f.Hom h.left h.right g.Hom := +theorem of_arrow {f g : Arrow C} (h : f ⟶ g) : CommSq f.hom h.left h.right g.hom := ⟨h.w.symm⟩ #align category_theory.comm_sq.of_arrow CategoryTheory.CommSq.of_arrow @@ -87,7 +87,7 @@ theorem map_commSq (s : CommSq f g h i) : CommSq (F.map f) (F.map g) (F.map h) ( end Functor -alias functor.map_comm_sq ← comm_sq.map +alias Functor.map_commSq ← CommSq.map #align category_theory.comm_sq.map CategoryTheory.CommSq.map namespace CommSq @@ -96,7 +96,8 @@ variable {A B X Y : C} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} /-- The datum of a lift in a commutative square, i.e. a up-right-diagonal morphism which makes both triangles commute. -/ -@[ext, nolint has_nonempty_instance] +-- Porting note: removed @[nolint has_nonempty_instance] +@[ext] structure LiftStruct (sq : CommSq f i p g) where l : B ⟶ X fac_left' : i ≫ l = f @@ -137,8 +138,8 @@ def opEquiv (sq : CommSq f i p g) : LiftStruct sq ≃ LiftStruct sq.op where toFun := op invFun := unop - left_inv := by tidy - right_inv := by tidy + left_inv := by aesop_cat + right_inv := by aesop_cat #align category_theory.comm_sq.lift_struct.op_equiv CategoryTheory.CommSq.LiftStruct.opEquiv /-- Equivalences of `lift_struct` for a square in the oppositive category and @@ -148,8 +149,8 @@ def unopEquiv {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : where toFun := unop invFun := op - left_inv := by tidy - right_inv := by tidy + left_inv := by aesop_cat + right_inv := by aesop_cat #align category_theory.comm_sq.lift_struct.unop_equiv CategoryTheory.CommSq.LiftStruct.unopEquiv end LiftStruct @@ -158,14 +159,17 @@ instance subsingleton_liftStruct_of_epi (sq : CommSq f i p g) [Epi i] : Subsingleton (LiftStruct sq) := ⟨fun l₁ l₂ => by ext - simp only [← cancel_epi i, lift_struct.fac_left]⟩ + rw [← cancel_epi i] + simp only [LiftStruct.fac_left] + ⟩ #align category_theory.comm_sq.subsingleton_lift_struct_of_epi CategoryTheory.CommSq.subsingleton_liftStruct_of_epi instance subsingleton_liftStruct_of_mono (sq : CommSq f i p g) [Mono p] : Subsingleton (LiftStruct sq) := ⟨fun l₁ l₂ => by ext - simp only [← cancel_mono p, lift_struct.fac_right]⟩ + rw [← cancel_mono p] + simp only [LiftStruct.fac_right]⟩ #align category_theory.comm_sq.subsingleton_lift_struct_of_mono CategoryTheory.CommSq.subsingleton_liftStruct_of_mono variable (sq : CommSq f i p g) @@ -191,14 +195,14 @@ theorem iff : HasLift sq ↔ Nonempty sq.LiftStruct := by #align category_theory.comm_sq.has_lift.iff CategoryTheory.CommSq.HasLift.iff theorem iff_op : HasLift sq ↔ HasLift sq.op := by - rw [Iff, Iff] - exact Nonempty.congr (lift_struct.op_equiv sq).toFun (lift_struct.op_equiv sq).invFun + rw [iff, iff] + exact Nonempty.congr (LiftStruct.opEquiv sq).toFun (LiftStruct.opEquiv sq).invFun #align category_theory.comm_sq.has_lift.iff_op CategoryTheory.CommSq.HasLift.iff_op theorem iff_unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} (sq : CommSq f i p g) : HasLift sq ↔ HasLift sq.unop := by - rw [Iff, Iff] - exact Nonempty.congr (lift_struct.unop_equiv sq).toFun (lift_struct.unop_equiv sq).invFun + rw [iff, iff] + exact Nonempty.congr (LiftStruct.unopEquiv sq).toFun (LiftStruct.unopEquiv sq).invFun #align category_theory.comm_sq.has_lift.iff_unop CategoryTheory.CommSq.HasLift.iff_unop end HasLift @@ -209,12 +213,12 @@ noncomputable def lift [hsq : HasLift sq] : B ⟶ X := hsq.exists_lift.some.l #align category_theory.comm_sq.lift CategoryTheory.CommSq.lift -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem fac_left [hsq : HasLift sq] : i ≫ sq.lift = f := hsq.exists_lift.some.fac_left #align category_theory.comm_sq.fac_left CategoryTheory.CommSq.fac_left -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem fac_right [hsq : HasLift sq] : sq.lift ≫ p = g := hsq.exists_lift.some.fac_right #align category_theory.comm_sq.fac_right CategoryTheory.CommSq.fac_right @@ -222,4 +226,3 @@ theorem fac_right [hsq : HasLift sq] : sq.lift ≫ p = g := end CommSq end CategoryTheory - From 2e1beb73c300a888ee6875801c111562ff2a543a Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 21:20:51 +0100 Subject: [PATCH 42/66] names and removing restate_axiom --- Mathlib/CategoryTheory/CommSq.lean | 25 +++++++++++-------------- 1 file changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index b4990b4b9a476..cd15b56718c84 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -15,15 +15,15 @@ import Mathlib.CategoryTheory.Arrow This file provide an API for commutative squares in categories. If `top`, `left`, `right` and `bottom` are four morphisms which are the edges -of a square, `comm_sq top left right bottom` is the predicate that this +of a square, `CommSq top left right bottom` is the predicate that this square is commutative. -The structure `comm_sq` is extended in `category_theory/shapes/limits/comm_sq.lean` -as `is_pullback` and `is_pushout` in order to define pullback and pushout squares. +The structure `CommSq` is extended in `CategoryTheory/Shapes/Limits/CommSq.lean` +as `IsPullback` and `IsPushout` in order to define pullback and pushout squares. ## Future work -Refactor `lift_struct` from `arrow.lean` and lifting properties using `comm_sq.lean`. +Refactor `LiftStruct` from `Arrow.lean` and lifting properties using `CommSq.lean`. -/ @@ -100,24 +100,20 @@ morphism which makes both triangles commute. -/ @[ext] structure LiftStruct (sq : CommSq f i p g) where l : B ⟶ X - fac_left' : i ≫ l = f - fac_right' : l ≫ p = g + fac_left : i ≫ l = f + fac_right: l ≫ p = g #align category_theory.comm_sq.lift_struct CategoryTheory.CommSq.LiftStruct namespace LiftStruct -restate_axiom fac_left' - -restate_axiom fac_right' - /-- A `lift_struct` for a commutative square gives a `lift_struct` for the corresponding square in the opposite category. -/ @[simps] def op {sq : CommSq f i p g} (l : LiftStruct sq) : LiftStruct sq.op where l := l.l.op - fac_left' := by rw [← op_comp, l.fac_right] - fac_right' := by rw [← op_comp, l.fac_left] + fac_left := by rw [← op_comp, l.fac_right] + fac_right := by rw [← op_comp, l.fac_left] #align category_theory.comm_sq.lift_struct.op CategoryTheory.CommSq.LiftStruct.op /-- A `lift_struct` for a commutative square in the opposite category @@ -127,8 +123,8 @@ def unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B (l : LiftStruct sq) : LiftStruct sq.unop where l := l.l.unop - fac_left' := by rw [← unop_comp, l.fac_right] - fac_right' := by rw [← unop_comp, l.fac_left] + fac_left := by rw [← unop_comp, l.fac_right] + fac_right := by rw [← unop_comp, l.fac_left] #align category_theory.comm_sq.lift_struct.unop CategoryTheory.CommSq.LiftStruct.unop /-- Equivalences of `lift_struct` for a square and the corresponding square @@ -226,3 +222,4 @@ theorem fac_right [hsq : HasLift sq] : sq.lift ≫ p = g := end CommSq end CategoryTheory +#lint From 45532e0e360404240de1f7548aa8f17cd8043765 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 21:38:40 +0100 Subject: [PATCH 43/66] fix lint --- Mathlib/CategoryTheory/CommSq.lean | 20 ++++++++++++++++++-- Mathlib/CategoryTheory/Functor/Category.lean | 1 + 2 files changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index cd15b56718c84..12ffba8bc2641 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -45,6 +45,7 @@ variable {C : Type _} [Category C] is a commuting square. -/ structure CommSq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y ⟶ Z) : Prop where + /-- The square commutes. -/ w : f ≫ h = g ≫ i #align category_theory.comm_sq CategoryTheory.CommSq @@ -92,15 +93,29 @@ alias Functor.map_commSq ← CommSq.map namespace CommSq + variable {A B X Y : C} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} -/-- The datum of a lift in a commutative square, i.e. a up-right-diagonal +/-- Now we consider a square: +``` + A ---f---> X + | | + i p + | | + v v + B ---g---> Y +``` + +The datum of a lift in a commutative square, i.e. a up-right-diagonal morphism which makes both triangles commute. -/ -- Porting note: removed @[nolint has_nonempty_instance] @[ext] structure LiftStruct (sq : CommSq f i p g) where + /-- The lift. -/ l : B ⟶ X + /-- The upper left triangle commutes. -/ fac_left : i ≫ l = f + /-- The lower right triangle commutes. -/ fac_right: l ≫ p = g #align category_theory.comm_sq.lift_struct CategoryTheory.CommSq.LiftStruct @@ -170,8 +185,10 @@ instance subsingleton_liftStruct_of_mono (sq : CommSq f i p g) [Mono p] : variable (sq : CommSq f i p g) + /-- The assertion that a square has a `lift_struct`. -/ class HasLift : Prop where + /-- Square has a `lift_struct`. -/ exists_lift : Nonempty sq.LiftStruct #align category_theory.comm_sq.has_lift CategoryTheory.CommSq.HasLift @@ -222,4 +239,3 @@ theorem fac_right [hsq : HasLift sq] : sq.lift ≫ p = g := end CommSq end CategoryTheory -#lint diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index 70904ae51e0f3..cff7a6d6d0aa1 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -168,3 +168,4 @@ theorem map_inv_hom_app (F : C ⥤ D ⥤ E) {X Y : C} (e : X ≅ Y) (Z : D) : #align category_theory.map_inv_hom_app_assoc CategoryTheory.map_inv_hom_app_assoc end CategoryTheory + From 332636e77bb683528ae78319fde579f5161ea414 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 16 Feb 2023 21:43:02 +0100 Subject: [PATCH 44/66] remove spurious edit --- Mathlib/CategoryTheory/Functor/Category.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index cff7a6d6d0aa1..70904ae51e0f3 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -168,4 +168,3 @@ theorem map_inv_hom_app (F : C ⥤ D ⥤ E) {X Y : C} (e : X ≅ Y) (Z : D) : #align category_theory.map_inv_hom_app_assoc CategoryTheory.map_inv_hom_app_assoc end CategoryTheory - From 4968e1930ea358a28f15e5c91ee3893641e13f84 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:25:50 -0500 Subject: [PATCH 45/66] feat: port CategoryTheory.LiftingProperties.Basic From e49ab39736ed2ab5eb8f0932df657b9f172be70a Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:25:51 -0500 Subject: [PATCH 46/66] Initial file copy from mathport --- .../LiftingProperties/Basic.lean | 153 ++++++++++++++++++ 1 file changed, 153 insertions(+) create mode 100644 Mathlib/CategoryTheory/LiftingProperties/Basic.lean diff --git a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean new file mode 100644 index 0000000000000..555c99dcf8892 --- /dev/null +++ b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean @@ -0,0 +1,153 @@ +/- +Copyright (c) 2021 Jakob Scholbach. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob Scholbach, Joël Riou + +! This file was ported from Lean 3 source module category_theory.lifting_properties.basic +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.CommSq + +/-! +# Lifting properties + +This file defines the lifting property of two morphisms in a category and +shows basic properties of this notion. + +## Main results +- `has_lifting_property`: the definition of the lifting property + +## Tags +lifting property + +@TODO : +1) define llp/rlp with respect to a `morphism_property` +2) retracts, direct/inverse images, (co)products, adjunctions + +-/ + + +universe v + +namespace CategoryTheory + +open Category + +variable {C : Type _} [Category C] {A B B' X Y Y' : C} (i : A ⟶ B) (i' : B ⟶ B') (p : X ⟶ Y) + (p' : Y ⟶ Y') + +/-- `has_lifting_property i p` means that `i` has the left lifting +property with respect to `p`, or equivalently that `p` has +the right lifting property with respect to `i`. -/ +class HasLiftingProperty : Prop where + sq_hasLift : ∀ {f : A ⟶ X} {g : B ⟶ Y} (sq : CommSq f i p g), sq.HasLift +#align category_theory.has_lifting_property CategoryTheory.HasLiftingProperty + +instance (priority := 100) sq_hasLift_of_hasLiftingProperty {f : A ⟶ X} {g : B ⟶ Y} + (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := by apply hip.sq_has_lift +#align category_theory.sq_has_lift_of_has_lifting_property CategoryTheory.sq_hasLift_of_hasLiftingProperty + +namespace HasLiftingProperty + +variable {i p} + +theorem op (h : HasLiftingProperty i p) : HasLiftingProperty p.op i.op := + ⟨fun f g sq => by + simp only [comm_sq.has_lift.iff_unop, Quiver.Hom.unop_op] + infer_instance⟩ +#align category_theory.has_lifting_property.op CategoryTheory.HasLiftingProperty.op + +theorem unop {A B X Y : Cᵒᵖ} {i : A ⟶ B} {p : X ⟶ Y} (h : HasLiftingProperty i p) : + HasLiftingProperty p.unop i.unop := + ⟨fun f g sq => by + rw [comm_sq.has_lift.iff_op] + simp only [Quiver.Hom.op_unop] + infer_instance⟩ +#align category_theory.has_lifting_property.unop CategoryTheory.HasLiftingProperty.unop + +theorem iff_op : HasLiftingProperty i p ↔ HasLiftingProperty p.op i.op := + ⟨op, unop⟩ +#align category_theory.has_lifting_property.iff_op CategoryTheory.HasLiftingProperty.iff_op + +theorem iff_unop {A B X Y : Cᵒᵖ} (i : A ⟶ B) (p : X ⟶ Y) : + HasLiftingProperty i p ↔ HasLiftingProperty p.unop i.unop := + ⟨unop, op⟩ +#align category_theory.has_lifting_property.iff_unop CategoryTheory.HasLiftingProperty.iff_unop + +variable (i p) + +instance (priority := 100) of_left_iso [IsIso i] : HasLiftingProperty i p := + ⟨fun f g sq => + CommSq.HasLift.mk' + { l := inv i ≫ f + fac_left' := by simp only [is_iso.hom_inv_id_assoc] + fac_right' := by simp only [sq.w, assoc, is_iso.inv_hom_id_assoc] }⟩ +#align category_theory.has_lifting_property.of_left_iso CategoryTheory.HasLiftingProperty.of_left_iso + +instance (priority := 100) of_right_iso [IsIso p] : HasLiftingProperty i p := + ⟨fun f g sq => + CommSq.HasLift.mk' + { l := g ≫ inv p + fac_left' := by simp only [← sq.w_assoc, is_iso.hom_inv_id, comp_id] + fac_right' := by simp only [assoc, is_iso.inv_hom_id, comp_id] }⟩ +#align category_theory.has_lifting_property.of_right_iso CategoryTheory.HasLiftingProperty.of_right_iso + +instance of_comp_left [HasLiftingProperty i p] [HasLiftingProperty i' p] : + HasLiftingProperty (i ≫ i') p := + ⟨fun f g sq => by + have fac := sq.w + rw [assoc] at fac + exact + comm_sq.has_lift.mk' + { l := (comm_sq.mk (comm_sq.mk fac).fac_right).lift + fac_left' := by simp only [assoc, comm_sq.fac_left] + fac_right' := by simp only [comm_sq.fac_right] }⟩ +#align category_theory.has_lifting_property.of_comp_left CategoryTheory.HasLiftingProperty.of_comp_left + +instance of_comp_right [HasLiftingProperty i p] [HasLiftingProperty i p'] : + HasLiftingProperty i (p ≫ p') := + ⟨fun f g sq => by + have fac := sq.w + rw [← assoc] at fac + let sq₂ := (comm_sq.mk (comm_sq.mk fac).fac_left.symm).lift + exact + comm_sq.has_lift.mk' + { l := (comm_sq.mk (comm_sq.mk fac).fac_left.symm).lift + fac_left' := by simp only [comm_sq.fac_left] + fac_right' := by simp only [comm_sq.fac_right_assoc, comm_sq.fac_right] }⟩ +#align category_theory.has_lifting_property.of_comp_right CategoryTheory.HasLiftingProperty.of_comp_right + +theorem of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} + (e : Arrow.mk i ≅ Arrow.mk i') (p : X ⟶ Y) [hip : HasLiftingProperty i p] : + HasLiftingProperty i' p := by + rw [arrow.iso_w' e] + infer_instance +#align category_theory.has_lifting_property.of_arrow_iso_left CategoryTheory.HasLiftingProperty.of_arrow_iso_left + +theorem of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} + (e : Arrow.mk p ≅ Arrow.mk p') [hip : HasLiftingProperty i p] : HasLiftingProperty i p' := + by + rw [arrow.iso_w' e] + infer_instance +#align category_theory.has_lifting_property.of_arrow_iso_right CategoryTheory.HasLiftingProperty.of_arrow_iso_right + +theorem iff_of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} + (e : Arrow.mk i ≅ Arrow.mk i') (p : X ⟶ Y) : HasLiftingProperty i p ↔ HasLiftingProperty i' p := + by + constructor <;> intro + exacts[of_arrow_iso_left e p, of_arrow_iso_left e.symm p] +#align category_theory.has_lifting_property.iff_of_arrow_iso_left CategoryTheory.HasLiftingProperty.iff_of_arrow_iso_left + +theorem iff_of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} + (e : Arrow.mk p ≅ Arrow.mk p') : HasLiftingProperty i p ↔ HasLiftingProperty i p' := + by + constructor <;> intro + exacts[of_arrow_iso_right i e, of_arrow_iso_right i e.symm] +#align category_theory.has_lifting_property.iff_of_arrow_iso_right CategoryTheory.HasLiftingProperty.iff_of_arrow_iso_right + +end HasLiftingProperty + +end CategoryTheory + From bd1174bfeac59e971bdd020bc8d5bada65d5c45b Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:25:51 -0500 Subject: [PATCH 47/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/LiftingProperties/Basic.lean | 8 +++----- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..b3871b58cb599 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -245,6 +245,7 @@ import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso +import Mathlib.CategoryTheory.LiftingProperties.Basic import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites diff --git a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean index 555c99dcf8892..08d78bd9dc45d 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean @@ -8,7 +8,7 @@ Authors: Jakob Scholbach, Joël Riou ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.CommSq +import Mathlib.CategoryTheory.CommSq /-! # Lifting properties @@ -127,8 +127,7 @@ theorem of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} #align category_theory.has_lifting_property.of_arrow_iso_left CategoryTheory.HasLiftingProperty.of_arrow_iso_left theorem of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} - (e : Arrow.mk p ≅ Arrow.mk p') [hip : HasLiftingProperty i p] : HasLiftingProperty i p' := - by + (e : Arrow.mk p ≅ Arrow.mk p') [hip : HasLiftingProperty i p] : HasLiftingProperty i p' := by rw [arrow.iso_w' e] infer_instance #align category_theory.has_lifting_property.of_arrow_iso_right CategoryTheory.HasLiftingProperty.of_arrow_iso_right @@ -141,8 +140,7 @@ theorem iff_of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} #align category_theory.has_lifting_property.iff_of_arrow_iso_left CategoryTheory.HasLiftingProperty.iff_of_arrow_iso_left theorem iff_of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} - (e : Arrow.mk p ≅ Arrow.mk p') : HasLiftingProperty i p ↔ HasLiftingProperty i p' := - by + (e : Arrow.mk p ≅ Arrow.mk p') : HasLiftingProperty i p ↔ HasLiftingProperty i p' := by constructor <;> intro exacts[of_arrow_iso_right i e, of_arrow_iso_right i e.symm] #align category_theory.has_lifting_property.iff_of_arrow_iso_right CategoryTheory.HasLiftingProperty.iff_of_arrow_iso_right From 6a1bbf0b253e18332c2458143174db9bb204d636 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:44:09 -0500 Subject: [PATCH 48/66] fix errors; lint --- .../LiftingProperties/Basic.lean | 55 ++++++++++--------- 1 file changed, 28 insertions(+), 27 deletions(-) diff --git a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean index 08d78bd9dc45d..7fc094780277e 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean @@ -17,7 +17,7 @@ This file defines the lifting property of two morphisms in a category and shows basic properties of this notion. ## Main results -- `has_lifting_property`: the definition of the lifting property +- `HasLiftingProperty`: the definition of the lifting property ## Tags lifting property @@ -38,15 +38,17 @@ open Category variable {C : Type _} [Category C] {A B B' X Y Y' : C} (i : A ⟶ B) (i' : B ⟶ B') (p : X ⟶ Y) (p' : Y ⟶ Y') -/-- `has_lifting_property i p` means that `i` has the left lifting +/-- `HasLiftingProperty i p` means that `i` has the left lifting property with respect to `p`, or equivalently that `p` has the right lifting property with respect to `i`. -/ class HasLiftingProperty : Prop where + /-- Unique field expressing the any commutative square built from `f` and `g` has a lift -/ sq_hasLift : ∀ {f : A ⟶ X} {g : B ⟶ Y} (sq : CommSq f i p g), sq.HasLift #align category_theory.has_lifting_property CategoryTheory.HasLiftingProperty +#align category_theory.has_lifting_property.sq_has_lift CategoryTheory.HasLiftingProperty.sq_hasLift instance (priority := 100) sq_hasLift_of_hasLiftingProperty {f : A ⟶ X} {g : B ⟶ Y} - (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := by apply hip.sq_has_lift + (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := by apply hip.sq_hasLift #align category_theory.sq_has_lift_of_has_lifting_property CategoryTheory.sq_hasLift_of_hasLiftingProperty namespace HasLiftingProperty @@ -54,15 +56,15 @@ namespace HasLiftingProperty variable {i p} theorem op (h : HasLiftingProperty i p) : HasLiftingProperty p.op i.op := - ⟨fun f g sq => by - simp only [comm_sq.has_lift.iff_unop, Quiver.Hom.unop_op] + ⟨fun {f} {g} sq => by + simp only [CommSq.HasLift.iff_unop, Quiver.Hom.unop_op] infer_instance⟩ #align category_theory.has_lifting_property.op CategoryTheory.HasLiftingProperty.op theorem unop {A B X Y : Cᵒᵖ} {i : A ⟶ B} {p : X ⟶ Y} (h : HasLiftingProperty i p) : HasLiftingProperty p.unop i.unop := - ⟨fun f g sq => by - rw [comm_sq.has_lift.iff_op] + ⟨fun {f} {g} sq => by + rw [CommSq.HasLift.iff_op] simp only [Quiver.Hom.op_unop] infer_instance⟩ #align category_theory.has_lifting_property.unop CategoryTheory.HasLiftingProperty.unop @@ -79,56 +81,56 @@ theorem iff_unop {A B X Y : Cᵒᵖ} (i : A ⟶ B) (p : X ⟶ Y) : variable (i p) instance (priority := 100) of_left_iso [IsIso i] : HasLiftingProperty i p := - ⟨fun f g sq => + ⟨fun {f} {g} sq => CommSq.HasLift.mk' { l := inv i ≫ f - fac_left' := by simp only [is_iso.hom_inv_id_assoc] - fac_right' := by simp only [sq.w, assoc, is_iso.inv_hom_id_assoc] }⟩ + fac_left := by simp only [IsIso.hom_inv_id_assoc] + fac_right := by simp only [sq.w, assoc, IsIso.inv_hom_id_assoc] }⟩ #align category_theory.has_lifting_property.of_left_iso CategoryTheory.HasLiftingProperty.of_left_iso instance (priority := 100) of_right_iso [IsIso p] : HasLiftingProperty i p := - ⟨fun f g sq => + ⟨fun {f} {g} sq => CommSq.HasLift.mk' { l := g ≫ inv p - fac_left' := by simp only [← sq.w_assoc, is_iso.hom_inv_id, comp_id] - fac_right' := by simp only [assoc, is_iso.inv_hom_id, comp_id] }⟩ + fac_left := by simp only [← sq.w_assoc, IsIso.hom_inv_id, comp_id] + fac_right := by simp only [assoc, IsIso.inv_hom_id, comp_id] }⟩ #align category_theory.has_lifting_property.of_right_iso CategoryTheory.HasLiftingProperty.of_right_iso instance of_comp_left [HasLiftingProperty i p] [HasLiftingProperty i' p] : HasLiftingProperty (i ≫ i') p := - ⟨fun f g sq => by + ⟨fun {f} {g} sq => by have fac := sq.w rw [assoc] at fac exact - comm_sq.has_lift.mk' - { l := (comm_sq.mk (comm_sq.mk fac).fac_right).lift - fac_left' := by simp only [assoc, comm_sq.fac_left] - fac_right' := by simp only [comm_sq.fac_right] }⟩ + CommSq.HasLift.mk' + { l := (CommSq.mk (CommSq.mk fac).fac_right).lift + fac_left := by simp only [assoc, CommSq.fac_left] + fac_right := by simp only [CommSq.fac_right] }⟩ #align category_theory.has_lifting_property.of_comp_left CategoryTheory.HasLiftingProperty.of_comp_left instance of_comp_right [HasLiftingProperty i p] [HasLiftingProperty i p'] : HasLiftingProperty i (p ≫ p') := - ⟨fun f g sq => by + ⟨fun {f} {g} sq => by have fac := sq.w rw [← assoc] at fac - let sq₂ := (comm_sq.mk (comm_sq.mk fac).fac_left.symm).lift + let _ := (CommSq.mk (CommSq.mk fac).fac_left.symm).lift exact - comm_sq.has_lift.mk' - { l := (comm_sq.mk (comm_sq.mk fac).fac_left.symm).lift - fac_left' := by simp only [comm_sq.fac_left] - fac_right' := by simp only [comm_sq.fac_right_assoc, comm_sq.fac_right] }⟩ + CommSq.HasLift.mk' + { l := (CommSq.mk (CommSq.mk fac).fac_left.symm).lift + fac_left := by simp only [CommSq.fac_left] + fac_right := by simp only [CommSq.fac_right_assoc, CommSq.fac_right] }⟩ #align category_theory.has_lifting_property.of_comp_right CategoryTheory.HasLiftingProperty.of_comp_right theorem of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} (e : Arrow.mk i ≅ Arrow.mk i') (p : X ⟶ Y) [hip : HasLiftingProperty i p] : HasLiftingProperty i' p := by - rw [arrow.iso_w' e] + rw [Arrow.iso_w' e] infer_instance #align category_theory.has_lifting_property.of_arrow_iso_left CategoryTheory.HasLiftingProperty.of_arrow_iso_left theorem of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} (e : Arrow.mk p ≅ Arrow.mk p') [hip : HasLiftingProperty i p] : HasLiftingProperty i p' := by - rw [arrow.iso_w' e] + rw [Arrow.iso_w' e] infer_instance #align category_theory.has_lifting_property.of_arrow_iso_right CategoryTheory.HasLiftingProperty.of_arrow_iso_right @@ -148,4 +150,3 @@ theorem iff_of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} { end HasLiftingProperty end CategoryTheory - From 4ebb6e167f419bf5ae0ae3dd1428b6a307bb4500 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:54:09 -0500 Subject: [PATCH 49/66] feat: port CategoryTheory.Limits.Shapes.StrongEpi From 1979ee1cc3e427404bba9a46653d436f51ee2960 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:54:09 -0500 Subject: [PATCH 50/66] Initial file copy from mathport --- .../Limits/Shapes/StrongEpi.lean | 246 ++++++++++++++++++ 1 file changed, 246 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean new file mode 100644 index 0000000000000..be937b430ea76 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -0,0 +1,246 @@ +/- +Copyright (c) 2020 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel + +! This file was ported from Lean 3 source module category_theory.limits.shapes.strong_epi +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Balanced +import Mathbin.CategoryTheory.LiftingProperties.Basic + +/-! +# Strong epimorphisms + +In this file, we define strong epimorphisms. A strong epimorphism is an epimorphism `f` +which has the (unique) left lifting property with respect to monomorphisms. Similarly, +a strong monomorphisms in a monomorphism which has the (unique) right lifting property +with respect to epimorphisms. + +## Main results + +Besides the definition, we show that +* the composition of two strong epimorphisms is a strong epimorphism, +* if `f ≫ g` is a strong epimorphism, then so is `g`, +* if `f` is both a strong epimorphism and a monomorphism, then it is an isomorphism + +We also define classes `strong_mono_category` and `strong_epi_category` for categories in which +every monomorphism or epimorphism is strong, and deduce that these categories are balanced. + +## TODO + +Show that the dual of a strong epimorphism is a strong monomorphism, and vice versa. + +## References + +* [F. Borceux, *Handbook of Categorical Algebra 1*][borceux-vol1] +-/ + + +universe v u + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] + +variable {P Q : C} + +/-- A strong epimorphism `f` is an epimorphism which has the left lifting property +with respect to monomorphisms. -/ +class StrongEpi (f : P ⟶ Q) : Prop where + Epi : Epi f + llp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Mono z], HasLiftingProperty f z +#align category_theory.strong_epi CategoryTheory.StrongEpi + +theorem StrongEpi.mk' {f : P ⟶ Q} [Epi f] + (hf : + ∀ (X Y : C) (z : X ⟶ Y) (hz : Mono z) (u : P ⟶ X) (v : Q ⟶ Y) (sq : CommSq u f z v), + sq.HasLift) : + StrongEpi f := + { Epi := inferInstance + llp := fun X Y z hz => ⟨fun u v sq => hf X Y z hz u v sq⟩ } +#align category_theory.strong_epi.mk' CategoryTheory.StrongEpi.mk' + +/-- A strong monomorphism `f` is a monomorphism which has the right lifting property +with respect to epimorphisms. -/ +class StrongMono (f : P ⟶ Q) : Prop where + Mono : Mono f + rlp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Epi z], HasLiftingProperty z f +#align category_theory.strong_mono CategoryTheory.StrongMono + +theorem StrongMono.mk' {f : P ⟶ Q} [Mono f] + (hf : + ∀ (X Y : C) (z : X ⟶ Y) (hz : Epi z) (u : X ⟶ P) (v : Y ⟶ Q) (sq : CommSq u z f v), + sq.HasLift) : + StrongMono f := + { Mono := inferInstance + rlp := fun X Y z hz => ⟨fun u v sq => hf X Y z hz u v sq⟩ } +#align category_theory.strong_mono.mk' CategoryTheory.StrongMono.mk' + +attribute [instance] strong_epi.llp + +attribute [instance] strong_mono.rlp + +instance (priority := 100) epi_of_strongEpi (f : P ⟶ Q) [StrongEpi f] : Epi f := + StrongEpi.epi +#align category_theory.epi_of_strong_epi CategoryTheory.epi_of_strongEpi + +instance (priority := 100) mono_of_strongMono (f : P ⟶ Q) [StrongMono f] : Mono f := + StrongMono.mono +#align category_theory.mono_of_strong_mono CategoryTheory.mono_of_strongMono + +section + +variable {R : C} (f : P ⟶ Q) (g : Q ⟶ R) + +/-- The composition of two strong epimorphisms is a strong epimorphism. -/ +theorem strongEpi_comp [StrongEpi f] [StrongEpi g] : StrongEpi (f ≫ g) := + { Epi := epi_comp _ _ + llp := by + intros + infer_instance } +#align category_theory.strong_epi_comp CategoryTheory.strongEpi_comp + +/-- The composition of two strong monomorphisms is a strong monomorphism. -/ +theorem strongMono_comp [StrongMono f] [StrongMono g] : StrongMono (f ≫ g) := + { Mono := mono_comp _ _ + rlp := by + intros + infer_instance } +#align category_theory.strong_mono_comp CategoryTheory.strongMono_comp + +/-- If `f ≫ g` is a strong epimorphism, then so is `g`. -/ +theorem strongEpi_of_strongEpi [StrongEpi (f ≫ g)] : StrongEpi g := + { Epi := epi_of_epi f g + llp := by + intros + constructor + intro u v sq + have h₀ : (f ≫ u) ≫ z = (f ≫ g) ≫ v := by simp only [category.assoc, sq.w] + exact + comm_sq.has_lift.mk' + ⟨(comm_sq.mk h₀).lift, by + simp only [← cancel_mono z, category.assoc, comm_sq.fac_right, sq.w], by simp⟩ } +#align category_theory.strong_epi_of_strong_epi CategoryTheory.strongEpi_of_strongEpi + +/-- If `f ≫ g` is a strong monomorphism, then so is `f`. -/ +theorem strongMono_of_strongMono [StrongMono (f ≫ g)] : StrongMono f := + { Mono := mono_of_mono f g + rlp := by + intros + constructor + intro u v sq + have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by rw [reassoc_of sq.w] + exact comm_sq.has_lift.mk' ⟨(comm_sq.mk h₀).lift, by simp, by simp [← cancel_epi z, sq.w]⟩ } +#align category_theory.strong_mono_of_strong_mono CategoryTheory.strongMono_of_strongMono + +/-- An isomorphism is in particular a strong epimorphism. -/ +instance (priority := 100) strongEpi_of_isIso [IsIso f] : StrongEpi f + where + Epi := by infer_instance + llp X Y z hz := HasLiftingProperty.of_left_iso _ _ +#align category_theory.strong_epi_of_is_iso CategoryTheory.strongEpi_of_isIso + +/-- An isomorphism is in particular a strong monomorphism. -/ +instance (priority := 100) strongMono_of_isIso [IsIso f] : StrongMono f + where + Mono := by infer_instance + rlp X Y z hz := HasLiftingProperty.of_right_iso _ _ +#align category_theory.strong_mono_of_is_iso CategoryTheory.strongMono_of_isIso + +theorem StrongEpi.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongEpi f] : StrongEpi g := + { Epi := by + rw [arrow.iso_w' e] + haveI := epi_comp f e.hom.right + apply epi_comp + llp := fun X Y z => by + intro + apply has_lifting_property.of_arrow_iso_left e z } +#align category_theory.strong_epi.of_arrow_iso CategoryTheory.StrongEpi.of_arrow_iso + +theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongMono f] : StrongMono g := + { Mono := by + rw [arrow.iso_w' e] + haveI := mono_comp f e.hom.right + apply mono_comp + rlp := fun X Y z => by + intro + apply has_lifting_property.of_arrow_iso_right z e } +#align category_theory.strong_mono.of_arrow_iso CategoryTheory.StrongMono.of_arrow_iso + +theorem StrongEpi.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : Arrow.mk f ≅ Arrow.mk g) : StrongEpi f ↔ StrongEpi g := + by + constructor <;> intro + exacts[strong_epi.of_arrow_iso e, strong_epi.of_arrow_iso e.symm] +#align category_theory.strong_epi.iff_of_arrow_iso CategoryTheory.StrongEpi.iff_of_arrow_iso + +theorem StrongMono.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : Arrow.mk f ≅ Arrow.mk g) : StrongMono f ↔ StrongMono g := + by + constructor <;> intro + exacts[strong_mono.of_arrow_iso e, strong_mono.of_arrow_iso e.symm] +#align category_theory.strong_mono.iff_of_arrow_iso CategoryTheory.StrongMono.iff_of_arrow_iso + +end + +/-- A strong epimorphism that is a monomorphism is an isomorphism. -/ +theorem isIso_of_mono_of_strongEpi (f : P ⟶ Q) [Mono f] [StrongEpi f] : IsIso f := + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by tidy⟩⟩ +#align category_theory.is_iso_of_mono_of_strong_epi CategoryTheory.isIso_of_mono_of_strongEpi + +/-- A strong monomorphism that is an epimorphism is an isomorphism. -/ +theorem isIso_of_epi_of_strongMono (f : P ⟶ Q) [Epi f] [StrongMono f] : IsIso f := + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by tidy⟩⟩ +#align category_theory.is_iso_of_epi_of_strong_mono CategoryTheory.isIso_of_epi_of_strongMono + +section + +variable (C) + +/-- A strong epi category is a category in which every epimorphism is strong. -/ +class StrongEpiCategory : Prop where + strongEpi_of_epi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], StrongEpi f +#align category_theory.strong_epi_category CategoryTheory.StrongEpiCategory + +/-- A strong mono category is a category in which every monomorphism is strong. -/ +class StrongMonoCategory : Prop where + strongMono_of_mono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], StrongMono f +#align category_theory.strong_mono_category CategoryTheory.StrongMonoCategory + +end + +theorem strongEpi_of_epi [StrongEpiCategory C] (f : P ⟶ Q) [Epi f] : StrongEpi f := + StrongEpiCategory.strongEpi_of_epi _ +#align category_theory.strong_epi_of_epi CategoryTheory.strongEpi_of_epi + +theorem strongMono_of_mono [StrongMonoCategory C] (f : P ⟶ Q) [Mono f] : StrongMono f := + StrongMonoCategory.strongMono_of_mono _ +#align category_theory.strong_mono_of_mono CategoryTheory.strongMono_of_mono + +section + +attribute [local instance] strong_epi_of_epi + +instance (priority := 100) balanced_of_strongEpiCategory [StrongEpiCategory C] : Balanced C + where isIso_of_mono_of_epi _ _ _ _ _ := is_iso_of_mono_of_strong_epi _ +#align category_theory.balanced_of_strong_epi_category CategoryTheory.balanced_of_strongEpiCategory + +end + +section + +attribute [local instance] strong_mono_of_mono + +instance (priority := 100) balanced_of_strongMonoCategory [StrongMonoCategory C] : Balanced C + where isIso_of_mono_of_epi _ _ _ _ _ := is_iso_of_epi_of_strong_mono _ +#align category_theory.balanced_of_strong_mono_category CategoryTheory.balanced_of_strongMonoCategory + +end + +end CategoryTheory + From d9d7366134998e4fac6df9bb9c98113d454ac39d Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 18:54:09 -0500 Subject: [PATCH 51/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean | 10 ++++------ 2 files changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..03102f881ae69 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -245,6 +245,7 @@ import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso +import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index be937b430ea76..08ab8806e2220 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -8,8 +8,8 @@ Authors: Markus Himmel ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Balanced -import Mathbin.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Balanced +import Mathlib.CategoryTheory.LiftingProperties.Basic /-! # Strong epimorphisms @@ -173,15 +173,13 @@ theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} #align category_theory.strong_mono.of_arrow_iso CategoryTheory.StrongMono.of_arrow_iso theorem StrongEpi.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} - (e : Arrow.mk f ≅ Arrow.mk g) : StrongEpi f ↔ StrongEpi g := - by + (e : Arrow.mk f ≅ Arrow.mk g) : StrongEpi f ↔ StrongEpi g := by constructor <;> intro exacts[strong_epi.of_arrow_iso e, strong_epi.of_arrow_iso e.symm] #align category_theory.strong_epi.iff_of_arrow_iso CategoryTheory.StrongEpi.iff_of_arrow_iso theorem StrongMono.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} - (e : Arrow.mk f ≅ Arrow.mk g) : StrongMono f ↔ StrongMono g := - by + (e : Arrow.mk f ≅ Arrow.mk g) : StrongMono f ↔ StrongMono g := by constructor <;> intro exacts[strong_mono.of_arrow_iso e, strong_mono.of_arrow_iso e.symm] #align category_theory.strong_mono.iff_of_arrow_iso CategoryTheory.StrongMono.iff_of_arrow_iso From 118d7ed86fd869f52d3f1a07923a4ebcf2242be8 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 19:57:39 -0500 Subject: [PATCH 52/66] fix errors; lint --- .../Limits/Shapes/StrongEpi.lean | 103 +++++++++--------- 1 file changed, 54 insertions(+), 49 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index 08ab8806e2220..a71d94949664b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -50,38 +50,41 @@ variable {P Q : C} /-- A strong epimorphism `f` is an epimorphism which has the left lifting property with respect to monomorphisms. -/ class StrongEpi (f : P ⟶ Q) : Prop where - Epi : Epi f + /-- The epimorphism condition on `f` -/ + epi : Epi f + /-- The left lifting property with respect to all monomorphism -/ llp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Mono z], HasLiftingProperty f z #align category_theory.strong_epi CategoryTheory.StrongEpi +#align category_theory.strong_epi.epi CategoryTheory.StrongEpi.epi + theorem StrongEpi.mk' {f : P ⟶ Q} [Epi f] - (hf : - ∀ (X Y : C) (z : X ⟶ Y) (hz : Mono z) (u : P ⟶ X) (v : Q ⟶ Y) (sq : CommSq u f z v), - sq.HasLift) : + (hf : ∀ (X Y : C) (z : X ⟶ Y) + (_ : Mono z) (u : P ⟶ X) (v : Q ⟶ Y) (sq : CommSq u f z v), sq.HasLift) : StrongEpi f := - { Epi := inferInstance - llp := fun X Y z hz => ⟨fun u v sq => hf X Y z hz u v sq⟩ } + { epi := inferInstance + llp := fun {X} {Y} z hz => ⟨fun {u} {v} sq => hf X Y z hz u v sq⟩ } #align category_theory.strong_epi.mk' CategoryTheory.StrongEpi.mk' /-- A strong monomorphism `f` is a monomorphism which has the right lifting property with respect to epimorphisms. -/ class StrongMono (f : P ⟶ Q) : Prop where - Mono : Mono f + /-- The monomorphism condition on `f` -/ + mono : Mono f + /-- The right lifting property with respect to all epimorphisms -/ rlp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Epi z], HasLiftingProperty z f #align category_theory.strong_mono CategoryTheory.StrongMono theorem StrongMono.mk' {f : P ⟶ Q} [Mono f] - (hf : - ∀ (X Y : C) (z : X ⟶ Y) (hz : Epi z) (u : X ⟶ P) (v : Y ⟶ Q) (sq : CommSq u z f v), - sq.HasLift) : - StrongMono f := - { Mono := inferInstance - rlp := fun X Y z hz => ⟨fun u v sq => hf X Y z hz u v sq⟩ } + (hf : ∀ (X Y : C) (z : X ⟶ Y) (_ : Epi z) (u : X ⟶ P) + (v : Y ⟶ Q) (sq : CommSq u z f v), sq.HasLift) : StrongMono f where + mono := inferInstance + rlp := fun {X} {Y} z hz => ⟨fun {u} {v} sq => hf X Y z hz u v sq⟩ #align category_theory.strong_mono.mk' CategoryTheory.StrongMono.mk' -attribute [instance] strong_epi.llp +attribute [instance] StrongEpi.llp -attribute [instance] strong_mono.rlp +attribute [instance] StrongMono.rlp instance (priority := 100) epi_of_strongEpi (f : P ⟶ Q) [StrongEpi f] : Epi f := StrongEpi.epi @@ -97,7 +100,7 @@ variable {R : C} (f : P ⟶ Q) (g : Q ⟶ R) /-- The composition of two strong epimorphisms is a strong epimorphism. -/ theorem strongEpi_comp [StrongEpi f] [StrongEpi g] : StrongEpi (f ≫ g) := - { Epi := epi_comp _ _ + { epi := epi_comp _ _ llp := by intros infer_instance } @@ -105,7 +108,7 @@ theorem strongEpi_comp [StrongEpi f] [StrongEpi g] : StrongEpi (f ≫ g) := /-- The composition of two strong monomorphisms is a strong monomorphism. -/ theorem strongMono_comp [StrongMono f] [StrongMono g] : StrongMono (f ≫ g) := - { Mono := mono_comp _ _ + { mono := mono_comp _ _ rlp := by intros infer_instance } @@ -113,87 +116,87 @@ theorem strongMono_comp [StrongMono f] [StrongMono g] : StrongMono (f ≫ g) := /-- If `f ≫ g` is a strong epimorphism, then so is `g`. -/ theorem strongEpi_of_strongEpi [StrongEpi (f ≫ g)] : StrongEpi g := - { Epi := epi_of_epi f g - llp := by - intros + { epi := epi_of_epi f g + llp := fun {X} {Y} z _ => by constructor intro u v sq - have h₀ : (f ≫ u) ≫ z = (f ≫ g) ≫ v := by simp only [category.assoc, sq.w] + have h₀ : (f ≫ u) ≫ z = (f ≫ g) ≫ v := by simp only [Category.assoc, sq.w] exact - comm_sq.has_lift.mk' - ⟨(comm_sq.mk h₀).lift, by - simp only [← cancel_mono z, category.assoc, comm_sq.fac_right, sq.w], by simp⟩ } + CommSq.HasLift.mk' + ⟨(CommSq.mk h₀).lift, by + simp only [← cancel_mono z, Category.assoc, CommSq.fac_right, sq.w], by simp⟩ } #align category_theory.strong_epi_of_strong_epi CategoryTheory.strongEpi_of_strongEpi /-- If `f ≫ g` is a strong monomorphism, then so is `f`. -/ theorem strongMono_of_strongMono [StrongMono (f ≫ g)] : StrongMono f := - { Mono := mono_of_mono f g - rlp := by + { mono := mono_of_mono f g + rlp := fun {X} {Y} z => by intros constructor intro u v sq - have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by rw [reassoc_of sq.w] - exact comm_sq.has_lift.mk' ⟨(comm_sq.mk h₀).lift, by simp, by simp [← cancel_epi z, sq.w]⟩ } + have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by + rw [←Category.assoc, eq_whisker sq.w, Category.assoc] + exact CommSq.HasLift.mk' ⟨(CommSq.mk h₀).lift, by simp, by simp [← cancel_epi z, sq.w]⟩ } #align category_theory.strong_mono_of_strong_mono CategoryTheory.strongMono_of_strongMono /-- An isomorphism is in particular a strong epimorphism. -/ instance (priority := 100) strongEpi_of_isIso [IsIso f] : StrongEpi f where - Epi := by infer_instance - llp X Y z hz := HasLiftingProperty.of_left_iso _ _ + epi := by infer_instance + llp {X} {Y} z := HasLiftingProperty.of_left_iso _ _ #align category_theory.strong_epi_of_is_iso CategoryTheory.strongEpi_of_isIso /-- An isomorphism is in particular a strong monomorphism. -/ instance (priority := 100) strongMono_of_isIso [IsIso f] : StrongMono f where - Mono := by infer_instance - rlp X Y z hz := HasLiftingProperty.of_right_iso _ _ + mono := by infer_instance + rlp {X} {Y} z := HasLiftingProperty.of_right_iso _ _ #align category_theory.strong_mono_of_is_iso CategoryTheory.strongMono_of_isIso theorem StrongEpi.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongEpi f] : StrongEpi g := - { Epi := by - rw [arrow.iso_w' e] + { epi := by + rw [Arrow.iso_w' e] haveI := epi_comp f e.hom.right apply epi_comp - llp := fun X Y z => by + llp := fun {X} {Y} z => by intro - apply has_lifting_property.of_arrow_iso_left e z } + apply HasLiftingProperty.of_arrow_iso_left e z } #align category_theory.strong_epi.of_arrow_iso CategoryTheory.StrongEpi.of_arrow_iso theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongMono f] : StrongMono g := - { Mono := by - rw [arrow.iso_w' e] + { mono := by + rw [Arrow.iso_w' e] haveI := mono_comp f e.hom.right apply mono_comp - rlp := fun X Y z => by + rlp := fun {X} {Y} z => by intro - apply has_lifting_property.of_arrow_iso_right z e } + apply HasLiftingProperty.of_arrow_iso_right z e } #align category_theory.strong_mono.of_arrow_iso CategoryTheory.StrongMono.of_arrow_iso theorem StrongEpi.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} (e : Arrow.mk f ≅ Arrow.mk g) : StrongEpi f ↔ StrongEpi g := by constructor <;> intro - exacts[strong_epi.of_arrow_iso e, strong_epi.of_arrow_iso e.symm] + exacts[StrongEpi.of_arrow_iso e, StrongEpi.of_arrow_iso e.symm] #align category_theory.strong_epi.iff_of_arrow_iso CategoryTheory.StrongEpi.iff_of_arrow_iso theorem StrongMono.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} (e : Arrow.mk f ≅ Arrow.mk g) : StrongMono f ↔ StrongMono g := by constructor <;> intro - exacts[strong_mono.of_arrow_iso e, strong_mono.of_arrow_iso e.symm] + exacts[StrongMono.of_arrow_iso e, StrongMono.of_arrow_iso e.symm] #align category_theory.strong_mono.iff_of_arrow_iso CategoryTheory.StrongMono.iff_of_arrow_iso end /-- A strong epimorphism that is a monomorphism is an isomorphism. -/ theorem isIso_of_mono_of_strongEpi (f : P ⟶ Q) [Mono f] [StrongEpi f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by tidy⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ #align category_theory.is_iso_of_mono_of_strong_epi CategoryTheory.isIso_of_mono_of_strongEpi /-- A strong monomorphism that is an epimorphism is an isomorphism. -/ theorem isIso_of_epi_of_strongMono (f : P ⟶ Q) [Epi f] [StrongMono f] : IsIso f := - ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by tidy⟩⟩ + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ #align category_theory.is_iso_of_epi_of_strong_mono CategoryTheory.isIso_of_epi_of_strongMono section @@ -202,11 +205,13 @@ variable (C) /-- A strong epi category is a category in which every epimorphism is strong. -/ class StrongEpiCategory : Prop where + /-- A strong epi category is a category in which every epimorphism is strong. -/ strongEpi_of_epi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], StrongEpi f #align category_theory.strong_epi_category CategoryTheory.StrongEpiCategory /-- A strong mono category is a category in which every monomorphism is strong. -/ class StrongMonoCategory : Prop where + /-- A strong mono category is a category in which every monomorphism is strong. -/ strongMono_of_mono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], StrongMono f #align category_theory.strong_mono_category CategoryTheory.StrongMonoCategory @@ -222,23 +227,23 @@ theorem strongMono_of_mono [StrongMonoCategory C] (f : P ⟶ Q) [Mono f] : Stron section -attribute [local instance] strong_epi_of_epi +attribute [local instance] strongEpi_of_epi instance (priority := 100) balanced_of_strongEpiCategory [StrongEpiCategory C] : Balanced C - where isIso_of_mono_of_epi _ _ _ _ _ := is_iso_of_mono_of_strong_epi _ + where isIso_of_mono_of_epi _ _ _ := isIso_of_mono_of_strongEpi _ #align category_theory.balanced_of_strong_epi_category CategoryTheory.balanced_of_strongEpiCategory end section -attribute [local instance] strong_mono_of_mono +attribute [local instance] strongMono_of_mono instance (priority := 100) balanced_of_strongMonoCategory [StrongMonoCategory C] : Balanced C - where isIso_of_mono_of_epi _ _ _ _ _ := is_iso_of_epi_of_strong_mono _ + where isIso_of_mono_of_epi _ _ _ := isIso_of_epi_of_strongMono _ #align category_theory.balanced_of_strong_mono_category CategoryTheory.balanced_of_strongMonoCategory end end CategoryTheory - +#lint From 5e558ae23775522c1c7774bada74f2a2477c7e4d Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> Date: Thu, 16 Feb 2023 20:14:04 -0500 Subject: [PATCH 53/66] Update Mathlib.lean --- Mathlib.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index 89cd9f20d1ef3..9db1b0dc6b36f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -246,8 +246,8 @@ import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso -import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi import Mathlib.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites From 12c4a61f8b1e0a74354e0e36a022320f822b915c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:28:27 -0500 Subject: [PATCH 54/66] feat: port CategoryTheory.LiftingProperties.Adjunction From fc177c8bfcfb00763f829f6c32ccdac3ae6121c3 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:28:27 -0500 Subject: [PATCH 55/66] Initial file copy from mathport --- .../LiftingProperties/Adjunction.lean | 152 ++++++++++++++++++ 1 file changed, 152 insertions(+) create mode 100644 Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean diff --git a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean new file mode 100644 index 0000000000000..3c36ed9072f29 --- /dev/null +++ b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean @@ -0,0 +1,152 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou + +! This file was ported from Lean 3 source module category_theory.lifting_properties.adjunction +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.LiftingProperties.Basic +import Mathbin.CategoryTheory.Adjunction.Basic + +/-! + +# Lifting properties and adjunction + +In this file, we obtain `adjunction.has_lifting_property_iff`, which states +that when we have an adjunction `adj : G ⊣ F` between two functors `G : C ⥤ D` +and `F : D ⥤ C`, then a morphism of the form `G.map i` has the left lifting +property in `D` with respect to a morphism `p` if and only the morphism `i` +has the left lifting property in `C` with respect to `F.map p`. + +-/ + + +namespace CategoryTheory + +open Category + +variable {C D : Type _} [Category C] [Category D] {G : C ⥤ D} {F : D ⥤ C} + +namespace CommSq + +section + +variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : G.obj A ⟶ X} {v : G.obj B ⟶ Y} + (sq : CommSq u (G.map i) p v) (adj : G ⊣ F) + +include sq + +/-- When we have an adjunction `G ⊣ F`, any commutative square where the left +map is of the form `G.map i` and the right map is `p` has an "adjoint" commutative +square whose left map is `i` and whose right map is `F.map p`. -/ +theorem right_adjoint : CommSq (adj.homEquiv _ _ u) i (F.map p) (adj.homEquiv _ _ v) := + ⟨by + simp only [adjunction.hom_equiv_unit, assoc, ← F.map_comp, sq.w] + rw [F.map_comp, adjunction.unit_naturality_assoc]⟩ +#align category_theory.comm_sq.right_adjoint CategoryTheory.CommSq.right_adjoint + +/-- The liftings of a commutative are in bijection with the liftings of its (right) +adjoint square. -/ +def rightAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.rightAdjoint adj).LiftStruct + where + toFun l := + { l := adj.homEquiv _ _ l.l + fac_left' := by rw [← adj.hom_equiv_naturality_left, l.fac_left] + fac_right' := by rw [← adjunction.hom_equiv_naturality_right, l.fac_right] } + invFun l := + { l := (adj.homEquiv _ _).symm l.l + fac_left' := by + rw [← adjunction.hom_equiv_naturality_left_symm, l.fac_left] + apply (adj.hom_equiv _ _).left_inv + fac_right' := by + rw [← adjunction.hom_equiv_naturality_right_symm, l.fac_right] + apply (adj.hom_equiv _ _).left_inv } + left_inv := by tidy + right_inv := by tidy +#align category_theory.comm_sq.right_adjoint_lift_struct_equiv CategoryTheory.CommSq.rightAdjointLiftStructEquiv + +/-- A square has a lifting if and only if its (right) adjoint square has a lifting. -/ +theorem right_adjoint_hasLift_iff : HasLift (sq.rightAdjoint adj) ↔ HasLift sq := + by + simp only [has_lift.iff] + exact Equiv.nonempty_congr (sq.right_adjoint_lift_struct_equiv adj).symm +#align category_theory.comm_sq.right_adjoint_has_lift_iff CategoryTheory.CommSq.right_adjoint_hasLift_iff + +instance [HasLift sq] : HasLift (sq.rightAdjoint adj) := + by + rw [right_adjoint_has_lift_iff] + infer_instance + +end + +section + +variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : A ⟶ F.obj X} {v : B ⟶ F.obj Y} + (sq : CommSq u i (F.map p) v) (adj : G ⊣ F) + +include sq + +/-- When we have an adjunction `G ⊣ F`, any commutative square where the left +map is of the form `i` and the right map is `F.map p` has an "adjoint" commutative +square whose left map is `G.map i` and whose right map is `p`. -/ +theorem left_adjoint : CommSq ((adj.homEquiv _ _).symm u) (G.map i) p ((adj.homEquiv _ _).symm v) := + ⟨by + simp only [adjunction.hom_equiv_counit, assoc, ← G.map_comp_assoc, ← sq.w] + rw [G.map_comp, assoc, adjunction.counit_naturality]⟩ +#align category_theory.comm_sq.left_adjoint CategoryTheory.CommSq.left_adjoint + +/-- The liftings of a commutative are in bijection with the liftings of its (left) +adjoint square. -/ +def leftAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.leftAdjoint adj).LiftStruct + where + toFun l := + { l := (adj.homEquiv _ _).symm l.l + fac_left' := by rw [← adj.hom_equiv_naturality_left_symm, l.fac_left] + fac_right' := by rw [← adj.hom_equiv_naturality_right_symm, l.fac_right] } + invFun l := + { l := (adj.homEquiv _ _) l.l + fac_left' := by + rw [← adj.hom_equiv_naturality_left, l.fac_left] + apply (adj.hom_equiv _ _).right_inv + fac_right' := by + rw [← adj.hom_equiv_naturality_right, l.fac_right] + apply (adj.hom_equiv _ _).right_inv } + left_inv := by tidy + right_inv := by tidy +#align category_theory.comm_sq.left_adjoint_lift_struct_equiv CategoryTheory.CommSq.leftAdjointLiftStructEquiv + +/-- A (left) adjoint square has a lifting if and only if the original square has a lifting. -/ +theorem left_adjoint_hasLift_iff : HasLift (sq.leftAdjoint adj) ↔ HasLift sq := + by + simp only [has_lift.iff] + exact Equiv.nonempty_congr (sq.left_adjoint_lift_struct_equiv adj).symm +#align category_theory.comm_sq.left_adjoint_has_lift_iff CategoryTheory.CommSq.left_adjoint_hasLift_iff + +instance [HasLift sq] : HasLift (sq.leftAdjoint adj) := + by + rw [left_adjoint_has_lift_iff] + infer_instance + +end + +end CommSq + +namespace Adjunction + +theorem hasLiftingProperty_iff (adj : G ⊣ F) {A B : C} {X Y : D} (i : A ⟶ B) (p : X ⟶ Y) : + HasLiftingProperty (G.map i) p ↔ HasLiftingProperty i (F.map p) := + by + constructor <;> intro <;> constructor <;> intro f g sq + · rw [← sq.left_adjoint_has_lift_iff adj] + infer_instance + · rw [← sq.right_adjoint_has_lift_iff adj] + infer_instance +#align category_theory.adjunction.has_lifting_property_iff CategoryTheory.Adjunction.hasLiftingProperty_iff + +end Adjunction + +end CategoryTheory + From f80301f4b52e3561d9cd011a37214151d164fb17 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:28:27 -0500 Subject: [PATCH 56/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + .../LiftingProperties/Adjunction.lean | 19 +++++++------------ 2 files changed, 8 insertions(+), 12 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..1ffd1b1ea583e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -245,6 +245,7 @@ import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso +import Mathlib.CategoryTheory.LiftingProperties.Adjunction import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites diff --git a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean index 3c36ed9072f29..e43790cfa2a11 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean @@ -8,8 +8,8 @@ Authors: Joël Riou ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.LiftingProperties.Basic -import Mathbin.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Adjunction.Basic /-! @@ -69,14 +69,12 @@ def rightAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.rightAdjoint adj).LiftSt #align category_theory.comm_sq.right_adjoint_lift_struct_equiv CategoryTheory.CommSq.rightAdjointLiftStructEquiv /-- A square has a lifting if and only if its (right) adjoint square has a lifting. -/ -theorem right_adjoint_hasLift_iff : HasLift (sq.rightAdjoint adj) ↔ HasLift sq := - by +theorem right_adjoint_hasLift_iff : HasLift (sq.rightAdjoint adj) ↔ HasLift sq := by simp only [has_lift.iff] exact Equiv.nonempty_congr (sq.right_adjoint_lift_struct_equiv adj).symm #align category_theory.comm_sq.right_adjoint_has_lift_iff CategoryTheory.CommSq.right_adjoint_hasLift_iff -instance [HasLift sq] : HasLift (sq.rightAdjoint adj) := - by +instance [HasLift sq] : HasLift (sq.rightAdjoint adj) := by rw [right_adjoint_has_lift_iff] infer_instance @@ -119,14 +117,12 @@ def leftAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.leftAdjoint adj).LiftStru #align category_theory.comm_sq.left_adjoint_lift_struct_equiv CategoryTheory.CommSq.leftAdjointLiftStructEquiv /-- A (left) adjoint square has a lifting if and only if the original square has a lifting. -/ -theorem left_adjoint_hasLift_iff : HasLift (sq.leftAdjoint adj) ↔ HasLift sq := - by +theorem left_adjoint_hasLift_iff : HasLift (sq.leftAdjoint adj) ↔ HasLift sq := by simp only [has_lift.iff] exact Equiv.nonempty_congr (sq.left_adjoint_lift_struct_equiv adj).symm #align category_theory.comm_sq.left_adjoint_has_lift_iff CategoryTheory.CommSq.left_adjoint_hasLift_iff -instance [HasLift sq] : HasLift (sq.leftAdjoint adj) := - by +instance [HasLift sq] : HasLift (sq.leftAdjoint adj) := by rw [left_adjoint_has_lift_iff] infer_instance @@ -137,8 +133,7 @@ end CommSq namespace Adjunction theorem hasLiftingProperty_iff (adj : G ⊣ F) {A B : C} {X Y : D} (i : A ⟶ B) (p : X ⟶ Y) : - HasLiftingProperty (G.map i) p ↔ HasLiftingProperty i (F.map p) := - by + HasLiftingProperty (G.map i) p ↔ HasLiftingProperty i (F.map p) := by constructor <;> intro <;> constructor <;> intro f g sq · rw [← sq.left_adjoint_has_lift_iff adj] infer_instance From e296a228228bacf1f27d6f9ac7a9c0478a05bea2 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:43:39 -0500 Subject: [PATCH 57/66] fix errors; lint --- .../LiftingProperties/Adjunction.lean | 82 +++++++++---------- 1 file changed, 39 insertions(+), 43 deletions(-) diff --git a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean index e43790cfa2a11..39edcc3687f81 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean @@ -15,7 +15,7 @@ import Mathlib.CategoryTheory.Adjunction.Basic # Lifting properties and adjunction -In this file, we obtain `adjunction.has_lifting_property_iff`, which states +In this file, we obtain `Adjunction.HasLiftingProperty_iff`, which states that when we have an adjunction `adj : G ⊣ F` between two functors `G : C ⥤ D` and `F : D ⥤ C`, then a morphism of the form `G.map i` has the left lifting property in `D` with respect to a morphism `p` if and only the morphism `i` @@ -37,45 +37,43 @@ section variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : G.obj A ⟶ X} {v : G.obj B ⟶ Y} (sq : CommSq u (G.map i) p v) (adj : G ⊣ F) -include sq - /-- When we have an adjunction `G ⊣ F`, any commutative square where the left map is of the form `G.map i` and the right map is `p` has an "adjoint" commutative square whose left map is `i` and whose right map is `F.map p`. -/ theorem right_adjoint : CommSq (adj.homEquiv _ _ u) i (F.map p) (adj.homEquiv _ _ v) := ⟨by - simp only [adjunction.hom_equiv_unit, assoc, ← F.map_comp, sq.w] - rw [F.map_comp, adjunction.unit_naturality_assoc]⟩ + simp only [Adjunction.homEquiv_unit, assoc, ← F.map_comp, sq.w] + rw [F.map_comp, Adjunction.unit_naturality_assoc]⟩ #align category_theory.comm_sq.right_adjoint CategoryTheory.CommSq.right_adjoint /-- The liftings of a commutative are in bijection with the liftings of its (right) adjoint square. -/ -def rightAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.rightAdjoint adj).LiftStruct +def rightAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.right_adjoint adj).LiftStruct where toFun l := { l := adj.homEquiv _ _ l.l - fac_left' := by rw [← adj.hom_equiv_naturality_left, l.fac_left] - fac_right' := by rw [← adjunction.hom_equiv_naturality_right, l.fac_right] } + fac_left := by rw [← adj.homEquiv_naturality_left, l.fac_left] + fac_right := by rw [← Adjunction.homEquiv_naturality_right, l.fac_right] } invFun l := { l := (adj.homEquiv _ _).symm l.l - fac_left' := by - rw [← adjunction.hom_equiv_naturality_left_symm, l.fac_left] - apply (adj.hom_equiv _ _).left_inv - fac_right' := by - rw [← adjunction.hom_equiv_naturality_right_symm, l.fac_right] - apply (adj.hom_equiv _ _).left_inv } - left_inv := by tidy - right_inv := by tidy + fac_left := by + rw [← Adjunction.homEquiv_naturality_left_symm, l.fac_left] + apply (adj.homEquiv _ _).left_inv + fac_right := by + rw [← Adjunction.homEquiv_naturality_right_symm, l.fac_right] + apply (adj.homEquiv _ _).left_inv } + left_inv := by aesop_cat + right_inv := by aesop_cat #align category_theory.comm_sq.right_adjoint_lift_struct_equiv CategoryTheory.CommSq.rightAdjointLiftStructEquiv /-- A square has a lifting if and only if its (right) adjoint square has a lifting. -/ -theorem right_adjoint_hasLift_iff : HasLift (sq.rightAdjoint adj) ↔ HasLift sq := by - simp only [has_lift.iff] - exact Equiv.nonempty_congr (sq.right_adjoint_lift_struct_equiv adj).symm +theorem right_adjoint_hasLift_iff : HasLift (sq.right_adjoint adj) ↔ HasLift sq := by + simp only [HasLift.iff] + exact Equiv.nonempty_congr (sq.rightAdjointLiftStructEquiv adj).symm #align category_theory.comm_sq.right_adjoint_has_lift_iff CategoryTheory.CommSq.right_adjoint_hasLift_iff -instance [HasLift sq] : HasLift (sq.rightAdjoint adj) := by - rw [right_adjoint_has_lift_iff] +instance [HasLift sq] : HasLift (sq.right_adjoint adj) := by + rw [right_adjoint_hasLift_iff] infer_instance end @@ -85,45 +83,43 @@ section variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : A ⟶ F.obj X} {v : B ⟶ F.obj Y} (sq : CommSq u i (F.map p) v) (adj : G ⊣ F) -include sq - /-- When we have an adjunction `G ⊣ F`, any commutative square where the left map is of the form `i` and the right map is `F.map p` has an "adjoint" commutative square whose left map is `G.map i` and whose right map is `p`. -/ theorem left_adjoint : CommSq ((adj.homEquiv _ _).symm u) (G.map i) p ((adj.homEquiv _ _).symm v) := ⟨by - simp only [adjunction.hom_equiv_counit, assoc, ← G.map_comp_assoc, ← sq.w] - rw [G.map_comp, assoc, adjunction.counit_naturality]⟩ + simp only [Adjunction.homEquiv_counit, assoc, ← G.map_comp_assoc, ← sq.w] + rw [G.map_comp, assoc, Adjunction.counit_naturality]⟩ #align category_theory.comm_sq.left_adjoint CategoryTheory.CommSq.left_adjoint /-- The liftings of a commutative are in bijection with the liftings of its (left) adjoint square. -/ -def leftAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.leftAdjoint adj).LiftStruct +def leftAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.left_adjoint adj).LiftStruct where toFun l := { l := (adj.homEquiv _ _).symm l.l - fac_left' := by rw [← adj.hom_equiv_naturality_left_symm, l.fac_left] - fac_right' := by rw [← adj.hom_equiv_naturality_right_symm, l.fac_right] } + fac_left := by rw [← adj.homEquiv_naturality_left_symm, l.fac_left] + fac_right := by rw [← adj.homEquiv_naturality_right_symm, l.fac_right] } invFun l := { l := (adj.homEquiv _ _) l.l - fac_left' := by - rw [← adj.hom_equiv_naturality_left, l.fac_left] - apply (adj.hom_equiv _ _).right_inv - fac_right' := by - rw [← adj.hom_equiv_naturality_right, l.fac_right] - apply (adj.hom_equiv _ _).right_inv } - left_inv := by tidy - right_inv := by tidy + fac_left := by + rw [← adj.homEquiv_naturality_left, l.fac_left] + apply (adj.homEquiv _ _).right_inv + fac_right := by + rw [← adj.homEquiv_naturality_right, l.fac_right] + apply (adj.homEquiv _ _).right_inv } + left_inv := by aesop_cat + right_inv := by aesop_cat #align category_theory.comm_sq.left_adjoint_lift_struct_equiv CategoryTheory.CommSq.leftAdjointLiftStructEquiv /-- A (left) adjoint square has a lifting if and only if the original square has a lifting. -/ -theorem left_adjoint_hasLift_iff : HasLift (sq.leftAdjoint adj) ↔ HasLift sq := by - simp only [has_lift.iff] - exact Equiv.nonempty_congr (sq.left_adjoint_lift_struct_equiv adj).symm +theorem left_adjoint_hasLift_iff : HasLift (sq.left_adjoint adj) ↔ HasLift sq := by + simp only [HasLift.iff] + exact Equiv.nonempty_congr (sq.leftAdjointLiftStructEquiv adj).symm #align category_theory.comm_sq.left_adjoint_has_lift_iff CategoryTheory.CommSq.left_adjoint_hasLift_iff -instance [HasLift sq] : HasLift (sq.leftAdjoint adj) := by - rw [left_adjoint_has_lift_iff] +instance [HasLift sq] : HasLift (sq.left_adjoint adj) := by + rw [left_adjoint_hasLift_iff] infer_instance end @@ -135,9 +131,9 @@ namespace Adjunction theorem hasLiftingProperty_iff (adj : G ⊣ F) {A B : C} {X Y : D} (i : A ⟶ B) (p : X ⟶ Y) : HasLiftingProperty (G.map i) p ↔ HasLiftingProperty i (F.map p) := by constructor <;> intro <;> constructor <;> intro f g sq - · rw [← sq.left_adjoint_has_lift_iff adj] + · rw [← sq.left_adjoint_hasLift_iff adj] infer_instance - · rw [← sq.right_adjoint_has_lift_iff adj] + · rw [← sq.right_adjoint_hasLift_iff adj] infer_instance #align category_theory.adjunction.has_lifting_property_iff CategoryTheory.Adjunction.hasLiftingProperty_iff From 084cffd36c2ebc0cef803306cb46853280ba0668 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> Date: Thu, 16 Feb 2023 20:52:10 -0500 Subject: [PATCH 58/66] Delete start_port-macos.sh --- scripts/start_port-macos.sh | 86 ------------------------------------- 1 file changed, 86 deletions(-) delete mode 100755 scripts/start_port-macos.sh diff --git a/scripts/start_port-macos.sh b/scripts/start_port-macos.sh deleted file mode 100755 index da56887ff18fb..0000000000000 --- a/scripts/start_port-macos.sh +++ /dev/null @@ -1,86 +0,0 @@ -#!/usr/bin/env bash - -set -e - -if [ ! -d Mathlib ] ; then - echo "No Mathlib/ directory; are you at the root of the mathlib4 repo?" - exit 1 -fi - -if [ ! $1 ] ; then - echo "usage: ./scripts/start_port.sh Mathlib/Foo/Bar/Baz.lean" - exit 1 -fi - -case $1 in - Mathlib/*) true ;; - *) echo "argument must begin with Mathlib/" - exit 1 ;; -esac - -if [ -f $1 ] ; then - echo "file already exists" - exit 1 -fi - -set -x -set -o pipefail - -MATHLIB3PORT_BASE_URL=https://raw.githubusercontent.com/leanprover-community/mathlib3port/master -PORT_STATUS_YAML=https://raw.githubusercontent.com/wiki/leanprover-community/mathlib/mathlib4-port-status.md -TMP_FILE=start_port_tmp.lean - -mathlib4_path=$1 -mathlib4_mod=$(basename $(echo "$mathlib4_path" | tr / .) .lean) - -mathlib3port_url=$MATHLIB3PORT_BASE_URL/Mathbin/${1#Mathlib/} - -curl --silent --show-error --fail -o "$TMP_FILE" "$mathlib3port_url" - -mathlib3_module=$(grep '^! .*source module ' <"$TMP_FILE" | gsed 's/.*source module \(.*\)$/\1/') - -if curl --silent --show-error --fail "$PORT_STATUS_YAML" | grep -F "$mathlib3_module: " | grep "mathlib4#" ; then - set +x - echo "WARNING: The file is already in the process of being ported." - echo "(See line above for PR number.)" - rm "$TMP_FILE" - exit 1 -fi - -mkdir -p $(dirname "$mathlib4_path") -mv "$TMP_FILE" "$mathlib4_path" - -git fetch -mathlib4_mod_tail=${mathlib4_mod#Mathlib.} -branch_name=port/${mathlib4_mod_tail} -git checkout --no-track -b "$branch_name" origin/master - -# Empty commit with nice title. Ugsed by gh and hub to suggest PR title. -git commit -m "feat: port $mathlib4_mod_tail" --allow-empty - -git add "$mathlib4_path" -git commit -m 'Initial file copy from mathport' - -gsed -i 's/Mathbin\./Mathlib\./g' "$mathlib4_path" -gsed -i '/^import/{s/[.]Gcd/.GCD/g; s/[.]Modeq/.ModEq/g; s/[.]Nary/.NAry/g; s/[.]Peq/.PEq/g; s/[.]Pfun/.PFun/g; s/[.]Pnat/.PNat/g; s/[.]Smul/.SMul/g; s/[.]Zmod/.ZMod/g}' "$mathlib4_path" - -# gawk script taken from https://github.com/leanprover-community/mathlib4/pull/1523 -gawk '{do {{if (match($0, "^ by$") && length(p) < 98 && (!(match(p, "^[ \t]*--.*$")))) {p=p " by";} else {if (NR!=1) {print p}; p=$0}}} while (getline == 1) if (getline==0) print p}' "$mathlib4_path" > "$mathlib4_path.tmp" -mv "$mathlib4_path.tmp" "$mathlib4_path" - -(echo "import $mathlib4_mod" ; cat Mathlib.lean) | LC_ALL=C sort | uniq > Mathlib.lean.tmp -mv -f Mathlib.lean.tmp Mathlib.lean - -git add Mathlib.lean "$mathlib4_path" -git commit \ - -m 'automated fixes' \ - -m '' \ - -m 'Mathbin -> Mathlib' \ - -m 'fix certain import statements' \ - -m 'move "by" to end of line' \ - -m 'add import to Mathlib.lean' - -set +x - -echo "After pushing, you can open a PR at:" -echo "https://github.com/leanprover-community/mathlib4/compare/$branch_name?expand=1&title=feat:+port+$mathlib4_mod_tail&labels=mathlib-port" From b71ba2c67dc73e042653f5323707acf7f97fc40c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:54:23 -0500 Subject: [PATCH 59/66] feat: port CategoryTheory.Functor.EpiMono From 1aaf0634f5ef32b714115b0b9269e372c33b4c71 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:54:23 -0500 Subject: [PATCH 60/66] Initial file copy from mathport --- Mathlib/CategoryTheory/Functor/EpiMono.lean | 343 ++++++++++++++++++++ 1 file changed, 343 insertions(+) create mode 100644 Mathlib/CategoryTheory/Functor/EpiMono.lean diff --git a/Mathlib/CategoryTheory/Functor/EpiMono.lean b/Mathlib/CategoryTheory/Functor/EpiMono.lean new file mode 100644 index 0000000000000..066061e6fefb5 --- /dev/null +++ b/Mathlib/CategoryTheory/Functor/EpiMono.lean @@ -0,0 +1,343 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel + +! This file was ported from Lean 3 source module category_theory.functor.epi_mono +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.EpiMono +import Mathbin.CategoryTheory.Limits.Shapes.StrongEpi +import Mathbin.CategoryTheory.LiftingProperties.Adjunction + +/-! +# Preservation and reflection of monomorphisms and epimorphisms + +We provide typeclasses that state that a functor preserves or reflects monomorphisms or +epimorphisms. +-/ + + +open CategoryTheory + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory.Functor + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {E : Type u₃} + [Category.{v₃} E] + +/-- A functor preserves monomorphisms if it maps monomorphisms to monomorphisms. -/ +class PreservesMonomorphisms (F : C ⥤ D) : Prop where + preserves : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], Mono (F.map f) +#align category_theory.functor.preserves_monomorphisms CategoryTheory.Functor.PreservesMonomorphisms + +instance map_mono (F : C ⥤ D) [PreservesMonomorphisms F] {X Y : C} (f : X ⟶ Y) [Mono f] : + Mono (F.map f) := + PreservesMonomorphisms.preserves f +#align category_theory.functor.map_mono CategoryTheory.Functor.map_mono + +/-- A functor preserves epimorphisms if it maps epimorphisms to epimorphisms. -/ +class PreservesEpimorphisms (F : C ⥤ D) : Prop where + preserves : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], Epi (F.map f) +#align category_theory.functor.preserves_epimorphisms CategoryTheory.Functor.PreservesEpimorphisms + +instance map_epi (F : C ⥤ D) [PreservesEpimorphisms F] {X Y : C} (f : X ⟶ Y) [Epi f] : + Epi (F.map f) := + PreservesEpimorphisms.preserves f +#align category_theory.functor.map_epi CategoryTheory.Functor.map_epi + +/-- A functor reflects monomorphisms if morphisms that are mapped to monomorphisms are themselves + monomorphisms. -/ +class ReflectsMonomorphisms (F : C ⥤ D) : Prop where + reflects : ∀ {X Y : C} (f : X ⟶ Y), Mono (F.map f) → Mono f +#align category_theory.functor.reflects_monomorphisms CategoryTheory.Functor.ReflectsMonomorphisms + +theorem mono_of_mono_map (F : C ⥤ D) [ReflectsMonomorphisms F] {X Y : C} {f : X ⟶ Y} + (h : Mono (F.map f)) : Mono f := + ReflectsMonomorphisms.reflects f h +#align category_theory.functor.mono_of_mono_map CategoryTheory.Functor.mono_of_mono_map + +/-- A functor reflects epimorphisms if morphisms that are mapped to epimorphisms are themselves + epimorphisms. -/ +class ReflectsEpimorphisms (F : C ⥤ D) : Prop where + reflects : ∀ {X Y : C} (f : X ⟶ Y), Epi (F.map f) → Epi f +#align category_theory.functor.reflects_epimorphisms CategoryTheory.Functor.ReflectsEpimorphisms + +theorem epi_of_epi_map (F : C ⥤ D) [ReflectsEpimorphisms F] {X Y : C} {f : X ⟶ Y} + (h : Epi (F.map f)) : Epi f := + ReflectsEpimorphisms.reflects f h +#align category_theory.functor.epi_of_epi_map CategoryTheory.Functor.epi_of_epi_map + +instance preservesMonomorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [PreservesMonomorphisms F] + [PreservesMonomorphisms G] : PreservesMonomorphisms (F ⋙ G) + where preserves X Y f h := by + rw [comp_map] + exact inferInstance +#align category_theory.functor.preserves_monomorphisms_comp CategoryTheory.Functor.preservesMonomorphisms_comp + +instance preservesEpimorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [PreservesEpimorphisms F] + [PreservesEpimorphisms G] : PreservesEpimorphisms (F ⋙ G) + where preserves X Y f h := by + rw [comp_map] + exact inferInstance +#align category_theory.functor.preserves_epimorphisms_comp CategoryTheory.Functor.preservesEpimorphisms_comp + +instance reflectsMonomorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [ReflectsMonomorphisms F] + [ReflectsMonomorphisms G] : ReflectsMonomorphisms (F ⋙ G) + where reflects X Y f h := F.mono_of_mono_map (G.mono_of_mono_map h) +#align category_theory.functor.reflects_monomorphisms_comp CategoryTheory.Functor.reflectsMonomorphisms_comp + +instance reflectsEpimorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [ReflectsEpimorphisms F] + [ReflectsEpimorphisms G] : ReflectsEpimorphisms (F ⋙ G) + where reflects X Y f h := F.epi_of_epi_map (G.epi_of_epi_map h) +#align category_theory.functor.reflects_epimorphisms_comp CategoryTheory.Functor.reflectsEpimorphisms_comp + +theorem preservesEpimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [PreservesEpimorphisms (F ⋙ G)] [ReflectsEpimorphisms G] : PreservesEpimorphisms F := + ⟨fun X Y f hf => G.epi_of_epi_map <| show Epi ((F ⋙ G).map f) from inferInstance⟩ +#align category_theory.functor.preserves_epimorphisms_of_preserves_of_reflects CategoryTheory.Functor.preservesEpimorphisms_of_preserves_of_reflects + +theorem preservesMonomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [PreservesMonomorphisms (F ⋙ G)] [ReflectsMonomorphisms G] : PreservesMonomorphisms F := + ⟨fun X Y f hf => G.mono_of_mono_map <| show Mono ((F ⋙ G).map f) from inferInstance⟩ +#align category_theory.functor.preserves_monomorphisms_of_preserves_of_reflects CategoryTheory.Functor.preservesMonomorphisms_of_preserves_of_reflects + +theorem reflectsEpimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [PreservesEpimorphisms G] [ReflectsEpimorphisms (F ⋙ G)] : ReflectsEpimorphisms F := + ⟨fun X Y f hf => (F ⋙ G).epi_of_epi_map <| show Epi (G.map (F.map f)) from inferInstance⟩ +#align category_theory.functor.reflects_epimorphisms_of_preserves_of_reflects CategoryTheory.Functor.reflectsEpimorphisms_of_preserves_of_reflects + +theorem reflectsMonomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [PreservesMonomorphisms G] [ReflectsMonomorphisms (F ⋙ G)] : ReflectsMonomorphisms F := + ⟨fun X Y f hf => (F ⋙ G).mono_of_mono_map <| show Mono (G.map (F.map f)) from inferInstance⟩ +#align category_theory.functor.reflects_monomorphisms_of_preserves_of_reflects CategoryTheory.Functor.reflectsMonomorphisms_of_preserves_of_reflects + +theorem PreservesMonomorphisms.of_iso {F G : C ⥤ D} [PreservesMonomorphisms F] (α : F ≅ G) : + PreservesMonomorphisms G := + { + preserves := fun X Y f h => + by + haveI : mono (F.map f ≫ (α.app Y).Hom) := mono_comp _ _ + convert (mono_comp _ _ : mono ((α.app X).inv ≫ F.map f ≫ (α.app Y).Hom)) + rw [iso.eq_inv_comp, iso.app_hom, iso.app_hom, nat_trans.naturality] } +#align category_theory.functor.preserves_monomorphisms.of_iso CategoryTheory.Functor.PreservesMonomorphisms.of_iso + +theorem PreservesMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : + PreservesMonomorphisms F ↔ PreservesMonomorphisms G := + ⟨fun h => preserves_monomorphisms.of_iso α, fun h => preserves_monomorphisms.of_iso α.symm⟩ +#align category_theory.functor.preserves_monomorphisms.iso_iff CategoryTheory.Functor.PreservesMonomorphisms.iso_iff + +theorem PreservesEpimorphisms.of_iso {F G : C ⥤ D} [PreservesEpimorphisms F] (α : F ≅ G) : + PreservesEpimorphisms G := + { + preserves := fun X Y f h => + by + haveI : epi (F.map f ≫ (α.app Y).Hom) := epi_comp _ _ + convert (epi_comp _ _ : epi ((α.app X).inv ≫ F.map f ≫ (α.app Y).Hom)) + rw [iso.eq_inv_comp, iso.app_hom, iso.app_hom, nat_trans.naturality] } +#align category_theory.functor.preserves_epimorphisms.of_iso CategoryTheory.Functor.PreservesEpimorphisms.of_iso + +theorem PreservesEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : + PreservesEpimorphisms F ↔ PreservesEpimorphisms G := + ⟨fun h => preserves_epimorphisms.of_iso α, fun h => preserves_epimorphisms.of_iso α.symm⟩ +#align category_theory.functor.preserves_epimorphisms.iso_iff CategoryTheory.Functor.PreservesEpimorphisms.iso_iff + +theorem ReflectsMonomorphisms.of_iso {F G : C ⥤ D} [ReflectsMonomorphisms F] (α : F ≅ G) : + ReflectsMonomorphisms G := + { + reflects := fun X Y f h => by + apply F.mono_of_mono_map + haveI : mono (G.map f ≫ (α.app Y).inv) := mono_comp _ _ + convert (mono_comp _ _ : mono ((α.app X).Hom ≫ G.map f ≫ (α.app Y).inv)) + rw [← category.assoc, iso.eq_comp_inv, iso.app_hom, iso.app_hom, nat_trans.naturality] } +#align category_theory.functor.reflects_monomorphisms.of_iso CategoryTheory.Functor.ReflectsMonomorphisms.of_iso + +theorem ReflectsMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : + ReflectsMonomorphisms F ↔ ReflectsMonomorphisms G := + ⟨fun h => reflects_monomorphisms.of_iso α, fun h => reflects_monomorphisms.of_iso α.symm⟩ +#align category_theory.functor.reflects_monomorphisms.iso_iff CategoryTheory.Functor.ReflectsMonomorphisms.iso_iff + +theorem ReflectsEpimorphisms.of_iso {F G : C ⥤ D} [ReflectsEpimorphisms F] (α : F ≅ G) : + ReflectsEpimorphisms G := + { + reflects := fun X Y f h => by + apply F.epi_of_epi_map + haveI : epi (G.map f ≫ (α.app Y).inv) := epi_comp _ _ + convert (epi_comp _ _ : epi ((α.app X).Hom ≫ G.map f ≫ (α.app Y).inv)) + rw [← category.assoc, iso.eq_comp_inv, iso.app_hom, iso.app_hom, nat_trans.naturality] } +#align category_theory.functor.reflects_epimorphisms.of_iso CategoryTheory.Functor.ReflectsEpimorphisms.of_iso + +theorem ReflectsEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : + ReflectsEpimorphisms F ↔ ReflectsEpimorphisms G := + ⟨fun h => reflects_epimorphisms.of_iso α, fun h => reflects_epimorphisms.of_iso α.symm⟩ +#align category_theory.functor.reflects_epimorphisms.iso_iff CategoryTheory.Functor.ReflectsEpimorphisms.iso_iff + +theorem preserves_epimorphsisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : + PreservesEpimorphisms F := + { + preserves := fun X Y f hf => + ⟨by + intro Z g h H + replace H := congr_arg (adj.hom_equiv X Z) H + rwa [adj.hom_equiv_naturality_left, adj.hom_equiv_naturality_left, cancel_epi, + Equiv.apply_eq_iff_eq] at H⟩ } +#align category_theory.functor.preserves_epimorphsisms_of_adjunction CategoryTheory.Functor.preserves_epimorphsisms_of_adjunction + +instance (priority := 100) preservesEpimorphisms_of_isLeftAdjoint (F : C ⥤ D) [IsLeftAdjoint F] : + PreservesEpimorphisms F := + preserves_epimorphsisms_of_adjunction (Adjunction.ofLeftAdjoint F) +#align category_theory.functor.preserves_epimorphisms_of_is_left_adjoint CategoryTheory.Functor.preservesEpimorphisms_of_isLeftAdjoint + +theorem preservesMonomorphisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : + PreservesMonomorphisms G := + { + preserves := fun X Y f hf => + ⟨by + intro Z g h H + replace H := congr_arg (adj.hom_equiv Z Y).symm H + rwa [adj.hom_equiv_naturality_right_symm, adj.hom_equiv_naturality_right_symm, cancel_mono, + Equiv.apply_eq_iff_eq] at H⟩ } +#align category_theory.functor.preserves_monomorphisms_of_adjunction CategoryTheory.Functor.preservesMonomorphisms_of_adjunction + +instance (priority := 100) preservesMonomorphisms_of_isRightAdjoint (F : C ⥤ D) [IsRightAdjoint F] : + PreservesMonomorphisms F := + preservesMonomorphisms_of_adjunction (Adjunction.ofRightAdjoint F) +#align category_theory.functor.preserves_monomorphisms_of_is_right_adjoint CategoryTheory.Functor.preservesMonomorphisms_of_isRightAdjoint + +instance (priority := 100) reflectsMonomorphisms_of_faithful (F : C ⥤ D) [Faithful F] : + ReflectsMonomorphisms F + where reflects X Y f hf := + ⟨fun Z g h hgh => + F.map_injective ((cancel_mono (F.map f)).1 (by rw [← F.map_comp, hgh, F.map_comp]))⟩ +#align category_theory.functor.reflects_monomorphisms_of_faithful CategoryTheory.Functor.reflectsMonomorphisms_of_faithful + +instance (priority := 100) reflectsEpimorphisms_of_faithful (F : C ⥤ D) [Faithful F] : + ReflectsEpimorphisms F + where reflects X Y f hf := + ⟨fun Z g h hgh => + F.map_injective ((cancel_epi (F.map f)).1 (by rw [← F.map_comp, hgh, F.map_comp]))⟩ +#align category_theory.functor.reflects_epimorphisms_of_faithful CategoryTheory.Functor.reflectsEpimorphisms_of_faithful + +section + +variable (F : C ⥤ D) {X Y : C} (f : X ⟶ Y) + +/-- If `F` is a fully faithful functor, split epimorphisms are preserved and reflected by `F`. -/ +def splitEpiEquiv [Full F] [Faithful F] : SplitEpi f ≃ SplitEpi (F.map f) + where + toFun f := f.map F + invFun s := by + refine' ⟨F.preimage s.section_, _⟩ + apply F.map_injective + simp only [map_comp, image_preimage, map_id] + apply split_epi.id + left_inv := by tidy + right_inv := by tidy +#align category_theory.functor.split_epi_equiv CategoryTheory.Functor.splitEpiEquiv + +@[simp] +theorem isSplitEpi_iff [Full F] [Faithful F] : IsSplitEpi (F.map f) ↔ IsSplitEpi f := + by + constructor + · intro h + exact is_split_epi.mk' ((split_epi_equiv F f).invFun h.exists_split_epi.some) + · intro h + exact is_split_epi.mk' ((split_epi_equiv F f).toFun h.exists_split_epi.some) +#align category_theory.functor.is_split_epi_iff CategoryTheory.Functor.isSplitEpi_iff + +/-- If `F` is a fully faithful functor, split monomorphisms are preserved and reflected by `F`. -/ +def splitMonoEquiv [Full F] [Faithful F] : SplitMono f ≃ SplitMono (F.map f) + where + toFun f := f.map F + invFun s := by + refine' ⟨F.preimage s.retraction, _⟩ + apply F.map_injective + simp only [map_comp, image_preimage, map_id] + apply split_mono.id + left_inv := by tidy + right_inv := by tidy +#align category_theory.functor.split_mono_equiv CategoryTheory.Functor.splitMonoEquiv + +@[simp] +theorem isSplitMono_iff [Full F] [Faithful F] : IsSplitMono (F.map f) ↔ IsSplitMono f := + by + constructor + · intro h + exact is_split_mono.mk' ((split_mono_equiv F f).invFun h.exists_split_mono.some) + · intro h + exact is_split_mono.mk' ((split_mono_equiv F f).toFun h.exists_split_mono.some) +#align category_theory.functor.is_split_mono_iff CategoryTheory.Functor.isSplitMono_iff + +@[simp] +theorem epi_map_iff_epi [hF₁ : PreservesEpimorphisms F] [hF₂ : ReflectsEpimorphisms F] : + Epi (F.map f) ↔ Epi f := by + constructor + · exact F.epi_of_epi_map + · intro h + exact F.map_epi f +#align category_theory.functor.epi_map_iff_epi CategoryTheory.Functor.epi_map_iff_epi + +@[simp] +theorem mono_map_iff_mono [hF₁ : PreservesMonomorphisms F] [hF₂ : ReflectsMonomorphisms F] : + Mono (F.map f) ↔ Mono f := by + constructor + · exact F.mono_of_mono_map + · intro h + exact F.map_mono f +#align category_theory.functor.mono_map_iff_mono CategoryTheory.Functor.mono_map_iff_mono + +/-- If `F : C ⥤ D` is an equivalence of categories and `C` is a `split_epi_category`, +then `D` also is. -/ +def splitEpiCategoryImpOfIsEquivalence [IsEquivalence F] [SplitEpiCategory C] : + SplitEpiCategory D := + ⟨fun X Y f => by + intro + rw [← F.inv.is_split_epi_iff f] + apply is_split_epi_of_epi⟩ +#align category_theory.functor.split_epi_category_imp_of_is_equivalence CategoryTheory.Functor.splitEpiCategoryImpOfIsEquivalence + +end + +end CategoryTheory.Functor + +namespace CategoryTheory.Adjunction + +variable {C D : Type _} [Category C] [Category D] {F : C ⥤ D} {F' : D ⥤ C} {A B : C} + +theorem strongEpi_map_of_strongEpi (adj : F ⊣ F') (f : A ⟶ B) [h₁ : F'.PreservesMonomorphisms] + [h₂ : F.PreservesEpimorphisms] [StrongEpi f] : StrongEpi (F.map f) := + ⟨inferInstance, fun X Y Z => by + intro + rw [adj.has_lifting_property_iff] + infer_instance⟩ +#align category_theory.adjunction.strong_epi_map_of_strong_epi CategoryTheory.Adjunction.strongEpi_map_of_strongEpi + +instance strongEpi_map_of_isEquivalence [IsEquivalence F] (f : A ⟶ B) [h : StrongEpi f] : + StrongEpi (F.map f) := + F.asEquivalence.toAdjunction.strongEpi_map_of_strongEpi f +#align category_theory.adjunction.strong_epi_map_of_is_equivalence CategoryTheory.Adjunction.strongEpi_map_of_isEquivalence + +end CategoryTheory.Adjunction + +namespace CategoryTheory.Functor + +variable {C D : Type _} [Category C] [Category D] {F : C ⥤ D} {A B : C} (f : A ⟶ B) + +@[simp] +theorem strongEpi_map_iff_strongEpi_of_isEquivalence [IsEquivalence F] : + StrongEpi (F.map f) ↔ StrongEpi f := by + constructor + · intro + have e : arrow.mk f ≅ arrow.mk (F.inv.map (F.map f)) := + arrow.iso_of_nat_iso F.as_equivalence.unit_iso (arrow.mk f) + rw [strong_epi.iff_of_arrow_iso e] + infer_instance + · intro + infer_instance +#align category_theory.functor.strong_epi_map_iff_strong_epi_of_is_equivalence CategoryTheory.Functor.strongEpi_map_iff_strongEpi_of_isEquivalence + +end CategoryTheory.Functor + From a8492d25b2be69fff94027f95d2a88610afafaaa Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:54:23 -0500 Subject: [PATCH 61/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Functor/EpiMono.lean | 12 +++++------- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..b7380aae2e4b5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -239,6 +239,7 @@ import Mathlib.CategoryTheory.Functor.Category import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.Functor.Currying import Mathlib.CategoryTheory.Functor.Default +import Mathlib.CategoryTheory.Functor.EpiMono import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial import Mathlib.CategoryTheory.Functor.Hom diff --git a/Mathlib/CategoryTheory/Functor/EpiMono.lean b/Mathlib/CategoryTheory/Functor/EpiMono.lean index 066061e6fefb5..aca32ce70636e 100644 --- a/Mathlib/CategoryTheory/Functor/EpiMono.lean +++ b/Mathlib/CategoryTheory/Functor/EpiMono.lean @@ -8,9 +8,9 @@ Authors: Markus Himmel ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.EpiMono -import Mathbin.CategoryTheory.Limits.Shapes.StrongEpi -import Mathbin.CategoryTheory.LiftingProperties.Adjunction +import Mathlib.CategoryTheory.EpiMono +import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi +import Mathlib.CategoryTheory.LiftingProperties.Adjunction /-! # Preservation and reflection of monomorphisms and epimorphisms @@ -239,8 +239,7 @@ def splitEpiEquiv [Full F] [Faithful F] : SplitEpi f ≃ SplitEpi (F.map f) #align category_theory.functor.split_epi_equiv CategoryTheory.Functor.splitEpiEquiv @[simp] -theorem isSplitEpi_iff [Full F] [Faithful F] : IsSplitEpi (F.map f) ↔ IsSplitEpi f := - by +theorem isSplitEpi_iff [Full F] [Faithful F] : IsSplitEpi (F.map f) ↔ IsSplitEpi f := by constructor · intro h exact is_split_epi.mk' ((split_epi_equiv F f).invFun h.exists_split_epi.some) @@ -262,8 +261,7 @@ def splitMonoEquiv [Full F] [Faithful F] : SplitMono f ≃ SplitMono (F.map f) #align category_theory.functor.split_mono_equiv CategoryTheory.Functor.splitMonoEquiv @[simp] -theorem isSplitMono_iff [Full F] [Faithful F] : IsSplitMono (F.map f) ↔ IsSplitMono f := - by +theorem isSplitMono_iff [Full F] [Faithful F] : IsSplitMono (F.map f) ↔ IsSplitMono f := by constructor · intro h exact is_split_mono.mk' ((split_mono_equiv F f).invFun h.exists_split_mono.some) From 936a01130bd821e2f77c219d863452cfe2442e32 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 21:45:18 -0500 Subject: [PATCH 62/66] fix errors; lint --- Mathlib/CategoryTheory/Arrow.lean | 1 - Mathlib/CategoryTheory/Functor/EpiMono.lean | 173 +++++++++++--------- 2 files changed, 94 insertions(+), 80 deletions(-) diff --git a/Mathlib/CategoryTheory/Arrow.lean b/Mathlib/CategoryTheory/Arrow.lean index 475368b88c326..41f7ee32c56d9 100644 --- a/Mathlib/CategoryTheory/Arrow.lean +++ b/Mathlib/CategoryTheory/Arrow.lean @@ -334,4 +334,3 @@ def Arrow.isoOfNatIso {C D : Type _} [Category C] [Category D] {F G : C ⥤ D} ( #align category_theory.arrow.iso_of_nat_iso CategoryTheory.Arrow.isoOfNatIso end CategoryTheory -#lint diff --git a/Mathlib/CategoryTheory/Functor/EpiMono.lean b/Mathlib/CategoryTheory/Functor/EpiMono.lean index aca32ce70636e..3c0d2c474c510 100644 --- a/Mathlib/CategoryTheory/Functor/EpiMono.lean +++ b/Mathlib/CategoryTheory/Functor/EpiMono.lean @@ -31,6 +31,7 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] /-- A functor preserves monomorphisms if it maps monomorphisms to monomorphisms. -/ class PreservesMonomorphisms (F : C ⥤ D) : Prop where + /-- A functor preserves monomorphisms if it maps monomorphisms to monomorphisms. -/ preserves : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], Mono (F.map f) #align category_theory.functor.preserves_monomorphisms CategoryTheory.Functor.PreservesMonomorphisms @@ -41,6 +42,7 @@ instance map_mono (F : C ⥤ D) [PreservesMonomorphisms F] {X Y : C} (f : X ⟶ /-- A functor preserves epimorphisms if it maps epimorphisms to epimorphisms. -/ class PreservesEpimorphisms (F : C ⥤ D) : Prop where + /-- A functor preserves epimorphisms if it maps epimorphisms to epimorphisms. -/ preserves : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], Epi (F.map f) #align category_theory.functor.preserves_epimorphisms CategoryTheory.Functor.PreservesEpimorphisms @@ -52,6 +54,8 @@ instance map_epi (F : C ⥤ D) [PreservesEpimorphisms F] {X Y : C} (f : X ⟶ Y) /-- A functor reflects monomorphisms if morphisms that are mapped to monomorphisms are themselves monomorphisms. -/ class ReflectsMonomorphisms (F : C ⥤ D) : Prop where + /-- A functor reflects monomorphisms if morphisms that are mapped to monomorphisms are themselves + monomorphisms. -/ reflects : ∀ {X Y : C} (f : X ⟶ Y), Mono (F.map f) → Mono f #align category_theory.functor.reflects_monomorphisms CategoryTheory.Functor.ReflectsMonomorphisms @@ -63,6 +67,8 @@ theorem mono_of_mono_map (F : C ⥤ D) [ReflectsMonomorphisms F] {X Y : C} {f : /-- A functor reflects epimorphisms if morphisms that are mapped to epimorphisms are themselves epimorphisms. -/ class ReflectsEpimorphisms (F : C ⥤ D) : Prop where + /-- A functor reflects epimorphisms if morphisms that are mapped to epimorphisms are themselves + epimorphisms. -/ reflects : ∀ {X Y : C} (f : X ⟶ Y), Epi (F.map f) → Epi f #align category_theory.functor.reflects_epimorphisms CategoryTheory.Functor.ReflectsEpimorphisms @@ -72,133 +78,133 @@ theorem epi_of_epi_map (F : C ⥤ D) [ReflectsEpimorphisms F] {X Y : C} {f : X #align category_theory.functor.epi_of_epi_map CategoryTheory.Functor.epi_of_epi_map instance preservesMonomorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [PreservesMonomorphisms F] - [PreservesMonomorphisms G] : PreservesMonomorphisms (F ⋙ G) - where preserves X Y f h := by + [PreservesMonomorphisms G] : PreservesMonomorphisms (F ⋙ G) where + preserves f h := by rw [comp_map] exact inferInstance #align category_theory.functor.preserves_monomorphisms_comp CategoryTheory.Functor.preservesMonomorphisms_comp instance preservesEpimorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [PreservesEpimorphisms F] - [PreservesEpimorphisms G] : PreservesEpimorphisms (F ⋙ G) - where preserves X Y f h := by + [PreservesEpimorphisms G] : PreservesEpimorphisms (F ⋙ G) where + preserves f h := by rw [comp_map] exact inferInstance #align category_theory.functor.preserves_epimorphisms_comp CategoryTheory.Functor.preservesEpimorphisms_comp instance reflectsMonomorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [ReflectsMonomorphisms F] - [ReflectsMonomorphisms G] : ReflectsMonomorphisms (F ⋙ G) - where reflects X Y f h := F.mono_of_mono_map (G.mono_of_mono_map h) + [ReflectsMonomorphisms G] : ReflectsMonomorphisms (F ⋙ G) where + reflects _ h := F.mono_of_mono_map (G.mono_of_mono_map h) #align category_theory.functor.reflects_monomorphisms_comp CategoryTheory.Functor.reflectsMonomorphisms_comp instance reflectsEpimorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [ReflectsEpimorphisms F] - [ReflectsEpimorphisms G] : ReflectsEpimorphisms (F ⋙ G) - where reflects X Y f h := F.epi_of_epi_map (G.epi_of_epi_map h) + [ReflectsEpimorphisms G] : ReflectsEpimorphisms (F ⋙ G) where + reflects _ h := F.epi_of_epi_map (G.epi_of_epi_map h) #align category_theory.functor.reflects_epimorphisms_comp CategoryTheory.Functor.reflectsEpimorphisms_comp theorem preservesEpimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) [PreservesEpimorphisms (F ⋙ G)] [ReflectsEpimorphisms G] : PreservesEpimorphisms F := - ⟨fun X Y f hf => G.epi_of_epi_map <| show Epi ((F ⋙ G).map f) from inferInstance⟩ + ⟨fun f _ => G.epi_of_epi_map <| show Epi ((F ⋙ G).map f) from inferInstance⟩ #align category_theory.functor.preserves_epimorphisms_of_preserves_of_reflects CategoryTheory.Functor.preservesEpimorphisms_of_preserves_of_reflects theorem preservesMonomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) [PreservesMonomorphisms (F ⋙ G)] [ReflectsMonomorphisms G] : PreservesMonomorphisms F := - ⟨fun X Y f hf => G.mono_of_mono_map <| show Mono ((F ⋙ G).map f) from inferInstance⟩ + ⟨fun f _ => G.mono_of_mono_map <| show Mono ((F ⋙ G).map f) from inferInstance⟩ #align category_theory.functor.preserves_monomorphisms_of_preserves_of_reflects CategoryTheory.Functor.preservesMonomorphisms_of_preserves_of_reflects theorem reflectsEpimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) [PreservesEpimorphisms G] [ReflectsEpimorphisms (F ⋙ G)] : ReflectsEpimorphisms F := - ⟨fun X Y f hf => (F ⋙ G).epi_of_epi_map <| show Epi (G.map (F.map f)) from inferInstance⟩ + ⟨fun f _ => (F ⋙ G).epi_of_epi_map <| show Epi (G.map (F.map f)) from inferInstance⟩ #align category_theory.functor.reflects_epimorphisms_of_preserves_of_reflects CategoryTheory.Functor.reflectsEpimorphisms_of_preserves_of_reflects theorem reflectsMonomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) [PreservesMonomorphisms G] [ReflectsMonomorphisms (F ⋙ G)] : ReflectsMonomorphisms F := - ⟨fun X Y f hf => (F ⋙ G).mono_of_mono_map <| show Mono (G.map (F.map f)) from inferInstance⟩ + ⟨fun f _ => (F ⋙ G).mono_of_mono_map <| show Mono (G.map (F.map f)) from inferInstance⟩ #align category_theory.functor.reflects_monomorphisms_of_preserves_of_reflects CategoryTheory.Functor.reflectsMonomorphisms_of_preserves_of_reflects -theorem PreservesMonomorphisms.of_iso {F G : C ⥤ D} [PreservesMonomorphisms F] (α : F ≅ G) : +theorem preservesMonomorphisms.of_iso {F G : C ⥤ D} [PreservesMonomorphisms F] (α : F ≅ G) : PreservesMonomorphisms G := { - preserves := fun X Y f h => + preserves := fun {X} {Y} f h => by - haveI : mono (F.map f ≫ (α.app Y).Hom) := mono_comp _ _ - convert (mono_comp _ _ : mono ((α.app X).inv ≫ F.map f ≫ (α.app Y).Hom)) - rw [iso.eq_inv_comp, iso.app_hom, iso.app_hom, nat_trans.naturality] } -#align category_theory.functor.preserves_monomorphisms.of_iso CategoryTheory.Functor.PreservesMonomorphisms.of_iso + haveI : Mono (F.map f ≫ (α.app Y).hom) := mono_comp _ _ + convert (mono_comp _ _ : Mono ((α.app X).inv ≫ F.map f ≫ (α.app Y).hom)) + rw [Iso.eq_inv_comp, Iso.app_hom, Iso.app_hom, NatTrans.naturality] } +#align category_theory.functor.preserves_monomorphisms.of_iso CategoryTheory.Functor.preservesMonomorphisms.of_iso -theorem PreservesMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : +theorem preservesMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : PreservesMonomorphisms F ↔ PreservesMonomorphisms G := - ⟨fun h => preserves_monomorphisms.of_iso α, fun h => preserves_monomorphisms.of_iso α.symm⟩ -#align category_theory.functor.preserves_monomorphisms.iso_iff CategoryTheory.Functor.PreservesMonomorphisms.iso_iff + ⟨fun _ => preservesMonomorphisms.of_iso α, fun _ => preservesMonomorphisms.of_iso α.symm⟩ +#align category_theory.functor.preserves_monomorphisms.iso_iff CategoryTheory.Functor.preservesMonomorphisms.iso_iff -theorem PreservesEpimorphisms.of_iso {F G : C ⥤ D} [PreservesEpimorphisms F] (α : F ≅ G) : +theorem preservesEpimorphisms.of_iso {F G : C ⥤ D} [PreservesEpimorphisms F] (α : F ≅ G) : PreservesEpimorphisms G := { - preserves := fun X Y f h => + preserves := fun {X} {Y} f h => by - haveI : epi (F.map f ≫ (α.app Y).Hom) := epi_comp _ _ - convert (epi_comp _ _ : epi ((α.app X).inv ≫ F.map f ≫ (α.app Y).Hom)) - rw [iso.eq_inv_comp, iso.app_hom, iso.app_hom, nat_trans.naturality] } -#align category_theory.functor.preserves_epimorphisms.of_iso CategoryTheory.Functor.PreservesEpimorphisms.of_iso + haveI : Epi (F.map f ≫ (α.app Y).hom) := epi_comp _ _ + convert (epi_comp _ _ : Epi ((α.app X).inv ≫ F.map f ≫ (α.app Y).hom)) + rw [Iso.eq_inv_comp, Iso.app_hom, Iso.app_hom, NatTrans.naturality] } +#align category_theory.functor.preserves_epimorphisms.of_iso CategoryTheory.Functor.preservesEpimorphisms.of_iso -theorem PreservesEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : +theorem preservesEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : PreservesEpimorphisms F ↔ PreservesEpimorphisms G := - ⟨fun h => preserves_epimorphisms.of_iso α, fun h => preserves_epimorphisms.of_iso α.symm⟩ -#align category_theory.functor.preserves_epimorphisms.iso_iff CategoryTheory.Functor.PreservesEpimorphisms.iso_iff + ⟨fun _ => preservesEpimorphisms.of_iso α, fun _ => preservesEpimorphisms.of_iso α.symm⟩ +#align category_theory.functor.preserves_epimorphisms.iso_iff CategoryTheory.Functor.preservesEpimorphisms.iso_iff -theorem ReflectsMonomorphisms.of_iso {F G : C ⥤ D} [ReflectsMonomorphisms F] (α : F ≅ G) : +theorem reflectsMonomorphisms.of_iso {F G : C ⥤ D} [ReflectsMonomorphisms F] (α : F ≅ G) : ReflectsMonomorphisms G := { - reflects := fun X Y f h => by + reflects := fun {X} {Y} f h => by apply F.mono_of_mono_map - haveI : mono (G.map f ≫ (α.app Y).inv) := mono_comp _ _ - convert (mono_comp _ _ : mono ((α.app X).Hom ≫ G.map f ≫ (α.app Y).inv)) - rw [← category.assoc, iso.eq_comp_inv, iso.app_hom, iso.app_hom, nat_trans.naturality] } -#align category_theory.functor.reflects_monomorphisms.of_iso CategoryTheory.Functor.ReflectsMonomorphisms.of_iso + haveI : Mono (G.map f ≫ (α.app Y).inv) := mono_comp _ _ + convert (mono_comp _ _ : Mono ((α.app X).hom ≫ G.map f ≫ (α.app Y).inv)) + rw [← Category.assoc, Iso.eq_comp_inv, Iso.app_hom, Iso.app_hom, NatTrans.naturality] } +#align category_theory.functor.reflects_monomorphisms.of_iso CategoryTheory.Functor.reflectsMonomorphisms.of_iso -theorem ReflectsMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : +theorem reflectsMonomorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : ReflectsMonomorphisms F ↔ ReflectsMonomorphisms G := - ⟨fun h => reflects_monomorphisms.of_iso α, fun h => reflects_monomorphisms.of_iso α.symm⟩ -#align category_theory.functor.reflects_monomorphisms.iso_iff CategoryTheory.Functor.ReflectsMonomorphisms.iso_iff + ⟨fun _ => reflectsMonomorphisms.of_iso α, fun _ => reflectsMonomorphisms.of_iso α.symm⟩ +#align category_theory.functor.reflects_monomorphisms.iso_iff CategoryTheory.Functor.reflectsMonomorphisms.iso_iff -theorem ReflectsEpimorphisms.of_iso {F G : C ⥤ D} [ReflectsEpimorphisms F] (α : F ≅ G) : +theorem reflectsEpimorphisms.of_iso {F G : C ⥤ D} [ReflectsEpimorphisms F] (α : F ≅ G) : ReflectsEpimorphisms G := { - reflects := fun X Y f h => by + reflects := fun {X} {Y} f h => by apply F.epi_of_epi_map - haveI : epi (G.map f ≫ (α.app Y).inv) := epi_comp _ _ - convert (epi_comp _ _ : epi ((α.app X).Hom ≫ G.map f ≫ (α.app Y).inv)) - rw [← category.assoc, iso.eq_comp_inv, iso.app_hom, iso.app_hom, nat_trans.naturality] } -#align category_theory.functor.reflects_epimorphisms.of_iso CategoryTheory.Functor.ReflectsEpimorphisms.of_iso + haveI : Epi (G.map f ≫ (α.app Y).inv) := epi_comp _ _ + convert (epi_comp _ _ : Epi ((α.app X).hom ≫ G.map f ≫ (α.app Y).inv)) + rw [← Category.assoc, Iso.eq_comp_inv, Iso.app_hom, Iso.app_hom, NatTrans.naturality] } +#align category_theory.functor.reflects_epimorphisms.of_iso CategoryTheory.Functor.reflectsEpimorphisms.of_iso -theorem ReflectsEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : +theorem reflectsEpimorphisms.iso_iff {F G : C ⥤ D} (α : F ≅ G) : ReflectsEpimorphisms F ↔ ReflectsEpimorphisms G := - ⟨fun h => reflects_epimorphisms.of_iso α, fun h => reflects_epimorphisms.of_iso α.symm⟩ -#align category_theory.functor.reflects_epimorphisms.iso_iff CategoryTheory.Functor.ReflectsEpimorphisms.iso_iff + ⟨fun _ => reflectsEpimorphisms.of_iso α, fun _ => reflectsEpimorphisms.of_iso α.symm⟩ +#align category_theory.functor.reflects_epimorphisms.iso_iff CategoryTheory.Functor.reflectsEpimorphisms.iso_iff -theorem preserves_epimorphsisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : +theorem preservesEpimorphsisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : PreservesEpimorphisms F := { - preserves := fun X Y f hf => + preserves := fun {X} {Y} f hf => ⟨by intro Z g h H - replace H := congr_arg (adj.hom_equiv X Z) H - rwa [adj.hom_equiv_naturality_left, adj.hom_equiv_naturality_left, cancel_epi, + replace H := congr_arg (adj.homEquiv X Z) H + rwa [adj.homEquiv_naturality_left, adj.homEquiv_naturality_left, cancel_epi, Equiv.apply_eq_iff_eq] at H⟩ } -#align category_theory.functor.preserves_epimorphsisms_of_adjunction CategoryTheory.Functor.preserves_epimorphsisms_of_adjunction +#align category_theory.functor.preserves_epimorphsisms_of_adjunction CategoryTheory.Functor.preservesEpimorphsisms_of_adjunction instance (priority := 100) preservesEpimorphisms_of_isLeftAdjoint (F : C ⥤ D) [IsLeftAdjoint F] : PreservesEpimorphisms F := - preserves_epimorphsisms_of_adjunction (Adjunction.ofLeftAdjoint F) + preservesEpimorphsisms_of_adjunction (Adjunction.ofLeftAdjoint F) #align category_theory.functor.preserves_epimorphisms_of_is_left_adjoint CategoryTheory.Functor.preservesEpimorphisms_of_isLeftAdjoint theorem preservesMonomorphisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : PreservesMonomorphisms G := { - preserves := fun X Y f hf => + preserves := fun {X} {Y} f hf => ⟨by intro Z g h H - replace H := congr_arg (adj.hom_equiv Z Y).symm H - rwa [adj.hom_equiv_naturality_right_symm, adj.hom_equiv_naturality_right_symm, cancel_mono, + replace H := congr_arg (adj.homEquiv Z Y).symm H + rwa [adj.homEquiv_naturality_right_symm, adj.homEquiv_naturality_right_symm, cancel_mono, Equiv.apply_eq_iff_eq] at H⟩ } #align category_theory.functor.preserves_monomorphisms_of_adjunction CategoryTheory.Functor.preservesMonomorphisms_of_adjunction @@ -209,15 +215,15 @@ instance (priority := 100) preservesMonomorphisms_of_isRightAdjoint (F : C ⥤ D instance (priority := 100) reflectsMonomorphisms_of_faithful (F : C ⥤ D) [Faithful F] : ReflectsMonomorphisms F - where reflects X Y f hf := - ⟨fun Z g h hgh => + where reflects {X} {Y} f hf := + ⟨fun {Z} g h hgh => F.map_injective ((cancel_mono (F.map f)).1 (by rw [← F.map_comp, hgh, F.map_comp]))⟩ #align category_theory.functor.reflects_monomorphisms_of_faithful CategoryTheory.Functor.reflectsMonomorphisms_of_faithful instance (priority := 100) reflectsEpimorphisms_of_faithful (F : C ⥤ D) [Faithful F] : ReflectsEpimorphisms F - where reflects X Y f hf := - ⟨fun Z g h hgh => + where reflects {X} {Y} f hf := + ⟨fun {Z} g h hgh => F.map_injective ((cancel_epi (F.map f)).1 (by rw [← F.map_comp, hgh, F.map_comp]))⟩ #align category_theory.functor.reflects_epimorphisms_of_faithful CategoryTheory.Functor.reflectsEpimorphisms_of_faithful @@ -233,18 +239,22 @@ def splitEpiEquiv [Full F] [Faithful F] : SplitEpi f ≃ SplitEpi (F.map f) refine' ⟨F.preimage s.section_, _⟩ apply F.map_injective simp only [map_comp, image_preimage, map_id] - apply split_epi.id - left_inv := by tidy - right_inv := by tidy + apply SplitEpi.id + left_inv := by aesop_cat + right_inv := by + simp only [Function.RightInverse,Function.LeftInverse] + intro x + simp only [SplitEpi.map, preimage] + aesop_cat #align category_theory.functor.split_epi_equiv CategoryTheory.Functor.splitEpiEquiv @[simp] theorem isSplitEpi_iff [Full F] [Faithful F] : IsSplitEpi (F.map f) ↔ IsSplitEpi f := by constructor · intro h - exact is_split_epi.mk' ((split_epi_equiv F f).invFun h.exists_split_epi.some) + exact IsSplitEpi.mk' ((splitEpiEquiv F f).invFun h.exists_splitEpi.some) · intro h - exact is_split_epi.mk' ((split_epi_equiv F f).toFun h.exists_split_epi.some) + exact IsSplitEpi.mk' ((splitEpiEquiv F f).toFun h.exists_splitEpi.some) #align category_theory.functor.is_split_epi_iff CategoryTheory.Functor.isSplitEpi_iff /-- If `F` is a fully faithful functor, split monomorphisms are preserved and reflected by `F`. -/ @@ -255,18 +265,23 @@ def splitMonoEquiv [Full F] [Faithful F] : SplitMono f ≃ SplitMono (F.map f) refine' ⟨F.preimage s.retraction, _⟩ apply F.map_injective simp only [map_comp, image_preimage, map_id] - apply split_mono.id - left_inv := by tidy - right_inv := by tidy + apply SplitMono.id + left_inv := by aesop_cat + right_inv := by + simp only [Function.RightInverse, Function.LeftInverse] + intro x + simp only [SplitMono.map,preimage] + aesop_cat + #align category_theory.functor.split_mono_equiv CategoryTheory.Functor.splitMonoEquiv @[simp] theorem isSplitMono_iff [Full F] [Faithful F] : IsSplitMono (F.map f) ↔ IsSplitMono f := by constructor · intro h - exact is_split_mono.mk' ((split_mono_equiv F f).invFun h.exists_split_mono.some) + exact IsSplitMono.mk' ((splitMonoEquiv F f).invFun h.exists_splitMono.some) · intro h - exact is_split_mono.mk' ((split_mono_equiv F f).toFun h.exists_split_mono.some) + exact IsSplitMono.mk' ((splitMonoEquiv F f).toFun h.exists_splitMono.some) #align category_theory.functor.is_split_mono_iff CategoryTheory.Functor.isSplitMono_iff @[simp] @@ -291,10 +306,10 @@ theorem mono_map_iff_mono [hF₁ : PreservesMonomorphisms F] [hF₂ : ReflectsMo then `D` also is. -/ def splitEpiCategoryImpOfIsEquivalence [IsEquivalence F] [SplitEpiCategory C] : SplitEpiCategory D := - ⟨fun X Y f => by + ⟨fun {X} {Y} f => by intro - rw [← F.inv.is_split_epi_iff f] - apply is_split_epi_of_epi⟩ + rw [← F.inv.isSplitEpi_iff f] + apply isSplitEpi_of_epi⟩ #align category_theory.functor.split_epi_category_imp_of_is_equivalence CategoryTheory.Functor.splitEpiCategoryImpOfIsEquivalence end @@ -309,11 +324,11 @@ theorem strongEpi_map_of_strongEpi (adj : F ⊣ F') (f : A ⟶ B) [h₁ : F'.Pre [h₂ : F.PreservesEpimorphisms] [StrongEpi f] : StrongEpi (F.map f) := ⟨inferInstance, fun X Y Z => by intro - rw [adj.has_lifting_property_iff] + rw [adj.hasLiftingProperty_iff] infer_instance⟩ #align category_theory.adjunction.strong_epi_map_of_strong_epi CategoryTheory.Adjunction.strongEpi_map_of_strongEpi -instance strongEpi_map_of_isEquivalence [IsEquivalence F] (f : A ⟶ B) [h : StrongEpi f] : +instance strongEpi_map_of_isEquivalence [IsEquivalence F] (f : A ⟶ B) [_h : StrongEpi f] : StrongEpi (F.map f) := F.asEquivalence.toAdjunction.strongEpi_map_of_strongEpi f #align category_theory.adjunction.strong_epi_map_of_is_equivalence CategoryTheory.Adjunction.strongEpi_map_of_isEquivalence @@ -329,9 +344,9 @@ theorem strongEpi_map_iff_strongEpi_of_isEquivalence [IsEquivalence F] : StrongEpi (F.map f) ↔ StrongEpi f := by constructor · intro - have e : arrow.mk f ≅ arrow.mk (F.inv.map (F.map f)) := - arrow.iso_of_nat_iso F.as_equivalence.unit_iso (arrow.mk f) - rw [strong_epi.iff_of_arrow_iso e] + have e : Arrow.mk f ≅ Arrow.mk (F.inv.map (F.map f)) := + Arrow.isoOfNatIso F.asEquivalence.unitIso (Arrow.mk f) + rw [StrongEpi.iff_of_arrow_iso e] infer_instance · intro infer_instance From e91cd26b896609fca086d8063f34df1e7c0bc8b3 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 21:51:28 -0500 Subject: [PATCH 63/66] feat: port CategoryTheory.Functor.ReflectsIsomorphisms From 1fe7733c0208090554e299b8064bd0c91dbe4d6c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 21:51:28 -0500 Subject: [PATCH 64/66] Initial file copy from mathport --- .../Functor/ReflectsIsomorphisms.lean | 80 +++++++++++++++++++ 1 file changed, 80 insertions(+) create mode 100644 Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean diff --git a/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean b/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean new file mode 100644 index 0000000000000..9f54a5c8e8e9e --- /dev/null +++ b/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2020 Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bhavik Mehta + +! This file was ported from Lean 3 source module category_theory.functor.reflects_isomorphisms +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Balanced +import Mathbin.CategoryTheory.Functor.EpiMono +import Mathbin.CategoryTheory.Functor.FullyFaithful + +/-! +# Functors which reflect isomorphisms + +A functor `F` reflects isomorphisms if whenever `F.map f` is an isomorphism, `f` was too. + +It is formalized as a `Prop` valued typeclass `reflects_isomorphisms F`. + +Any fully faithful functor reflects isomorphisms. +-/ + + +open CategoryTheory CategoryTheory.Functor + +namespace CategoryTheory + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +variable {C : Type u₁} [Category.{v₁} C] + +section ReflectsIso + +variable {D : Type u₂} [Category.{v₂} D] + +variable {E : Type u₃} [Category.{v₃} E] + +/-- Define what it means for a functor `F : C ⥤ D` to reflect isomorphisms: for any +morphism `f : A ⟶ B`, if `F.map f` is an isomorphism then `f` is as well. +Note that we do not assume or require that `F` is faithful. +-/ +class ReflectsIsomorphisms (F : C ⥤ D) : Prop where + reflects : ∀ {A B : C} (f : A ⟶ B) [IsIso (F.map f)], IsIso f +#align category_theory.reflects_isomorphisms CategoryTheory.ReflectsIsomorphisms + +/-- If `F` reflects isos and `F.map f` is an iso, then `f` is an iso. -/ +theorem isIso_of_reflects_iso {A B : C} (f : A ⟶ B) (F : C ⥤ D) [IsIso (F.map f)] + [ReflectsIsomorphisms F] : IsIso f := + ReflectsIsomorphisms.reflects F f +#align category_theory.is_iso_of_reflects_iso CategoryTheory.isIso_of_reflects_iso + +instance (priority := 100) of_full_and_faithful (F : C ⥤ D) [Full F] [Faithful F] : + ReflectsIsomorphisms F + where reflects X Y f i := + ⟨⟨F.preimage (inv (F.map f)), ⟨F.map_injective (by simp), F.map_injective (by simp)⟩⟩⟩ +#align category_theory.of_full_and_faithful CategoryTheory.of_full_and_faithful + +instance (F : C ⥤ D) (G : D ⥤ E) [ReflectsIsomorphisms F] [ReflectsIsomorphisms G] : + ReflectsIsomorphisms (F ⋙ G) := + ⟨fun _ _ f (hf : IsIso (G.map _)) => by + skip + haveI := is_iso_of_reflects_iso (F.map f) G + exact is_iso_of_reflects_iso f F⟩ + +instance (priority := 100) reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms + [Balanced C] (F : C ⥤ D) [ReflectsMonomorphisms F] [ReflectsEpimorphisms F] : + ReflectsIsomorphisms F + where reflects A B f hf := by + skip + haveI : epi f := epi_of_epi_map F inferInstance + haveI : mono f := mono_of_mono_map F inferInstance + exact is_iso_of_mono_of_epi f +#align category_theory.reflects_isomorphisms_of_reflects_monomorphisms_of_reflects_epimorphisms CategoryTheory.reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms + +end ReflectsIso + +end CategoryTheory + From f10353648b41fda6604f9f113a28f0cf4a98d9fa Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 21:51:28 -0500 Subject: [PATCH 65/66] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean | 6 +++--- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..308953915d8bd 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -243,6 +243,7 @@ import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos +import Mathlib.CategoryTheory.Functor.ReflectsIsomorphisms import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.NatIso diff --git a/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean b/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean index 9f54a5c8e8e9e..8197884ac4029 100644 --- a/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean +++ b/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean @@ -8,9 +8,9 @@ Authors: Bhavik Mehta ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Balanced -import Mathbin.CategoryTheory.Functor.EpiMono -import Mathbin.CategoryTheory.Functor.FullyFaithful +import Mathlib.CategoryTheory.Balanced +import Mathlib.CategoryTheory.Functor.EpiMono +import Mathlib.CategoryTheory.Functor.FullyFaithful /-! # Functors which reflect isomorphisms From 7f41cf3e83e6504ef6a51c98a584bfeddb7e8e6d Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 21:56:57 -0500 Subject: [PATCH 66/66] fix errors; lint; shorten filename --- Mathlib.lean | 2 +- ...ectsIsomorphisms.lean => ReflectsIso.lean} | 21 ++++++++++--------- 2 files changed, 12 insertions(+), 11 deletions(-) rename Mathlib/CategoryTheory/Functor/{ReflectsIsomorphisms.lean => ReflectsIso.lean} (84%) diff --git a/Mathlib.lean b/Mathlib.lean index 1c583a7b5c193..59319f46b5e60 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -246,7 +246,7 @@ import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos -import Mathlib.CategoryTheory.Functor.ReflectsIsomorphisms +import Mathlib.CategoryTheory.Functor.ReflectsIso import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.LiftingProperties.Adjunction diff --git a/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean b/Mathlib/CategoryTheory/Functor/ReflectsIso.lean similarity index 84% rename from Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean rename to Mathlib/CategoryTheory/Functor/ReflectsIso.lean index 8197884ac4029..34192a02d8c1d 100644 --- a/Mathlib/CategoryTheory/Functor/ReflectsIsomorphisms.lean +++ b/Mathlib/CategoryTheory/Functor/ReflectsIso.lean @@ -17,7 +17,7 @@ import Mathlib.CategoryTheory.Functor.FullyFaithful A functor `F` reflects isomorphisms if whenever `F.map f` is an isomorphism, `f` was too. -It is formalized as a `Prop` valued typeclass `reflects_isomorphisms F`. +It is formalized as a `Prop` valued typeclass `ReflectsIsomorphisms F`. Any fully faithful functor reflects isomorphisms. -/ @@ -42,6 +42,7 @@ morphism `f : A ⟶ B`, if `F.map f` is an isomorphism then `f` is as well. Note that we do not assume or require that `F` is faithful. -/ class ReflectsIsomorphisms (F : C ⥤ D) : Prop where + /-- For any `f`, if `F.map f` is an iso, then so was `f`-/ reflects : ∀ {A B : C} (f : A ⟶ B) [IsIso (F.map f)], IsIso f #align category_theory.reflects_isomorphisms CategoryTheory.ReflectsIsomorphisms @@ -53,25 +54,25 @@ theorem isIso_of_reflects_iso {A B : C} (f : A ⟶ B) (F : C ⥤ D) [IsIso (F.ma instance (priority := 100) of_full_and_faithful (F : C ⥤ D) [Full F] [Faithful F] : ReflectsIsomorphisms F - where reflects X Y f i := + where reflects f i := ⟨⟨F.preimage (inv (F.map f)), ⟨F.map_injective (by simp), F.map_injective (by simp)⟩⟩⟩ #align category_theory.of_full_and_faithful CategoryTheory.of_full_and_faithful instance (F : C ⥤ D) (G : D ⥤ E) [ReflectsIsomorphisms F] [ReflectsIsomorphisms G] : ReflectsIsomorphisms (F ⋙ G) := - ⟨fun _ _ f (hf : IsIso (G.map _)) => by + ⟨fun f (hf : IsIso (G.map _)) => by skip - haveI := is_iso_of_reflects_iso (F.map f) G - exact is_iso_of_reflects_iso f F⟩ + haveI := isIso_of_reflects_iso (F.map f) G + exact isIso_of_reflects_iso f F⟩ instance (priority := 100) reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms [Balanced C] (F : C ⥤ D) [ReflectsMonomorphisms F] [ReflectsEpimorphisms F] : - ReflectsIsomorphisms F - where reflects A B f hf := by + ReflectsIsomorphisms F where + reflects f hf := by skip - haveI : epi f := epi_of_epi_map F inferInstance - haveI : mono f := mono_of_mono_map F inferInstance - exact is_iso_of_mono_of_epi f + haveI : Epi f := epi_of_epi_map F inferInstance + haveI : Mono f := mono_of_mono_map F inferInstance + exact isIso_of_mono_of_epi f #align category_theory.reflects_isomorphisms_of_reflects_monomorphisms_of_reflects_epimorphisms CategoryTheory.reflectsIsomorphisms_of_reflectsMonomorphisms_of_reflectsEpimorphisms end ReflectsIso