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 001/126] 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 002/126] 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 003/126] 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 004/126] 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 005/126] 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 006/126] 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 007/126] 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 008/126] 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 009/126] 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 010/126] 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 011/126] 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 012/126] 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 013/126] 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 014/126] 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 015/126] 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 016/126] 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 017/126] 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 018/126] 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 019/126] 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 020/126] 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 021/126] 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 022/126] 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 023/126] 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 024/126] 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 025/126] 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 026/126] 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 027/126] 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 028/126] 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 b3e68ca52a36b1639c2e24c8b03664f55658d05f Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Sat, 11 Feb 2023 20:38:10 -0500 Subject: [PATCH 029/126] feat: port CategoryTheory.Pi.Basic From 8414826065094327362452132ae9b8020f89df88 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Sat, 11 Feb 2023 20:38:10 -0500 Subject: [PATCH 030/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Pi/Basic.lean | 263 +++++++++++++++++++++++++++ 1 file changed, 263 insertions(+) create mode 100644 Mathlib/CategoryTheory/Pi/Basic.lean diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean new file mode 100644 index 0000000000000..39f2de4b84654 --- /dev/null +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -0,0 +1,263 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Simon Hudon, Scott Morrison + +! This file was ported from Lean 3 source module category_theory.pi.basic +! leanprover-community/mathlib commit dc6c365e751e34d100e80fe6e314c3c3e0fd2988 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.NaturalIsomorphism +import Mathbin.CategoryTheory.EqToHom +import Mathbin.Data.Sum.Basic + +/-! +# Categories of indexed families of objects. + +We define the pointwise category structure on indexed families of objects in a category +(and also the dependent generalization). + +-/ + + +namespace CategoryTheory + +universe w₀ w₁ w₂ v₁ v₂ u₁ u₂ + +variable {I : Type w₀} (C : I → Type u₁) [∀ i, Category.{v₁} (C i)] + +/-- `pi C` gives the cartesian product of an indexed family of categories. +-/ +instance pi : Category.{max w₀ v₁} (∀ i, C i) + where + Hom X Y := ∀ i, X i ⟶ Y i + id X i := 𝟙 (X i) + comp X Y Z f g i := f i ≫ g i +#align category_theory.pi CategoryTheory.pi + +/-- This provides some assistance to typeclass search in a common situation, +which otherwise fails. (Without this `category_theory.pi.has_limit_of_has_limit_comp_eval` fails.) +-/ +abbrev pi' {I : Type v₁} (C : I → Type u₁) [∀ i, Category.{v₁} (C i)] : Category.{v₁} (∀ i, C i) := + CategoryTheory.pi C +#align category_theory.pi' CategoryTheory.pi' + +attribute [instance] pi' + +namespace Pi + +@[simp] +theorem id_apply (X : ∀ i, C i) (i) : (𝟙 X : ∀ i, X i ⟶ X i) i = 𝟙 (X i) := + rfl +#align category_theory.pi.id_apply CategoryTheory.pi.id_apply + +@[simp] +theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : + (f ≫ g : ∀ i, X i ⟶ Z i) i = f i ≫ g i := + rfl +#align category_theory.pi.comp_apply CategoryTheory.pi.comp_apply + +/-- +The evaluation functor at `i : I`, sending an `I`-indexed family of objects to the object over `i`. +-/ +@[simps] +def eval (i : I) : (∀ i, C i) ⥤ C i where + obj f := f i + map f g α := α i +#align category_theory.pi.eval CategoryTheory.pi.eval + +section + +variable {J : Type w₁} + +/-- Pull back an `I`-indexed family of objects to an `J`-indexed family, along a function `J → I`. +-/ +@[simps] +def comap (h : J → I) : (∀ i, C i) ⥤ ∀ j, C (h j) + where + obj f i := f (h i) + map f g α i := α (h i) +#align category_theory.pi.comap CategoryTheory.pi.comap + +variable (I) + +/-- The natural isomorphism between +pulling back a grading along the identity function, +and the identity functor. -/ +@[simps] +def comapId : comap C (id : I → I) ≅ 𝟭 (∀ i, C i) + where + Hom := { app := fun X => 𝟙 X } + inv := { app := fun X => 𝟙 X } +#align category_theory.pi.comap_id CategoryTheory.pi.comapId + +variable {I} + +variable {K : Type w₂} + +/-- The natural isomorphism comparing between +pulling back along two successive functions, and +pulling back along their composition +-/ +@[simps] +def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ comap C (g ∘ f) + where + Hom := { app := fun X b => 𝟙 (X (g (f b))) } + inv := { app := fun X b => 𝟙 (X (g (f b))) } +#align category_theory.pi.comap_comp CategoryTheory.pi.comapComp + +/-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ +@[simps] +def comapEvalIsoEval (h : J → I) (j : J) : comap C h ⋙ eval (C ∘ h) j ≅ eval C (h j) := + NatIso.ofComponents (fun f => Iso.refl _) (by tidy) +#align category_theory.pi.comap_eval_iso_eval CategoryTheory.pi.comapEvalIsoEval + +end + +section + +variable {J : Type w₀} {D : J → Type u₁} [∀ j, Category.{v₁} (D j)] + +/- warning: category_theory.pi.sum_elim_category -> CategoryTheory.pi.sumElimCategory is a dubious translation: +lean 3 declaration is + forall {I : Type.{u1}} (C : I -> Type.{u3}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u3} (C i)] {J : Type.{u1}} {D : J -> Type.{u3}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u3} (D j)] (s : Sum.{u1, u1} I J), CategoryTheory.Category.{u2, u3} (Sum.elim.{u1, u1, succ (succ u3)} I J Type.{u3} C D s) +but is expected to have type + forall {I : Type.{u3}} (C : I -> Type.{u1}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u1} (C i)] {J : Type.{u3}} {D : J -> Type.{u1}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u1} (D j)] (s : Sum.{u3, u3} I J), CategoryTheory.Category.{u2, u1} (Sum.elim.{u3, u3, succ (succ u1)} I J Type.{u1} C D s) +Case conversion may be inaccurate. Consider using '#align category_theory.pi.sum_elim_category CategoryTheory.pi.sumElimCategoryₓ'. -/ +instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) + | Sum.inl i => by + dsimp + infer_instance + | Sum.inr j => by + dsimp + infer_instance +#align category_theory.pi.sum_elim_category CategoryTheory.pi.sumElimCategory + +/-- The bifunctor combining an `I`-indexed family of objects with a `J`-indexed family of objects +to obtain an `I ⊕ J`-indexed family of objects. +-/ +@[simps] +def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s + where + obj f := + { obj := fun g s => Sum.rec f g s + map := fun g g' α s => Sum.rec (fun i => 𝟙 (f i)) α s } + map f f' α := { app := fun g s => Sum.rec α (fun j => 𝟙 (g j)) s } +#align category_theory.pi.sum CategoryTheory.pi.sum + +end + +variable {C} + +/-- An isomorphism between `I`-indexed objects gives an isomorphism between each +pair of corresponding components. -/ +@[simps] +def isoApp {X Y : ∀ i, C i} (f : X ≅ Y) (i : I) : X i ≅ Y i := + ⟨f.Hom i, f.inv i, by + dsimp + rw [← comp_apply, iso.hom_inv_id, id_apply], + by + dsimp + rw [← comp_apply, iso.inv_hom_id, id_apply]⟩ +#align category_theory.pi.iso_app CategoryTheory.pi.isoApp + +@[simp] +theorem isoApp_refl (X : ∀ i, C i) (i : I) : isoApp (Iso.refl X) i = Iso.refl (X i) := + rfl +#align category_theory.pi.iso_app_refl CategoryTheory.pi.isoApp_refl + +@[simp] +theorem isoApp_symm {X Y : ∀ i, C i} (f : X ≅ Y) (i : I) : isoApp f.symm i = (isoApp f i).symm := + rfl +#align category_theory.pi.iso_app_symm CategoryTheory.pi.isoApp_symm + +@[simp] +theorem isoApp_trans {X Y Z : ∀ i, C i} (f : X ≅ Y) (g : Y ≅ Z) (i : I) : + isoApp (f ≪≫ g) i = isoApp f i ≪≫ isoApp g i := + rfl +#align category_theory.pi.iso_app_trans CategoryTheory.pi.isoApp_trans + +end Pi + +namespace Functor + +variable {C} + +variable {D : I → Type u₁} [∀ i, Category.{v₁} (D i)] {A : Type u₁} [Category.{u₁} A] + +/-- Assemble an `I`-indexed family of functors into a functor between the pi types. +-/ +@[simps] +def pi (F : ∀ i, C i ⥤ D i) : (∀ i, C i) ⥤ ∀ i, D i + where + obj f i := (F i).obj (f i) + map f g α i := (F i).map (α i) +#align category_theory.functor.pi CategoryTheory.Functor.pi + +/-- Similar to `pi`, but all functors come from the same category `A` +-/ +@[simps] +def pi' (f : ∀ i, A ⥤ C i) : A ⥤ ∀ i, C i + where + obj a i := (f i).obj a + map a₁ a₂ h i := (f i).map h +#align category_theory.functor.pi' CategoryTheory.Functor.pi' + +section EqToHom + +@[simp] +theorem eqToHom_proj {x x' : ∀ i, C i} (h : x = x') (i : I) : + (eqToHom h : x ⟶ x') i = eqToHom (Function.funext_iff.mp h i) := + by + subst h + rfl +#align category_theory.functor.eq_to_hom_proj CategoryTheory.Functor.eqToHom_proj + +end EqToHom + +-- One could add some natural isomorphisms showing +-- how `functor.pi` commutes with `pi.eval` and `pi.comap`. +@[simp] +theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ pi.eval C i = f i := + by + apply Functor.ext <;> intros + · simp; · rfl +#align category_theory.functor.pi'_eval CategoryTheory.Functor.pi'_eval + +/-- Two functors to a product category are equal iff they agree on every coordinate. -/ +theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ pi.eval C i = f' ⋙ pi.eval C i) : f = f' := + by + apply Functor.ext; swap + · intro X + ext i + specialize h i + have := congr_obj h X + simpa + · intro x y p + ext i + specialize h i + have := congr_hom h p + simpa +#align category_theory.functor.pi_ext CategoryTheory.Functor.pi_ext + +end Functor + +namespace NatTrans + +variable {C} + +variable {D : I → Type u₁} [∀ i, Category.{v₁} (D i)] + +variable {F G : ∀ i, C i ⥤ D i} + +/-- Assemble an `I`-indexed family of natural transformations into a single natural transformation. +-/ +@[simps] +def pi (α : ∀ i, F i ⟶ G i) : Functor.pi F ⟶ Functor.pi G where app f i := (α i).app (f i) +#align category_theory.nat_trans.pi CategoryTheory.NatTrans.pi + +end NatTrans + +end CategoryTheory + From a4ea90f81f619c7cba6abb7545982d69a5d7ada4 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Sat, 11 Feb 2023 20:38:11 -0500 Subject: [PATCH 031/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Pi/Basic.lean | 15 ++++++--------- 2 files changed, 7 insertions(+), 9 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 8e0b03cddc495..42760ada3d9db 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -236,6 +236,7 @@ import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites +import Mathlib.CategoryTheory.Pi.Basic import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.Thin import Mathlib.CategoryTheory.Whiskering diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 39f2de4b84654..7a9fdf385511c 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -8,9 +8,9 @@ Authors: Simon Hudon, Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.NaturalIsomorphism -import Mathbin.CategoryTheory.EqToHom -import Mathbin.Data.Sum.Basic +import Mathlib.CategoryTheory.NaturalIsomorphism +import Mathlib.CategoryTheory.EqToHom +import Mathlib.Data.Sum.Basic /-! # Categories of indexed families of objects. @@ -208,8 +208,7 @@ section EqToHom @[simp] theorem eqToHom_proj {x x' : ∀ i, C i} (h : x = x') (i : I) : - (eqToHom h : x ⟶ x') i = eqToHom (Function.funext_iff.mp h i) := - by + (eqToHom h : x ⟶ x') i = eqToHom (Function.funext_iff.mp h i) := by subst h rfl #align category_theory.functor.eq_to_hom_proj CategoryTheory.Functor.eqToHom_proj @@ -219,15 +218,13 @@ end EqToHom -- One could add some natural isomorphisms showing -- how `functor.pi` commutes with `pi.eval` and `pi.comap`. @[simp] -theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ pi.eval C i = f i := - by +theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ pi.eval C i = f i := by apply Functor.ext <;> intros · simp; · rfl #align category_theory.functor.pi'_eval CategoryTheory.Functor.pi'_eval /-- Two functors to a product category are equal iff they agree on every coordinate. -/ -theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ pi.eval C i = f' ⋙ pi.eval C i) : f = f' := - by +theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ pi.eval C i = f' ⋙ pi.eval C i) : f = f' := by apply Functor.ext; swap · intro X ext i From 813630737bce52dc18e8f0ca4ce2e52838d44b10 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Sat, 11 Feb 2023 22:25:59 -0500 Subject: [PATCH 032/126] fix some errors --- Mathlib/CategoryTheory/Pi/Basic.lean | 116 ++++++++++++++++----------- 1 file changed, 69 insertions(+), 47 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 7a9fdf385511c..8077e7fcf12f8 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -8,7 +8,7 @@ Authors: Simon Hudon, Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathlib.CategoryTheory.NaturalIsomorphism +import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.EqToHom import Mathlib.Data.Sum.Basic @@ -20,12 +20,12 @@ We define the pointwise category structure on indexed families of objects in a c -/ - namespace CategoryTheory universe w₀ w₁ w₂ v₁ v₂ u₁ u₂ -variable {I : Type w₀} (C : I → Type u₁) [∀ i, Category.{v₁} (C i)] +variable {I : Type w₀} {J : Type w₁} (C : I → Type u₁) [∀ i, Category.{v₁} (C i)] + /-- `pi C` gives the cartesian product of an indexed family of categories. -/ @@ -33,7 +33,7 @@ instance pi : Category.{max w₀ v₁} (∀ i, C i) where Hom X Y := ∀ i, X i ⟶ Y i id X i := 𝟙 (X i) - comp X Y Z f g i := f i ≫ g i + comp f g i := f i ≫ g i #align category_theory.pi CategoryTheory.pi /-- This provides some assistance to typeclass search in a common situation, @@ -50,47 +50,56 @@ namespace Pi @[simp] theorem id_apply (X : ∀ i, C i) (i) : (𝟙 X : ∀ i, X i ⟶ X i) i = 𝟙 (X i) := rfl -#align category_theory.pi.id_apply CategoryTheory.pi.id_apply +#align category_theory.pi.id_apply CategoryTheory.Pi.id_apply @[simp] theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : (f ≫ g : ∀ i, X i ⟶ Z i) i = f i ≫ g i := rfl -#align category_theory.pi.comp_apply CategoryTheory.pi.comp_apply +#align category_theory.pi.comp_apply CategoryTheory.Pi.comp_apply /-- The evaluation functor at `i : I`, sending an `I`-indexed family of objects to the object over `i`. -/ -@[simps] +@[simps!] def eval (i : I) : (∀ i, C i) ⥤ C i where obj f := f i - map f g α := α i -#align category_theory.pi.eval CategoryTheory.pi.eval + map α := α i +#align category_theory.pi.eval CategoryTheory.Pi.eval section variable {J : Type w₁} +/- Porting note: add this because Lean cannot see directly through the `∘` for +`Function.comp` -/ + +instance (f : J → I) : (j : J) → Category ((C ∘ f) j) := by + dsimp + infer_instance + /-- Pull back an `I`-indexed family of objects to an `J`-indexed family, along a function `J → I`. -/ -@[simps] -def comap (h : J → I) : (∀ i, C i) ⥤ ∀ j, C (h j) +@[simps!] +def comap (h : J → I) : (∀ i, C i) ⥤ (∀ j, C (h j)) where obj f i := f (h i) - map f g α i := α (h i) -#align category_theory.pi.comap CategoryTheory.pi.comap + map α i := α (h i) +#align category_theory.pi.comap CategoryTheory.Pi.comap variable (I) /-- The natural isomorphism between pulling back a grading along the identity function, and the identity functor. -/ -@[simps] +@[simps!] def comapId : comap C (id : I → I) ≅ 𝟭 (∀ i, C i) where - Hom := { app := fun X => 𝟙 X } + hom := { app := fun X => 𝟙 X, naturality := by simp only [comap]; aesop_cat} inv := { app := fun X => 𝟙 X } -#align category_theory.pi.comap_id CategoryTheory.pi.comapId +#align category_theory.pi.comap_id CategoryTheory.Pi.comapId + +example (g : J → I) : (j : J) → Category (C (g j)) := by infer_instance variable {I} @@ -100,18 +109,26 @@ variable {K : Type w₂} pulling back along two successive functions, and pulling back along their composition -/ -@[simps] +@[simps!] def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ comap C (g ∘ f) where - Hom := { app := fun X b => 𝟙 (X (g (f b))) } - inv := { app := fun X b => 𝟙 (X (g (f b))) } -#align category_theory.pi.comap_comp CategoryTheory.pi.comapComp + hom := { + app := fun X b => 𝟙 (X (g (f b))) + naturality := by intro X Y f'; simp [comap,Function.comp]; aesop_cat + } + inv := { + app := fun X b => 𝟙 (X (g (f b))) + naturality := by aesop_cat + } + hom_inv_id := by aesop_cat + inv_hom_id := by aesop_cat +#align category_theory.pi.comap_comp CategoryTheory.Pi.comapComp /-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ -@[simps] +@[simps!] def comapEvalIsoEval (h : J → I) (j : J) : comap C h ⋙ eval (C ∘ h) j ≅ eval C (h j) := - NatIso.ofComponents (fun f => Iso.refl _) (by tidy) -#align category_theory.pi.comap_eval_iso_eval CategoryTheory.pi.comapEvalIsoEval + NatIso.ofComponents (fun f => Iso.refl _) (by simp only [Iso.refl]; aesop_cat) +#align category_theory.pi.comap_eval_iso_eval CategoryTheory.Pi.comapEvalIsoEval end @@ -119,12 +136,12 @@ section variable {J : Type w₀} {D : J → Type u₁} [∀ j, Category.{v₁} (D j)] -/- warning: category_theory.pi.sum_elim_category -> CategoryTheory.pi.sumElimCategory is a dubious translation: +/- warning: category_theory.pi.sum_elim_category -> CategoryTheory.Pi.sumElimCategory is a dubious translation: lean 3 declaration is forall {I : Type.{u1}} (C : I -> Type.{u3}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u3} (C i)] {J : Type.{u1}} {D : J -> Type.{u3}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u3} (D j)] (s : Sum.{u1, u1} I J), CategoryTheory.Category.{u2, u3} (Sum.elim.{u1, u1, succ (succ u3)} I J Type.{u3} C D s) but is expected to have type forall {I : Type.{u3}} (C : I -> Type.{u1}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u1} (C i)] {J : Type.{u3}} {D : J -> Type.{u1}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u1} (D j)] (s : Sum.{u3, u3} I J), CategoryTheory.Category.{u2, u1} (Sum.elim.{u3, u3, succ (succ u1)} I J Type.{u1} C D s) -Case conversion may be inaccurate. Consider using '#align category_theory.pi.sum_elim_category CategoryTheory.pi.sumElimCategoryₓ'. -/ +Case conversion may be inaccurate. Consider using '#align category_theory.pi.sum_elim_category CategoryTheory.Pi.sumElimCategoryₓ'. -/ instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) | Sum.inl i => by dsimp @@ -132,19 +149,19 @@ instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) | Sum.inr j => by dsimp infer_instance -#align category_theory.pi.sum_elim_category CategoryTheory.pi.sumElimCategory +#align category_theory.pi.sum_elim_category CategoryTheory.Pi.sumElimCategory /-- The bifunctor combining an `I`-indexed family of objects with a `J`-indexed family of objects to obtain an `I ⊕ J`-indexed family of objects. -/ -@[simps] +@[simps!] def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s where obj f := { obj := fun g s => Sum.rec f g s - map := fun g g' α s => Sum.rec (fun i => 𝟙 (f i)) α s } + map := fun α s => Sum.rec (fun i => 𝟙 (f i)) α s } map f f' α := { app := fun g s => Sum.rec α (fun j => 𝟙 (g j)) s } -#align category_theory.pi.sum CategoryTheory.pi.sum +#align category_theory.pi.sum CategoryTheory.Pi.sum end @@ -154,29 +171,29 @@ variable {C} pair of corresponding components. -/ @[simps] def isoApp {X Y : ∀ i, C i} (f : X ≅ Y) (i : I) : X i ≅ Y i := - ⟨f.Hom i, f.inv i, by + ⟨f.hom i, f.inv i, by dsimp - rw [← comp_apply, iso.hom_inv_id, id_apply], + rw [← comp_apply, Iso.hom_inv_id, id_apply], by dsimp - rw [← comp_apply, iso.inv_hom_id, id_apply]⟩ -#align category_theory.pi.iso_app CategoryTheory.pi.isoApp + rw [← comp_apply, Iso.inv_hom_id, id_apply]⟩ +#align category_theory.pi.iso_app CategoryTheory.Pi.isoApp @[simp] theorem isoApp_refl (X : ∀ i, C i) (i : I) : isoApp (Iso.refl X) i = Iso.refl (X i) := rfl -#align category_theory.pi.iso_app_refl CategoryTheory.pi.isoApp_refl +#align category_theory.pi.iso_app_refl CategoryTheory.Pi.isoApp_refl @[simp] theorem isoApp_symm {X Y : ∀ i, C i} (f : X ≅ Y) (i : I) : isoApp f.symm i = (isoApp f i).symm := rfl -#align category_theory.pi.iso_app_symm CategoryTheory.pi.isoApp_symm +#align category_theory.pi.iso_app_symm CategoryTheory.Pi.isoApp_symm @[simp] theorem isoApp_trans {X Y Z : ∀ i, C i} (f : X ≅ Y) (g : Y ≅ Z) (i : I) : isoApp (f ≪≫ g) i = isoApp f i ≪≫ isoApp g i := rfl -#align category_theory.pi.iso_app_trans CategoryTheory.pi.isoApp_trans +#align category_theory.pi.iso_app_trans CategoryTheory.Pi.isoApp_trans end Pi @@ -188,20 +205,20 @@ variable {D : I → Type u₁} [∀ i, Category.{v₁} (D i)] {A : Type u₁} [C /-- Assemble an `I`-indexed family of functors into a functor between the pi types. -/ -@[simps] +@[simps!] def pi (F : ∀ i, C i ⥤ D i) : (∀ i, C i) ⥤ ∀ i, D i where obj f i := (F i).obj (f i) - map f g α i := (F i).map (α i) + map α i := (F i).map (α i) #align category_theory.functor.pi CategoryTheory.Functor.pi /-- Similar to `pi`, but all functors come from the same category `A` -/ -@[simps] +@[simps!] def pi' (f : ∀ i, A ⥤ C i) : A ⥤ ∀ i, C i where obj a i := (f i).obj a - map a₁ a₂ h i := (f i).map h + map h i := (f i).map h #align category_theory.functor.pi' CategoryTheory.Functor.pi' section EqToHom @@ -218,14 +235,17 @@ end EqToHom -- One could add some natural isomorphisms showing -- how `functor.pi` commutes with `pi.eval` and `pi.comap`. @[simp] -theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ pi.eval C i = f i := by - apply Functor.ext <;> intros - · simp; · rfl +theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ Pi.eval C i = f i := by + apply Functor.ext + intro _ _ _ + · simp + · intro _ + rfl #align category_theory.functor.pi'_eval CategoryTheory.Functor.pi'_eval /-- Two functors to a product category are equal iff they agree on every coordinate. -/ -theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ pi.eval C i = f' ⋙ pi.eval C i) : f = f' := by - apply Functor.ext; swap +theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ (Pi.eval C i) = f' ⋙ (Pi.eval C i)) : f = f' := by + apply Functor.ext; rotate_left · intro X ext i specialize h i @@ -250,8 +270,10 @@ variable {F G : ∀ i, C i ⥤ D i} /-- Assemble an `I`-indexed family of natural transformations into a single natural transformation. -/ -@[simps] -def pi (α : ∀ i, F i ⟶ G i) : Functor.pi F ⟶ Functor.pi G where app f i := (α i).app (f i) +@[simps!] +def pi (α : ∀ i, F i ⟶ G i) : Functor.pi F ⟶ Functor.pi G where + app f i := (α i).app (f i) + naturality := by intro X Y p; simp [Functor.pi]; aesop_cat #align category_theory.nat_trans.pi CategoryTheory.NatTrans.pi end NatTrans From 1f56a41b344cb16296a6c6fc4260aa26fdd0f809 Mon Sep 17 00:00:00 2001 From: mpenciak Date: Mon, 13 Feb 2023 01:14:36 -0500 Subject: [PATCH 033/126] 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 034/126] 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 035/126] 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 036/126] 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 037/126] 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 200151823064ee08abfe63067eaadad8cd4f4875 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 07:30:31 -0500 Subject: [PATCH 038/126] some more fixes --- Mathlib/CategoryTheory/Pi/Basic.lean | 26 +++++++++++++++----------- 1 file changed, 15 insertions(+), 11 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 8077e7fcf12f8..47c8bf4e382ae 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -61,7 +61,7 @@ theorem comp_apply {X Y Z : ∀ i, C i} (f : X ⟶ Y) (g : Y ⟶ Z) (i) : /-- The evaluation functor at `i : I`, sending an `I`-indexed family of objects to the object over `i`. -/ -@[simps!] +@[simps] def eval (i : I) : (∀ i, C i) ⥤ C i where obj f := f i map α := α i @@ -80,7 +80,7 @@ instance (f : J → I) : (j : J) → Category ((C ∘ f) j) := by /-- Pull back an `I`-indexed family of objects to an `J`-indexed family, along a function `J → I`. -/ -@[simps!] +@[simps] def comap (h : J → I) : (∀ i, C i) ⥤ (∀ j, C (h j)) where obj f i := f (h i) @@ -92,7 +92,7 @@ variable (I) /-- The natural isomorphism between pulling back a grading along the identity function, and the identity functor. -/ -@[simps!] +@[simps] def comapId : comap C (id : I → I) ≅ 𝟭 (∀ i, C i) where hom := { app := fun X => 𝟙 X, naturality := by simp only [comap]; aesop_cat} @@ -114,14 +114,14 @@ def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ where hom := { app := fun X b => 𝟙 (X (g (f b))) - naturality := by intro X Y f'; simp [comap,Function.comp]; aesop_cat + naturality := fun X Y f' => by simp only [comap,Function.comp]; funext; simp } inv := { app := fun X b => 𝟙 (X (g (f b))) - naturality := by aesop_cat + naturality := fun X Y f' => by simp only [comap,Function.comp]; funext; simp } - hom_inv_id := by aesop_cat - inv_hom_id := by aesop_cat + hom_inv_id := by simp only [comap]; aesop_cat; sorry + inv_hom_id := by simp only [comap]; sorry #align category_theory.pi.comap_comp CategoryTheory.Pi.comapComp /-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ @@ -157,10 +157,14 @@ to obtain an `I ⊕ J`-indexed family of objects. @[simps!] def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s where - obj f := - { obj := fun g s => Sum.rec f g s - map := fun α s => Sum.rec (fun i => 𝟙 (f i)) α s } - map f f' α := { app := fun g s => Sum.rec α (fun j => 𝟙 (g j)) s } + obj X := + { obj := fun Y s => Sum.rec X Y s + map := fun {Y} {Y'} f s => Sum.rec (fun i => 𝟙 (X i)) f s + map_id := fun Y => sorry + map_comp := sorry } + map {X} {X'} f := { app := fun Y s => Sum.rec f (fun j => 𝟙 (Y j)) s } + map_id := sorry + map_comp := sorry #align category_theory.pi.sum CategoryTheory.Pi.sum end From 23193580c24720bcbb06941b3ec1ceb2cdb3e3f8 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 10:07:44 -0500 Subject: [PATCH 039/126] more fixes --- Mathlib/CategoryTheory/Pi/Basic.lean | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 47c8bf4e382ae..100236a7456ec 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -120,7 +120,7 @@ def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ app := fun X b => 𝟙 (X (g (f b))) naturality := fun X Y f' => by simp only [comap,Function.comp]; funext; simp } - hom_inv_id := by simp only [comap]; aesop_cat; sorry + hom_inv_id := by simp only [comap]; aesop_cat; sorry inv_hom_id := by simp only [comap]; sorry #align category_theory.pi.comap_comp CategoryTheory.Pi.comapComp @@ -160,9 +160,17 @@ def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s obj X := { obj := fun Y s => Sum.rec X Y s map := fun {Y} {Y'} f s => Sum.rec (fun i => 𝟙 (X i)) f s - map_id := fun Y => sorry - map_comp := sorry } - map {X} {X'} f := { app := fun Y s => Sum.rec f (fun j => 𝟙 (Y j)) s } + map_id := fun Y => by + dsimp + simp only [CategoryStruct.id] + funext s + match s with + | .inl i => simp + | .inr j => simp + map_comp := fun {Y₁} {Y₂} {Y₃} f g => by dsimp; funext s; cases s; repeat {simp} } + map {X} {X'} f := + { app := fun Y s => Sum.rec f (fun j => 𝟙 (Y j)) s + naturality := sorry } map_id := sorry map_comp := sorry #align category_theory.pi.sum CategoryTheory.Pi.sum From ebdee62603ecb810215bf8efb9c21fce4d702269 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 13:39:09 -0500 Subject: [PATCH 040/126] finally fixed --- Mathlib/CategoryTheory/Pi/Basic.lean | 50 +++++++++++++++++++--------- 1 file changed, 34 insertions(+), 16 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 100236a7456ec..b112927dfd99a 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -120,8 +120,14 @@ def comapComp (f : K → J) (g : J → I) : comap C g ⋙ comap (C ∘ g) f ≅ app := fun X b => 𝟙 (X (g (f b))) naturality := fun X Y f' => by simp only [comap,Function.comp]; funext; simp } - hom_inv_id := by simp only [comap]; aesop_cat; sorry - inv_hom_id := by simp only [comap]; sorry + hom_inv_id := by + simp only [comap] + ext Y + simp [CategoryStruct.comp,CategoryStruct.id] + inv_hom_id := by + simp only [comap] + ext X + simp [CategoryStruct.comp,CategoryStruct.id] #align category_theory.pi.comap_comp CategoryTheory.Pi.comapComp /-- The natural isomorphism between pulling back then evaluating, and just evaluating. -/ @@ -154,12 +160,18 @@ instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) /-- The bifunctor combining an `I`-indexed family of objects with a `J`-indexed family of objects to obtain an `I ⊕ J`-indexed family of objects. -/ -@[simps!] +@[simps] def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s where obj X := - { obj := fun Y s => Sum.rec X Y s - map := fun {Y} {Y'} f s => Sum.rec (fun i => 𝟙 (X i)) f s + { obj := fun Y s => + match s with + | .inl i => X i + | .inr j => Y j + map := fun {Y} {Y'} f s => + match s with + | .inl i => 𝟙 (X i) + | .inr j => f j map_id := fun Y => by dsimp simp only [CategoryStruct.id] @@ -167,12 +179,17 @@ def sum : (∀ i, C i) ⥤ (∀ j, D j) ⥤ ∀ s : Sum I J, Sum.elim C D s match s with | .inl i => simp | .inr j => simp - map_comp := fun {Y₁} {Y₂} {Y₃} f g => by dsimp; funext s; cases s; repeat {simp} } + map_comp := fun {Y₁} {Y₂} {Y₃} f g => by funext s; cases s; repeat {simp} } map {X} {X'} f := - { app := fun Y s => Sum.rec f (fun j => 𝟙 (Y j)) s - naturality := sorry } - map_id := sorry - map_comp := sorry + { app := fun Y s => + match s with + | .inl i => f i + | .inr j => 𝟙 (Y j) + naturality := fun {Y} {Y'} g => by funext s; cases s; repeat {simp} } + map_id := fun X => by + ext Y; dsimp; simp only [CategoryStruct.id]; funext s; cases s; repeat {simp} + map_comp := fun f g => by + ext Y; dsimp; simp only [CategoryStruct.comp]; funext s; cases s; repeat {simp} #align category_theory.pi.sum CategoryTheory.Pi.sum end @@ -217,7 +234,7 @@ variable {D : I → Type u₁} [∀ i, Category.{v₁} (D i)] {A : Type u₁} [C /-- Assemble an `I`-indexed family of functors into a functor between the pi types. -/ -@[simps!] +@[simps] def pi (F : ∀ i, C i ⥤ D i) : (∀ i, C i) ⥤ ∀ i, D i where obj f i := (F i).obj (f i) @@ -226,7 +243,7 @@ def pi (F : ∀ i, C i ⥤ D i) : (∀ i, C i) ⥤ ∀ i, D i /-- Similar to `pi`, but all functors come from the same category `A` -/ -@[simps!] +@[simps] def pi' (f : ∀ i, A ⥤ C i) : A ⥤ ∀ i, C i where obj a i := (f i).obj a @@ -263,10 +280,11 @@ theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ (Pi.eval C i) = f' specialize h i have := congr_obj h X simpa - · intro x y p - ext i + · intro X Y g + dsimp + funext i specialize h i - have := congr_hom h p + have := congr_hom h g simpa #align category_theory.functor.pi_ext CategoryTheory.Functor.pi_ext @@ -285,7 +303,7 @@ variable {F G : ∀ i, C i ⥤ D i} @[simps!] def pi (α : ∀ i, F i ⟶ G i) : Functor.pi F ⟶ Functor.pi G where app f i := (α i).app (f i) - naturality := by intro X Y p; simp [Functor.pi]; aesop_cat + naturality := fun X Y f => by simp [Functor.pi,CategoryStruct.comp] #align category_theory.nat_trans.pi CategoryTheory.NatTrans.pi end NatTrans From c1897788f9f7e0db3e9e597eceba0c9b9dcb33ad Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 13:50:07 -0500 Subject: [PATCH 041/126] lint --- Mathlib/CategoryTheory/Pi/Basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index b112927dfd99a..d34a8fd7271ef 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -37,7 +37,7 @@ instance pi : Category.{max w₀ v₁} (∀ i, C i) #align category_theory.pi CategoryTheory.pi /-- This provides some assistance to typeclass search in a common situation, -which otherwise fails. (Without this `category_theory.pi.has_limit_of_has_limit_comp_eval` fails.) +which otherwise fails. (Without this `CategoryTheory.Pi.has_limit_of_has_limit_comp_eval` fails.) -/ abbrev pi' {I : Type v₁} (C : I → Type u₁) [∀ i, Category.{v₁} (C i)] : Category.{v₁} (∀ i, C i) := CategoryTheory.pi C @@ -262,7 +262,7 @@ theorem eqToHom_proj {x x' : ∀ i, C i} (h : x = x') (i : I) : end EqToHom -- One could add some natural isomorphisms showing --- how `functor.pi` commutes with `pi.eval` and `pi.comap`. +-- how `Functor.pi` commutes with `Pi.eval` and `Pi.comap`. @[simp] theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ Pi.eval C i = f i := by apply Functor.ext @@ -273,7 +273,8 @@ theorem pi'_eval (f : ∀ i, A ⥤ C i) (i : I) : pi' f ⋙ Pi.eval C i = f i : #align category_theory.functor.pi'_eval CategoryTheory.Functor.pi'_eval /-- Two functors to a product category are equal iff they agree on every coordinate. -/ -theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ (Pi.eval C i) = f' ⋙ (Pi.eval C i)) : f = f' := by +theorem pi_ext (f f' : A ⥤ ∀ i, C i) (h : ∀ i, f ⋙ (Pi.eval C i) = f' ⋙ (Pi.eval C i)) + : f = f' := by apply Functor.ext; rotate_left · intro X ext i @@ -309,4 +310,3 @@ def pi (α : ∀ i, F i ⟶ G i) : Functor.pi F ⟶ Functor.pi G where end NatTrans end CategoryTheory - From 7cbd34da5d3af38cdf96b25c863e22158f4acc99 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 13:55:57 -0500 Subject: [PATCH 042/126] add porting note for scary warning --- Mathlib/CategoryTheory/Pi/Basic.lean | 23 +++++++++++++++++------ 1 file changed, 17 insertions(+), 6 deletions(-) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index d34a8fd7271ef..3df30bcd5a306 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -142,12 +142,23 @@ section variable {J : Type w₀} {D : J → Type u₁} [∀ j, Category.{v₁} (D j)] -/- warning: category_theory.pi.sum_elim_category -> CategoryTheory.Pi.sumElimCategory is a dubious translation: -lean 3 declaration is - forall {I : Type.{u1}} (C : I -> Type.{u3}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u3} (C i)] {J : Type.{u1}} {D : J -> Type.{u3}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u3} (D j)] (s : Sum.{u1, u1} I J), CategoryTheory.Category.{u2, u3} (Sum.elim.{u1, u1, succ (succ u3)} I J Type.{u3} C D s) -but is expected to have type - forall {I : Type.{u3}} (C : I -> Type.{u1}) [_inst_1 : forall (i : I), CategoryTheory.Category.{u2, u1} (C i)] {J : Type.{u3}} {D : J -> Type.{u1}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u1} (D j)] (s : Sum.{u3, u3} I J), CategoryTheory.Category.{u2, u1} (Sum.elim.{u3, u3, succ (succ u1)} I J Type.{u1} C D s) -Case conversion may be inaccurate. Consider using '#align category_theory.pi.sum_elim_category CategoryTheory.Pi.sumElimCategoryₓ'. -/ +/- Porting note: We have the following warning but it just looks +to me to be permutation of the universe labels + +warning: category_theory.pi.sum_elim_category -> +\ CategoryTheory.Pi.sumElimCategory is a dubious translation: +lean 3 declaration is forall {I : Type.{u1}} (C : I -> Type.{u3}) [_inst_1 : +forall (i : I), CategoryTheory.Category.{u2, u3} (C i)] {J : Type.{u1}} {D : J +-> Type.{u3}} [_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u3} (D +j)] (s : Sum.{u1, u1} I J), CategoryTheory.Category.{u2, u3} (Sum.elim.{u1, u1, +succ (succ u3)} I J Type.{u3} C D s) but is expected to have type forall {I : +Type.{u3}} (C : I -> Type.{u1}) [_inst_1 : forall (i : I), +CategoryTheory.Category.{u2, u1} (C i)] {J : Type.{u3}} {D : J -> Type.{u1}} +[_inst_2 : forall (j : J), CategoryTheory.Category.{u2, u1} (D j)] (s : +Sum.{u3, u3} I J), CategoryTheory.Category.{u2, u1} (Sum.elim.{u3, u3, succ +(succ u1)} I J Type.{u1} C D s) Case conversion may be inaccurate. Consider +using '#align category_theory.pi.sum_elim_category +CategoryTheory.Pi.sumElimCategoryₓ'. -/ instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) | Sum.inl i => by dsimp From b0fba3552d171b19772d83849291ef8100853706 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 13:57:46 -0500 Subject: [PATCH 043/126] add porting note about proliferation of match --- Mathlib/CategoryTheory/Pi/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index 3df30bcd5a306..19d9e10022e0c 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -168,6 +168,9 @@ instance sumElimCategory : ∀ s : Sum I J, Category.{v₁} (Sum.elim C D s) infer_instance #align category_theory.pi.sum_elim_category CategoryTheory.Pi.sumElimCategory +/- Porting note: replaced `Sum.rec` with `match`'s per the error about +current state of code generation -/ + /-- The bifunctor combining an `I`-indexed family of objects with a `J`-indexed family of objects to obtain an `I ⊕ J`-indexed family of objects. -/ From 5d61758e863afe09236dc99758cfb32a5912ca0e Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 14:10:41 -0500 Subject: [PATCH 044/126] feat: port CategoryTheory.Groupoid From 466cb9fa0a505cd7c33e616a612acdfc51b18746 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 14:10:41 -0500 Subject: [PATCH 045/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Groupoid.lean | 169 +++++++++++++++++++++++++++ 1 file changed, 169 insertions(+) create mode 100644 Mathlib/CategoryTheory/Groupoid.lean diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean new file mode 100644 index 0000000000000..6593a09b137e2 --- /dev/null +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -0,0 +1,169 @@ +/- +Copyright (c) 2018 Reid Barton All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Reid Barton, Scott Morrison, David Wärn + +! This file was ported from Lean 3 source module category_theory.groupoid +! leanprover-community/mathlib commit dc6c365e751e34d100e80fe6e314c3c3e0fd2988 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.FullSubcategory +import Mathbin.CategoryTheory.Products.Basic +import Mathbin.CategoryTheory.Pi.Basic +import Mathbin.CategoryTheory.Category.Basic +import Mathbin.Combinatorics.Quiver.ConnectedComponent + +/-! +# Groupoids + +We define `groupoid` as a typeclass extending `category`, +asserting that all morphisms have inverses. + +The instance `is_iso.of_groupoid (f : X ⟶ Y) : is_iso f` means that you can then write +`inv f` to access the inverse of any morphism `f`. + +`groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y)` provides the equivalence between +isomorphisms and morphisms in a groupoid. + +We provide a (non-instance) constructor `groupoid.of_is_iso` from an existing category +with `is_iso f` for every `f`. + +## See also + +See also `category_theory.core` for the groupoid of isomorphisms in a category. +-/ + + +namespace CategoryTheory + +universe v v₂ u u₂ + +-- morphism levels before object levels. See note [category_theory universes]. +/-- A `groupoid` is a category such that all morphisms are isomorphisms. -/ +class Groupoid (obj : Type u) extends Category.{v} obj : Type max u (v + 1) where + inv : ∀ {X Y : obj}, (X ⟶ Y) → (Y ⟶ X) + inv_comp' : ∀ {X Y : obj} (f : X ⟶ Y), comp (inv f) f = id Y := by obviously + comp_inv' : ∀ {X Y : obj} (f : X ⟶ Y), comp f (inv f) = id X := by obviously +#align category_theory.groupoid CategoryTheory.Groupoid + +restate_axiom groupoid.inv_comp' + +restate_axiom groupoid.comp_inv' + +/-- A `large_groupoid` is a groupoid +where the objects live in `Type (u+1)` while the morphisms live in `Type u`. +-/ +abbrev LargeGroupoid (C : Type (u + 1)) : Type (u + 1) := + Groupoid.{u} C +#align category_theory.large_groupoid CategoryTheory.LargeGroupoid + +/-- A `small_groupoid` is a groupoid +where the objects and morphisms live in the same universe. +-/ +abbrev SmallGroupoid (C : Type u) : Type (u + 1) := + Groupoid.{u} C +#align category_theory.small_groupoid CategoryTheory.SmallGroupoid + +section + +variable {C : Type u} [Groupoid.{v} C] {X Y : C} + +-- see Note [lower instance priority] +instance (priority := 100) IsIso.of_groupoid (f : X ⟶ Y) : IsIso f := + ⟨⟨Groupoid.inv f, Groupoid.comp_inv f, Groupoid.inv_comp f⟩⟩ +#align category_theory.is_iso.of_groupoid CategoryTheory.IsIso.of_groupoid + +@[simp] +theorem Groupoid.inv_eq_inv (f : X ⟶ Y) : Groupoid.inv f = inv f := + IsIso.eq_inv_of_hom_inv_id <| Groupoid.comp_inv f +#align category_theory.groupoid.inv_eq_inv CategoryTheory.Groupoid.inv_eq_inv + +/-- `groupoid.inv` is involutive. -/ +@[simps] +def Groupoid.invEquiv : (X ⟶ Y) ≃ (Y ⟶ X) := + ⟨Groupoid.inv, Groupoid.inv, fun f => by simp, fun f => by simp⟩ +#align category_theory.groupoid.inv_equiv CategoryTheory.Groupoid.invEquiv + +instance (priority := 100) groupoidHasInvolutiveReverse : Quiver.HasInvolutiveReverse C + where + reverse' X Y f := Groupoid.inv f + inv' X Y f := by + dsimp [Quiver.reverse] + simp +#align category_theory.groupoid_has_involutive_reverse CategoryTheory.groupoidHasInvolutiveReverse + +@[simp] +theorem Groupoid.reverse_eq_inv (f : X ⟶ Y) : Quiver.reverse f = Groupoid.inv f := + rfl +#align category_theory.groupoid.reverse_eq_inv CategoryTheory.Groupoid.reverse_eq_inv + +instance functorMapReverse {D : Type _} [Groupoid D] (F : C ⥤ D) : F.toPrefunctor.MapReverse + where map_reverse' X Y f := by + simp only [Quiver.reverse, Quiver.HasReverse.reverse', groupoid.inv_eq_inv, + functor.to_prefunctor_map, functor.map_inv] +#align category_theory.functor_map_reverse CategoryTheory.functorMapReverse + +variable (X Y) + +/-- In a groupoid, isomorphisms are equivalent to morphisms. -/ +def Groupoid.isoEquivHom : (X ≅ Y) ≃ (X ⟶ Y) + where + toFun := Iso.hom + invFun f := ⟨f, Groupoid.inv f⟩ + left_inv i := Iso.ext rfl + right_inv f := rfl +#align category_theory.groupoid.iso_equiv_hom CategoryTheory.Groupoid.isoEquivHom + +variable (C) + +/-- The functor from a groupoid `C` to its opposite sending every morphism to its inverse. -/ +@[simps] +noncomputable def Groupoid.invFunctor : C ⥤ Cᵒᵖ + where + obj := Opposite.op + map {X Y} f := (inv f).op +#align category_theory.groupoid.inv_functor CategoryTheory.Groupoid.invFunctor + +end + +section + +variable {C : Type u} [Category.{v} C] + +/-- A category where every morphism `is_iso` is a groupoid. -/ +noncomputable def Groupoid.ofIsIso (all_is_iso : ∀ {X Y : C} (f : X ⟶ Y), IsIso f) : Groupoid.{v} C + where inv X Y f := inv f +#align category_theory.groupoid.of_is_iso CategoryTheory.Groupoid.ofIsIso + +/-- A category with a unique morphism between any two objects is a groupoid -/ +def Groupoid.ofHomUnique (all_unique : ∀ {X Y : C}, Unique (X ⟶ Y)) : Groupoid.{v} C + where inv X Y f := all_unique.default +#align category_theory.groupoid.of_hom_unique CategoryTheory.Groupoid.ofHomUnique + +end + +instance InducedCategory.groupoid {C : Type u} (D : Type u₂) [Groupoid.{v} D] (F : C → D) : + Groupoid.{v} (InducedCategory D F) := + { InducedCategory.category F with + inv := fun X Y f => Groupoid.inv f + inv_comp' := fun X Y f => Groupoid.inv_comp f + comp_inv' := fun X Y f => Groupoid.comp_inv f } +#align category_theory.induced_category.groupoid CategoryTheory.InducedCategory.groupoid + +section + +instance groupoidPi {I : Type u} {J : I → Type u₂} [∀ i, Groupoid.{v} (J i)] : + Groupoid.{max u v} (∀ i : I, J i) + where inv (x y : ∀ i, J i) (f : ∀ i, x i ⟶ y i) := fun i : I => Groupoid.inv (f i) +#align category_theory.groupoid_pi CategoryTheory.groupoidPi + +instance groupoidProd {α : Type u} {β : Type v} [Groupoid.{u₂} α] [Groupoid.{v₂} β] : + Groupoid.{max u₂ v₂} (α × β) + where inv (x y : α × β) (f : x ⟶ y) := (Groupoid.inv f.1, Groupoid.inv f.2) +#align category_theory.groupoid_prod CategoryTheory.groupoidProd + +end + +end CategoryTheory + From e2cfafa230c27afef0986c8b42a131c46ce1a1c3 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 14:10:41 -0500 Subject: [PATCH 046/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Groupoid.lean | 10 +++++----- 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index a02eb99136995..3802a1d4c476e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -235,6 +235,7 @@ import Mathlib.CategoryTheory.Functor.Default import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial import Mathlib.CategoryTheory.Functor.InvIsos +import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index 6593a09b137e2..0c1e2ab87e1e7 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -8,11 +8,11 @@ Authors: Reid Barton, Scott Morrison, David Wärn ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.FullSubcategory -import Mathbin.CategoryTheory.Products.Basic -import Mathbin.CategoryTheory.Pi.Basic -import Mathbin.CategoryTheory.Category.Basic -import Mathbin.Combinatorics.Quiver.ConnectedComponent +import Mathlib.CategoryTheory.FullSubcategory +import Mathlib.CategoryTheory.Products.Basic +import Mathlib.CategoryTheory.Pi.Basic +import Mathlib.CategoryTheory.Category.Basic +import Mathlib.Combinatorics.Quiver.ConnectedComponent /-! # Groupoids From ee114ee578264c457aeacbf16885addedb398722 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 14:39:58 -0500 Subject: [PATCH 047/126] initial pass --- Mathlib/CategoryTheory/Groupoid.lean | 53 +++++++++++++++------------- 1 file changed, 28 insertions(+), 25 deletions(-) diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index 0c1e2ab87e1e7..d9a91a84bcf97 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -20,18 +20,18 @@ import Mathlib.Combinatorics.Quiver.ConnectedComponent We define `groupoid` as a typeclass extending `category`, asserting that all morphisms have inverses. -The instance `is_iso.of_groupoid (f : X ⟶ Y) : is_iso f` means that you can then write +The instance `IsIso.ofGroupoid (f : X ⟶ Y) : IsIso f` means that you can then write `inv f` to access the inverse of any morphism `f`. -`groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y)` provides the equivalence between +`Groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y)` provides the equivalence between isomorphisms and morphisms in a groupoid. -We provide a (non-instance) constructor `groupoid.of_is_iso` from an existing category -with `is_iso f` for every `f`. +We provide a (non-instance) constructor `Groupoid.ofIsIso` from an existing category +with `IsIso f` for every `f`. ## See also -See also `category_theory.core` for the groupoid of isomorphisms in a category. +See also `CategoryTheory.Core` for the groupoid of isomorphisms in a category. -/ @@ -43,14 +43,10 @@ universe v v₂ u u₂ /-- A `groupoid` is a category such that all morphisms are isomorphisms. -/ class Groupoid (obj : Type u) extends Category.{v} obj : Type max u (v + 1) where inv : ∀ {X Y : obj}, (X ⟶ Y) → (Y ⟶ X) - inv_comp' : ∀ {X Y : obj} (f : X ⟶ Y), comp (inv f) f = id Y := by obviously - comp_inv' : ∀ {X Y : obj} (f : X ⟶ Y), comp f (inv f) = id X := by obviously + inv_comp : ∀ {X Y : obj} (f : X ⟶ Y), comp (inv f) f = id Y := by aesop_cat + comp_inv : ∀ {X Y : obj} (f : X ⟶ Y), comp f (inv f) = id X := by aesop_cat #align category_theory.groupoid CategoryTheory.Groupoid -restate_axiom groupoid.inv_comp' - -restate_axiom groupoid.comp_inv' - /-- A `large_groupoid` is a groupoid where the objects live in `Type (u+1)` while the morphisms live in `Type u`. -/ @@ -75,7 +71,7 @@ instance (priority := 100) IsIso.of_groupoid (f : X ⟶ Y) : IsIso f := #align category_theory.is_iso.of_groupoid CategoryTheory.IsIso.of_groupoid @[simp] -theorem Groupoid.inv_eq_inv (f : X ⟶ Y) : Groupoid.inv f = inv f := +theorem Groupoid.inv_eq_inv (f : X ⟶ Y) : Groupoid.inv f = CategoryTheory.inv f := IsIso.eq_inv_of_hom_inv_id <| Groupoid.comp_inv f #align category_theory.groupoid.inv_eq_inv CategoryTheory.Groupoid.inv_eq_inv @@ -87,8 +83,8 @@ def Groupoid.invEquiv : (X ⟶ Y) ≃ (Y ⟶ X) := instance (priority := 100) groupoidHasInvolutiveReverse : Quiver.HasInvolutiveReverse C where - reverse' X Y f := Groupoid.inv f - inv' X Y f := by + reverse' f := Groupoid.inv f + inv' f := by dsimp [Quiver.reverse] simp #align category_theory.groupoid_has_involutive_reverse CategoryTheory.groupoidHasInvolutiveReverse @@ -99,9 +95,10 @@ theorem Groupoid.reverse_eq_inv (f : X ⟶ Y) : Quiver.reverse f = Groupoid.inv #align category_theory.groupoid.reverse_eq_inv CategoryTheory.Groupoid.reverse_eq_inv instance functorMapReverse {D : Type _} [Groupoid D] (F : C ⥤ D) : F.toPrefunctor.MapReverse - where map_reverse' X Y f := by - simp only [Quiver.reverse, Quiver.HasReverse.reverse', groupoid.inv_eq_inv, - functor.to_prefunctor_map, functor.map_inv] + where + map_reverse' f := by + simp only [Quiver.reverse, Quiver.HasReverse.reverse', Groupoid.inv_eq_inv, + Functor.map_inv] #align category_theory.functor_map_reverse CategoryTheory.functorMapReverse variable (X Y) @@ -110,7 +107,7 @@ variable (X Y) def Groupoid.isoEquivHom : (X ≅ Y) ≃ (X ⟶ Y) where toFun := Iso.hom - invFun f := ⟨f, Groupoid.inv f⟩ + invFun f := ⟨f, Groupoid.inv f, (by aesop_cat), (by aesop_cat)⟩ left_inv i := Iso.ext rfl right_inv f := rfl #align category_theory.groupoid.iso_equiv_hom CategoryTheory.Groupoid.isoEquivHom @@ -133,12 +130,15 @@ variable {C : Type u} [Category.{v} C] /-- A category where every morphism `is_iso` is a groupoid. -/ noncomputable def Groupoid.ofIsIso (all_is_iso : ∀ {X Y : C} (f : X ⟶ Y), IsIso f) : Groupoid.{v} C - where inv X Y f := inv f + where + inv f := inv f + comp_inv := sorry + inv_comp := sorry #align category_theory.groupoid.of_is_iso CategoryTheory.Groupoid.ofIsIso /-- A category with a unique morphism between any two objects is a groupoid -/ def Groupoid.ofHomUnique (all_unique : ∀ {X Y : C}, Unique (X ⟶ Y)) : Groupoid.{v} C - where inv X Y f := all_unique.default + where inv _ := all_unique.default #align category_theory.groupoid.of_hom_unique CategoryTheory.Groupoid.ofHomUnique end @@ -146,21 +146,24 @@ end instance InducedCategory.groupoid {C : Type u} (D : Type u₂) [Groupoid.{v} D] (F : C → D) : Groupoid.{v} (InducedCategory D F) := { InducedCategory.category F with - inv := fun X Y f => Groupoid.inv f - inv_comp' := fun X Y f => Groupoid.inv_comp f - comp_inv' := fun X Y f => Groupoid.comp_inv f } + inv := fun f => Groupoid.inv f + inv_comp := fun f => Groupoid.inv_comp f + comp_inv := fun f => Groupoid.comp_inv f } #align category_theory.induced_category.groupoid CategoryTheory.InducedCategory.groupoid section instance groupoidPi {I : Type u} {J : I → Type u₂} [∀ i, Groupoid.{v} (J i)] : Groupoid.{max u v} (∀ i : I, J i) - where inv (x y : ∀ i, J i) (f : ∀ i, x i ⟶ y i) := fun i : I => Groupoid.inv (f i) + where + inv f := fun i : I => Groupoid.inv (f i) + comp_inv := sorry + inv_comp := sorry #align category_theory.groupoid_pi CategoryTheory.groupoidPi instance groupoidProd {α : Type u} {β : Type v} [Groupoid.{u₂} α] [Groupoid.{v₂} β] : Groupoid.{max u₂ v₂} (α × β) - where inv (x y : α × β) (f : x ⟶ y) := (Groupoid.inv f.1, Groupoid.inv f.2) + where inv f := (Groupoid.inv f.1, Groupoid.inv f.2) #align category_theory.groupoid_prod CategoryTheory.groupoidProd end From 7598ce4d98b3d216174eb82b967c5ca242edece0 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 19:35:41 -0500 Subject: [PATCH 048/126] fix errors --- Mathlib/CategoryTheory/Groupoid.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index d9a91a84bcf97..2888ef8cb7bfb 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -131,9 +131,9 @@ variable {C : Type u} [Category.{v} C] /-- A category where every morphism `is_iso` is a groupoid. -/ noncomputable def Groupoid.ofIsIso (all_is_iso : ∀ {X Y : C} (f : X ⟶ Y), IsIso f) : Groupoid.{v} C where - inv f := inv f - comp_inv := sorry - inv_comp := sorry + inv := fun f => CategoryTheory.inv f + comp_inv := by aesop_cat + inv_comp := fun f => Classical.choose_spec (all_is_iso f).out|>.right #align category_theory.groupoid.of_is_iso CategoryTheory.Groupoid.ofIsIso /-- A category with a unique morphism between any two objects is a groupoid -/ @@ -157,8 +157,8 @@ instance groupoidPi {I : Type u} {J : I → Type u₂} [∀ i, Groupoid.{v} (J i Groupoid.{max u v} (∀ i : I, J i) where inv f := fun i : I => Groupoid.inv (f i) - comp_inv := sorry - inv_comp := sorry + comp_inv := fun f => by funext i; apply Groupoid.comp_inv + inv_comp := fun f => by funext i; apply Groupoid.inv_comp #align category_theory.groupoid_pi CategoryTheory.groupoidPi instance groupoidProd {α : Type u} {β : Type v} [Groupoid.{u₂} α] [Groupoid.{v₂} β] : From 9b2cb68ad3aad071c74083893ff77406d07b53ce Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 19:46:46 -0500 Subject: [PATCH 049/126] lint --- Mathlib/CategoryTheory/Groupoid.lean | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) diff --git a/Mathlib/CategoryTheory/Groupoid.lean b/Mathlib/CategoryTheory/Groupoid.lean index 2888ef8cb7bfb..8ebf75d278f76 100644 --- a/Mathlib/CategoryTheory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid.lean @@ -23,7 +23,7 @@ asserting that all morphisms have inverses. The instance `IsIso.ofGroupoid (f : X ⟶ Y) : IsIso f` means that you can then write `inv f` to access the inverse of any morphism `f`. -`Groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y)` provides the equivalence between +`Groupoid.isoEquivHom : (X ≅ Y) ≃ (X ⟶ Y)` provides the equivalence between isomorphisms and morphisms in a groupoid. We provide a (non-instance) constructor `Groupoid.ofIsIso` from an existing category @@ -39,22 +39,25 @@ namespace CategoryTheory universe v v₂ u u₂ --- morphism levels before object levels. See note [category_theory universes]. +-- morphism levels before object levels. See note [CategoryTheory universes]. /-- A `groupoid` is a category such that all morphisms are isomorphisms. -/ class Groupoid (obj : Type u) extends Category.{v} obj : Type max u (v + 1) where + /-- The inverse morphism -/ inv : ∀ {X Y : obj}, (X ⟶ Y) → (Y ⟶ X) + /-- `inv f` composed `f` is the identity -/ inv_comp : ∀ {X Y : obj} (f : X ⟶ Y), comp (inv f) f = id Y := by aesop_cat + /-- `f` composed with `inv f` is the identity -/ comp_inv : ∀ {X Y : obj} (f : X ⟶ Y), comp f (inv f) = id X := by aesop_cat #align category_theory.groupoid CategoryTheory.Groupoid -/-- A `large_groupoid` is a groupoid +/-- A `LargeGroupoid` is a groupoid where the objects live in `Type (u+1)` while the morphisms live in `Type u`. -/ abbrev LargeGroupoid (C : Type (u + 1)) : Type (u + 1) := Groupoid.{u} C #align category_theory.large_groupoid CategoryTheory.LargeGroupoid -/-- A `small_groupoid` is a groupoid +/-- A `SmallGroupoid` is a groupoid where the objects and morphisms live in the same universe. -/ abbrev SmallGroupoid (C : Type u) : Type (u + 1) := @@ -128,7 +131,7 @@ section variable {C : Type u} [Category.{v} C] -/-- A category where every morphism `is_iso` is a groupoid. -/ +/-- A category where every morphism `IsIso` is a groupoid. -/ noncomputable def Groupoid.ofIsIso (all_is_iso : ∀ {X Y : C} (f : X ⟶ Y), IsIso f) : Groupoid.{v} C where inv := fun f => CategoryTheory.inv f @@ -169,4 +172,3 @@ instance groupoidProd {α : Type u} {β : Type v} [Groupoid.{u₂} α] [Groupoid end end CategoryTheory - From e2aebba7aad66d6feb4b27767e682faadaf7834d Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 20:02:31 -0500 Subject: [PATCH 050/126] feat: port CategoryTheory.EpiMono From 7f4258066570dbb31d80a289e302c5359055b327 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 20:02:31 -0500 Subject: [PATCH 051/126] Initial file copy from mathport --- Mathlib/CategoryTheory/EpiMono.lean | 269 ++++++++++++++++++++++++++++ 1 file changed, 269 insertions(+) create mode 100644 Mathlib/CategoryTheory/EpiMono.lean diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean new file mode 100644 index 0000000000000..4b39ca355dab1 --- /dev/null +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -0,0 +1,269 @@ +/- +Copyright (c) 2019 Reid Barton. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Reid Barton, Scott Morrison + +! This file was ported from Lean 3 source module category_theory.epi_mono +! leanprover-community/mathlib commit 48085f140e684306f9e7da907cd5932056d1aded +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Opposites +import Mathbin.CategoryTheory.Groupoid + +/-! +# Facts about epimorphisms and monomorphisms. + +The definitions of `epi` and `mono` are in `category_theory.category`, +since they are used by some lemmas for `iso`, which is used everywhere. +-/ + + +universe v₁ v₂ u₁ u₂ + +namespace CategoryTheory + +variable {C : Type u₁} [Category.{v₁} C] + +instance unop_mono_of_epi {A B : Cᵒᵖ} (f : A ⟶ B) [Epi f] : Mono f.unop := + ⟨fun Z g h eq => Quiver.Hom.op_inj ((cancel_epi f).1 (Quiver.Hom.unop_inj Eq))⟩ +#align category_theory.unop_mono_of_epi CategoryTheory.unop_mono_of_epi + +instance unop_epi_of_mono {A B : Cᵒᵖ} (f : A ⟶ B) [Mono f] : Epi f.unop := + ⟨fun Z g h eq => Quiver.Hom.op_inj ((cancel_mono f).1 (Quiver.Hom.unop_inj Eq))⟩ +#align category_theory.unop_epi_of_mono CategoryTheory.unop_epi_of_mono + +instance op_mono_of_epi {A B : C} (f : A ⟶ B) [Epi f] : Mono f.op := + ⟨fun Z g h eq => Quiver.Hom.unop_inj ((cancel_epi f).1 (Quiver.Hom.op_inj Eq))⟩ +#align category_theory.op_mono_of_epi CategoryTheory.op_mono_of_epi + +instance op_epi_of_mono {A B : C} (f : A ⟶ B) [Mono f] : Epi f.op := + ⟨fun Z g h eq => Quiver.Hom.unop_inj ((cancel_mono f).1 (Quiver.Hom.op_inj Eq))⟩ +#align category_theory.op_epi_of_mono CategoryTheory.op_epi_of_mono + +/-- A split monomorphism is a morphism `f : X ⟶ Y` with a given retraction `retraction f : Y ⟶ X` +such that `f ≫ retraction f = 𝟙 X`. + +Every split monomorphism is a monomorphism. +-/ +@[ext, nolint has_nonempty_instance] +structure SplitMono {X Y : C} (f : X ⟶ Y) where + retraction : Y ⟶ X + id' : f ≫ retraction = 𝟙 X := by obviously +#align category_theory.split_mono CategoryTheory.SplitMono + +restate_axiom split_mono.id' + +attribute [simp, reassoc.1] split_mono.id + +/-- `is_split_mono f` is the assertion that `f` admits a retraction -/ +class IsSplitMono {X Y : C} (f : X ⟶ Y) : Prop where + exists_splitMono : Nonempty (SplitMono f) +#align category_theory.is_split_mono CategoryTheory.IsSplitMono + +/-- A constructor for `is_split_mono f` taking a `split_mono f` as an argument -/ +theorem IsSplitMono.mk' {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : IsSplitMono f := + ⟨Nonempty.intro sm⟩ +#align category_theory.is_split_mono.mk' CategoryTheory.IsSplitMono.mk' + +/-- A split epimorphism is a morphism `f : X ⟶ Y` with a given section `section_ f : Y ⟶ X` +such that `section_ f ≫ f = 𝟙 Y`. +(Note that `section` is a reserved keyword, so we append an underscore.) + +Every split epimorphism is an epimorphism. +-/ +@[ext, nolint has_nonempty_instance] +structure SplitEpi {X Y : C} (f : X ⟶ Y) where + section_ : Y ⟶ X + id' : section_ ≫ f = 𝟙 Y := by obviously +#align category_theory.split_epi CategoryTheory.SplitEpi + +restate_axiom split_epi.id' + +attribute [simp, reassoc.1] split_epi.id + +/-- `is_split_epi f` is the assertion that `f` admits a section -/ +class IsSplitEpi {X Y : C} (f : X ⟶ Y) : Prop where + exists_splitEpi : Nonempty (SplitEpi f) +#align category_theory.is_split_epi CategoryTheory.IsSplitEpi + +/-- A constructor for `is_split_epi f` taking a `split_epi f` as an argument -/ +theorem IsSplitEpi.mk' {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : IsSplitEpi f := + ⟨Nonempty.intro se⟩ +#align category_theory.is_split_epi.mk' CategoryTheory.IsSplitEpi.mk' + +/-- The chosen retraction of a split monomorphism. -/ +noncomputable def retraction {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : Y ⟶ X := + hf.exists_splitMono.some.retraction +#align category_theory.retraction CategoryTheory.retraction + +@[simp, reassoc.1] +theorem IsSplitMono.id {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : f ≫ retraction f = 𝟙 X := + hf.exists_splitMono.some.id +#align category_theory.is_split_mono.id CategoryTheory.IsSplitMono.id + +/-- The retraction of a split monomorphism has an obvious section. -/ +def SplitMono.splitEpi {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : SplitEpi sm.retraction + where section_ := f +#align category_theory.split_mono.split_epi CategoryTheory.SplitMono.splitEpi + +/-- The retraction of a split monomorphism is itself a split epimorphism. -/ +instance retraction_isSplitEpi {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : + IsSplitEpi (retraction f) := + IsSplitEpi.mk' (SplitMono.splitEpi _) +#align category_theory.retraction_is_split_epi CategoryTheory.retraction_isSplitEpi + +/-- A split mono which is epi is an iso. -/ +theorem isIso_of_epi_of_isSplitMono {X Y : C} (f : X ⟶ Y) [IsSplitMono f] [Epi f] : IsIso f := + ⟨⟨retraction f, ⟨by simp, by simp [← cancel_epi f]⟩⟩⟩ +#align category_theory.is_iso_of_epi_of_is_split_mono CategoryTheory.isIso_of_epi_of_isSplitMono + +/-- The chosen section of a split epimorphism. +(Note that `section` is a reserved keyword, so we append an underscore.) +-/ +noncomputable def section_ {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : Y ⟶ X := + hf.exists_splitEpi.some.section_ +#align category_theory.section_ CategoryTheory.section_ + +@[simp, reassoc.1] +theorem IsSplitEpi.id {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : section_ f ≫ f = 𝟙 Y := + hf.exists_splitEpi.some.id +#align category_theory.is_split_epi.id CategoryTheory.IsSplitEpi.id + +/-- The section of a split epimorphism has an obvious retraction. -/ +def SplitEpi.splitMono {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : SplitMono se.section_ + where retraction := f +#align category_theory.split_epi.split_mono CategoryTheory.SplitEpi.splitMono + +/-- The section of a split epimorphism is itself a split monomorphism. -/ +instance section_isSplitMono {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : IsSplitMono (section_ f) := + IsSplitMono.mk' (SplitEpi.splitMono _) +#align category_theory.section_is_split_mono CategoryTheory.section_isSplitMono + +/-- A split epi which is mono is an iso. -/ +theorem isIso_of_mono_of_isSplitEpi {X Y : C} (f : X ⟶ Y) [Mono f] [IsSplitEpi f] : IsIso f := + ⟨⟨section_ f, ⟨by simp [← cancel_mono f], by simp⟩⟩⟩ +#align category_theory.is_iso_of_mono_of_is_split_epi CategoryTheory.isIso_of_mono_of_isSplitEpi + +/-- Every iso is a split mono. -/ +instance (priority := 100) IsSplitMono.of_iso {X Y : C} (f : X ⟶ Y) [IsIso f] : IsSplitMono f := + IsSplitMono.mk' { retraction := inv f } +#align category_theory.is_split_mono.of_iso CategoryTheory.IsSplitMono.of_iso + +/-- Every iso is a split epi. -/ +instance (priority := 100) IsSplitEpi.of_iso {X Y : C} (f : X ⟶ Y) [IsIso f] : IsSplitEpi f := + IsSplitEpi.mk' { section_ := inv f } +#align category_theory.is_split_epi.of_iso CategoryTheory.IsSplitEpi.of_iso + +theorem SplitMono.mono {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : Mono f := + { right_cancellation := fun Z g h w => by replace w := w =≫ sm.retraction; simpa using w } +#align category_theory.split_mono.mono CategoryTheory.SplitMono.mono + +/-- Every split mono is a mono. -/ +instance (priority := 100) IsSplitMono.mono {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : Mono f := + hf.exists_splitMono.some.Mono +#align category_theory.is_split_mono.mono CategoryTheory.IsSplitMono.mono + +theorem SplitEpi.epi {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : Epi f := + { left_cancellation := fun Z g h w => by replace w := se.section_ ≫= w; simpa using w } +#align category_theory.split_epi.epi CategoryTheory.SplitEpi.epi + +/-- Every split epi is an epi. -/ +instance (priority := 100) IsSplitEpi.epi {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : Epi f := + hf.exists_splitEpi.some.Epi +#align category_theory.is_split_epi.epi CategoryTheory.IsSplitEpi.epi + +/-- Every split mono whose retraction is mono is an iso. -/ +theorem IsIso.of_mono_retraction' {X Y : C} {f : X ⟶ Y} (hf : SplitMono f) [Mono <| hf.retraction] : + IsIso f := + ⟨⟨hf.retraction, ⟨by simp, (cancel_mono_id <| hf.retraction).mp (by simp)⟩⟩⟩ +#align category_theory.is_iso.of_mono_retraction' CategoryTheory.IsIso.of_mono_retraction' + +/-- Every split mono whose retraction is mono is an iso. -/ +theorem IsIso.of_mono_retraction {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] + [hf' : Mono <| retraction f] : IsIso f := + @IsIso.of_mono_retraction' _ _ _ _ _ hf.exists_splitMono.some hf' +#align category_theory.is_iso.of_mono_retraction CategoryTheory.IsIso.of_mono_retraction + +/-- Every split epi whose section is epi is an iso. -/ +theorem IsIso.of_epi_section' {X Y : C} {f : X ⟶ Y} (hf : SplitEpi f) [Epi <| hf.section_] : + IsIso f := + ⟨⟨hf.section_, ⟨(cancel_epi_id <| hf.section_).mp (by simp), by simp⟩⟩⟩ +#align category_theory.is_iso.of_epi_section' CategoryTheory.IsIso.of_epi_section' + +/-- Every split epi whose section is epi is an iso. -/ +theorem IsIso.of_epi_section {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] [hf' : Epi <| section_ f] : + IsIso f := + @IsIso.of_epi_section' _ _ _ _ _ hf.exists_splitEpi.some hf' +#align category_theory.is_iso.of_epi_section CategoryTheory.IsIso.of_epi_section + +-- FIXME this has unnecessarily become noncomputable! +/-- A category where every morphism has a `trunc` retraction is computably a groupoid. -/ +noncomputable def Groupoid.ofTruncSplitMono + (all_split_mono : ∀ {X Y : C} (f : X ⟶ Y), Trunc (IsSplitMono f)) : Groupoid.{v₁} C := + by + apply groupoid.of_is_iso + intro X Y f + trunc_cases all_split_mono f + trunc_cases all_split_mono (retraction f) + apply is_iso.of_mono_retraction +#align category_theory.groupoid.of_trunc_split_mono CategoryTheory.Groupoid.ofTruncSplitMono + +section + +variable (C) + +/-- A split mono category is a category in which every monomorphism is split. -/ +class SplitMonoCategory where + isSplitMono_of_mono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], IsSplitMono f +#align category_theory.split_mono_category CategoryTheory.SplitMonoCategory + +/-- A split epi category is a category in which every epimorphism is split. -/ +class SplitEpiCategory where + isSplitEpi_of_epi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], IsSplitEpi f +#align category_theory.split_epi_category CategoryTheory.SplitEpiCategory + +end + +/-- In a category in which every monomorphism is split, every monomorphism splits. This is not an + instance because it would create an instance loop. -/ +theorem isSplitMono_of_mono [SplitMonoCategory C] {X Y : C} (f : X ⟶ Y) [Mono f] : IsSplitMono f := + SplitMonoCategory.isSplitMono_of_mono _ +#align category_theory.is_split_mono_of_mono CategoryTheory.isSplitMono_of_mono + +/-- In a category in which every epimorphism is split, every epimorphism splits. This is not an + instance because it would create an instance loop. -/ +theorem isSplitEpi_of_epi [SplitEpiCategory C] {X Y : C} (f : X ⟶ Y) [Epi f] : IsSplitEpi f := + SplitEpiCategory.isSplitEpi_of_epi _ +#align category_theory.is_split_epi_of_epi CategoryTheory.isSplitEpi_of_epi + +section + +variable {D : Type u₂} [Category.{v₂} D] + +/-- Split monomorphisms are also absolute monomorphisms. -/ +@[simps] +def SplitMono.map {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) (F : C ⥤ D) : SplitMono (F.map f) + where + retraction := F.map sm.retraction + id' := by rw [← functor.map_comp, split_mono.id, Functor.map_id] +#align category_theory.split_mono.map CategoryTheory.SplitMono.map + +/-- Split epimorphisms are also absolute epimorphisms. -/ +@[simps] +def SplitEpi.map {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) (F : C ⥤ D) : SplitEpi (F.map f) + where + section_ := F.map se.section_ + id' := by rw [← functor.map_comp, split_epi.id, Functor.map_id] +#align category_theory.split_epi.map CategoryTheory.SplitEpi.map + +instance {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] (F : C ⥤ D) : IsSplitMono (F.map f) := + IsSplitMono.mk' (hf.exists_splitMono.some.map F) + +instance {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] (F : C ⥤ D) : IsSplitEpi (F.map f) := + IsSplitEpi.mk' (hf.exists_splitEpi.some.map F) + +end + +end CategoryTheory + From 985720e0a47244e6a348049f8583ebc479cf58c0 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 20:02:31 -0500 Subject: [PATCH 052/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/EpiMono.lean | 7 +++---- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1fe25060bf045..fc9401c89f8c6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -225,6 +225,7 @@ import Mathlib.CategoryTheory.Category.KleisliCat import Mathlib.CategoryTheory.Category.RelCat import Mathlib.CategoryTheory.Comma import Mathlib.CategoryTheory.ConcreteCategory.Bundled +import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Equivalence import Mathlib.CategoryTheory.EssentialImage diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean index 4b39ca355dab1..8ded9add79687 100644 --- a/Mathlib/CategoryTheory/EpiMono.lean +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -8,8 +8,8 @@ Authors: Reid Barton, Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Opposites -import Mathbin.CategoryTheory.Groupoid +import Mathlib.CategoryTheory.Opposites +import Mathlib.CategoryTheory.Groupoid /-! # Facts about epimorphisms and monomorphisms. @@ -200,8 +200,7 @@ theorem IsIso.of_epi_section {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] [hf' : -- FIXME this has unnecessarily become noncomputable! /-- A category where every morphism has a `trunc` retraction is computably a groupoid. -/ noncomputable def Groupoid.ofTruncSplitMono - (all_split_mono : ∀ {X Y : C} (f : X ⟶ Y), Trunc (IsSplitMono f)) : Groupoid.{v₁} C := - by + (all_split_mono : ∀ {X Y : C} (f : X ⟶ Y), Trunc (IsSplitMono f)) : Groupoid.{v₁} C := by apply groupoid.of_is_iso intro X Y f trunc_cases all_split_mono f From e1d3cf34b386125ca57fbc8a6a939f5a51b4c3e5 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 20:57:45 -0500 Subject: [PATCH 053/126] fix errors; lint; add porting notes --- Mathlib/CategoryTheory/EpiMono.lean | 79 ++++++++++++++++------------- 1 file changed, 43 insertions(+), 36 deletions(-) diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean index 8ded9add79687..565bbbeb662d2 100644 --- a/Mathlib/CategoryTheory/EpiMono.lean +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -14,8 +14,8 @@ import Mathlib.CategoryTheory.Groupoid /-! # Facts about epimorphisms and monomorphisms. -The definitions of `epi` and `mono` are in `category_theory.category`, -since they are used by some lemmas for `iso`, which is used everywhere. +The definitions of `Epi` and `Mono` are in `CategoryTheory.Category`, +since they are used by some lemmas for `Iso`, which is used everywhere. -/ @@ -26,19 +26,19 @@ namespace CategoryTheory variable {C : Type u₁} [Category.{v₁} C] instance unop_mono_of_epi {A B : Cᵒᵖ} (f : A ⟶ B) [Epi f] : Mono f.unop := - ⟨fun Z g h eq => Quiver.Hom.op_inj ((cancel_epi f).1 (Quiver.Hom.unop_inj Eq))⟩ + ⟨fun _ _ eq => Quiver.Hom.op_inj ((cancel_epi f).1 (Quiver.Hom.unop_inj eq))⟩ #align category_theory.unop_mono_of_epi CategoryTheory.unop_mono_of_epi instance unop_epi_of_mono {A B : Cᵒᵖ} (f : A ⟶ B) [Mono f] : Epi f.unop := - ⟨fun Z g h eq => Quiver.Hom.op_inj ((cancel_mono f).1 (Quiver.Hom.unop_inj Eq))⟩ + ⟨fun _ _ eq => Quiver.Hom.op_inj ((cancel_mono f).1 (Quiver.Hom.unop_inj eq))⟩ #align category_theory.unop_epi_of_mono CategoryTheory.unop_epi_of_mono instance op_mono_of_epi {A B : C} (f : A ⟶ B) [Epi f] : Mono f.op := - ⟨fun Z g h eq => Quiver.Hom.unop_inj ((cancel_epi f).1 (Quiver.Hom.op_inj Eq))⟩ + ⟨fun _ _ eq => Quiver.Hom.unop_inj ((cancel_epi f).1 (Quiver.Hom.op_inj eq))⟩ #align category_theory.op_mono_of_epi CategoryTheory.op_mono_of_epi instance op_epi_of_mono {A B : C} (f : A ⟶ B) [Mono f] : Epi f.op := - ⟨fun Z g h eq => Quiver.Hom.unop_inj ((cancel_mono f).1 (Quiver.Hom.op_inj Eq))⟩ + ⟨fun _ _ eq => Quiver.Hom.unop_inj ((cancel_mono f).1 (Quiver.Hom.op_inj eq))⟩ #align category_theory.op_epi_of_mono CategoryTheory.op_epi_of_mono /-- A split monomorphism is a morphism `f : X ⟶ Y` with a given retraction `retraction f : Y ⟶ X` @@ -46,22 +46,25 @@ such that `f ≫ retraction f = 𝟙 X`. Every split monomorphism is a monomorphism. -/ -@[ext, nolint has_nonempty_instance] +/- Porting note: @[nolint has_nonempty_instance] -/ +/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ +@[aesop apply safe (rule_sets [CategoryTheory])] structure SplitMono {X Y : C} (f : X ⟶ Y) where + /-- The map splitting `f` -/ retraction : Y ⟶ X - id' : f ≫ retraction = 𝟙 X := by obviously + /-- `f` composed with `retraction` is the identity -/ + id : f ≫ retraction = 𝟙 X := by aesop_cat #align category_theory.split_mono CategoryTheory.SplitMono -restate_axiom split_mono.id' +attribute [reassoc (attr := simp)] SplitMono.id -attribute [simp, reassoc.1] split_mono.id - -/-- `is_split_mono f` is the assertion that `f` admits a retraction -/ +/-- `IsSplitMono f` is the assertion that `f` admits a retraction -/ class IsSplitMono {X Y : C} (f : X ⟶ Y) : Prop where + /-- There is a splitting -/ exists_splitMono : Nonempty (SplitMono f) #align category_theory.is_split_mono CategoryTheory.IsSplitMono -/-- A constructor for `is_split_mono f` taking a `split_mono f` as an argument -/ +/-- A constructor for `IsSplitMono f` taking a `SplitMono f` as an argument -/ theorem IsSplitMono.mk' {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : IsSplitMono f := ⟨Nonempty.intro sm⟩ #align category_theory.is_split_mono.mk' CategoryTheory.IsSplitMono.mk' @@ -72,22 +75,25 @@ such that `section_ f ≫ f = 𝟙 Y`. Every split epimorphism is an epimorphism. -/ -@[ext, nolint has_nonempty_instance] +/- Porting note: @[nolint has_nonempty_instance] -/ +/- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule -/ +@[aesop apply safe (rule_sets [CategoryTheory])] structure SplitEpi {X Y : C} (f : X ⟶ Y) where + /-- The map splitting `f` -/ section_ : Y ⟶ X - id' : section_ ≫ f = 𝟙 Y := by obviously + /-- `section_` composed with `f` is the identity -/ + id : section_ ≫ f = 𝟙 Y := by aesop_cat #align category_theory.split_epi CategoryTheory.SplitEpi -restate_axiom split_epi.id' - -attribute [simp, reassoc.1] split_epi.id +attribute [reassoc (attr := simp)] SplitEpi.id -/-- `is_split_epi f` is the assertion that `f` admits a section -/ +/-- `IsSplitEpi f` is the assertion that `f` admits a section -/ class IsSplitEpi {X Y : C} (f : X ⟶ Y) : Prop where + /-- There is a splitting -/ exists_splitEpi : Nonempty (SplitEpi f) #align category_theory.is_split_epi CategoryTheory.IsSplitEpi -/-- A constructor for `is_split_epi f` taking a `split_epi f` as an argument -/ +/-- A constructor for `IsSplitEpi f` taking a `SplitEpi f` as an argument -/ theorem IsSplitEpi.mk' {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : IsSplitEpi f := ⟨Nonempty.intro se⟩ #align category_theory.is_split_epi.mk' CategoryTheory.IsSplitEpi.mk' @@ -97,18 +103,18 @@ noncomputable def retraction {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : Y hf.exists_splitMono.some.retraction #align category_theory.retraction CategoryTheory.retraction -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem IsSplitMono.id {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : f ≫ retraction f = 𝟙 X := hf.exists_splitMono.some.id #align category_theory.is_split_mono.id CategoryTheory.IsSplitMono.id /-- The retraction of a split monomorphism has an obvious section. -/ -def SplitMono.splitEpi {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : SplitEpi sm.retraction +def SplitMono.splitEpi {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : SplitEpi sm.retraction where section_ := f #align category_theory.split_mono.split_epi CategoryTheory.SplitMono.splitEpi /-- The retraction of a split monomorphism is itself a split epimorphism. -/ -instance retraction_isSplitEpi {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : +instance retraction_isSplitEpi {X Y : C} (f : X ⟶ Y) [IsSplitMono f] : IsSplitEpi (retraction f) := IsSplitEpi.mk' (SplitMono.splitEpi _) #align category_theory.retraction_is_split_epi CategoryTheory.retraction_isSplitEpi @@ -125,7 +131,7 @@ noncomputable def section_ {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : Y ⟶ X hf.exists_splitEpi.some.section_ #align category_theory.section_ CategoryTheory.section_ -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem IsSplitEpi.id {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : section_ f ≫ f = 𝟙 Y := hf.exists_splitEpi.some.id #align category_theory.is_split_epi.id CategoryTheory.IsSplitEpi.id @@ -136,7 +142,7 @@ def SplitEpi.splitMono {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : SplitMono se. #align category_theory.split_epi.split_mono CategoryTheory.SplitEpi.splitMono /-- The section of a split epimorphism is itself a split monomorphism. -/ -instance section_isSplitMono {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : IsSplitMono (section_ f) := +instance section_isSplitMono {X Y : C} (f : X ⟶ Y) [IsSplitEpi f] : IsSplitMono (section_ f) := IsSplitMono.mk' (SplitEpi.splitMono _) #align category_theory.section_is_split_mono CategoryTheory.section_isSplitMono @@ -156,21 +162,21 @@ instance (priority := 100) IsSplitEpi.of_iso {X Y : C} (f : X ⟶ Y) [IsIso f] : #align category_theory.is_split_epi.of_iso CategoryTheory.IsSplitEpi.of_iso theorem SplitMono.mono {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) : Mono f := - { right_cancellation := fun Z g h w => by replace w := w =≫ sm.retraction; simpa using w } + { right_cancellation := fun g h w => by replace w := w =≫ sm.retraction; simpa using w } #align category_theory.split_mono.mono CategoryTheory.SplitMono.mono /-- Every split mono is a mono. -/ instance (priority := 100) IsSplitMono.mono {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] : Mono f := - hf.exists_splitMono.some.Mono + hf.exists_splitMono.some.mono #align category_theory.is_split_mono.mono CategoryTheory.IsSplitMono.mono theorem SplitEpi.epi {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) : Epi f := - { left_cancellation := fun Z g h w => by replace w := se.section_ ≫= w; simpa using w } + { left_cancellation := fun g h w => by replace w := se.section_ ≫= w; simpa using w } #align category_theory.split_epi.epi CategoryTheory.SplitEpi.epi /-- Every split epi is an epi. -/ instance (priority := 100) IsSplitEpi.epi {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] : Epi f := - hf.exists_splitEpi.some.Epi + hf.exists_splitEpi.some.epi #align category_theory.is_split_epi.epi CategoryTheory.IsSplitEpi.epi /-- Every split mono whose retraction is mono is an iso. -/ @@ -201,11 +207,11 @@ theorem IsIso.of_epi_section {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] [hf' : /-- A category where every morphism has a `trunc` retraction is computably a groupoid. -/ noncomputable def Groupoid.ofTruncSplitMono (all_split_mono : ∀ {X Y : C} (f : X ⟶ Y), Trunc (IsSplitMono f)) : Groupoid.{v₁} C := by - apply groupoid.of_is_iso + apply Groupoid.ofIsIso intro X Y f - trunc_cases all_split_mono f - trunc_cases all_split_mono (retraction f) - apply is_iso.of_mono_retraction + have ⟨a,_⟩:= Trunc.exists_rep <| all_split_mono f + have ⟨b,_⟩:= Trunc.exists_rep <| all_split_mono <| retraction f + apply IsIso.of_mono_retraction #align category_theory.groupoid.of_trunc_split_mono CategoryTheory.Groupoid.ofTruncSplitMono section @@ -214,11 +220,13 @@ variable (C) /-- A split mono category is a category in which every monomorphism is split. -/ class SplitMonoCategory where + /-- All monos are split -/ isSplitMono_of_mono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], IsSplitMono f #align category_theory.split_mono_category CategoryTheory.SplitMonoCategory /-- A split epi category is a category in which every epimorphism is split. -/ class SplitEpiCategory where + /-- All epis are split -/ isSplitEpi_of_epi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], IsSplitEpi f #align category_theory.split_epi_category CategoryTheory.SplitEpiCategory @@ -245,7 +253,7 @@ variable {D : Type u₂} [Category.{v₂} D] def SplitMono.map {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) (F : C ⥤ D) : SplitMono (F.map f) where retraction := F.map sm.retraction - id' := by rw [← functor.map_comp, split_mono.id, Functor.map_id] + id := by rw [← Functor.map_comp, SplitMono.id, Functor.map_id] #align category_theory.split_mono.map CategoryTheory.SplitMono.map /-- Split epimorphisms are also absolute epimorphisms. -/ @@ -253,7 +261,7 @@ def SplitMono.map {X Y : C} {f : X ⟶ Y} (sm : SplitMono f) (F : C ⥤ D) : Spl def SplitEpi.map {X Y : C} {f : X ⟶ Y} (se : SplitEpi f) (F : C ⥤ D) : SplitEpi (F.map f) where section_ := F.map se.section_ - id' := by rw [← functor.map_comp, split_epi.id, Functor.map_id] + id := by rw [← Functor.map_comp, SplitEpi.id, Functor.map_id] #align category_theory.split_epi.map CategoryTheory.SplitEpi.map instance {X Y : C} (f : X ⟶ Y) [hf : IsSplitMono f] (F : C ⥤ D) : IsSplitMono (F.map f) := @@ -265,4 +273,3 @@ instance {X Y : C} (f : X ⟶ Y) [hf : IsSplitEpi f] (F : C ⥤ D) : IsSplitEpi end end CategoryTheory - From 8c3a9c0724b27236fa3bdb41b132ac246195f16e Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 21:13:52 -0500 Subject: [PATCH 054/126] feat: port CategoryTheory.Types From e25276059ba4479b7ff463e5a3db3da6a803e2ed Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 21:13:52 -0500 Subject: [PATCH 055/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Types.lean | 405 ++++++++++++++++++++++++++++++ 1 file changed, 405 insertions(+) create mode 100644 Mathlib/CategoryTheory/Types.lean diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean new file mode 100644 index 0000000000000..a09c499a70887 --- /dev/null +++ b/Mathlib/CategoryTheory/Types.lean @@ -0,0 +1,405 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Stephen Morgan, Scott Morrison, Johannes Hölzl + +! This file was ported from Lean 3 source module category_theory.types +! leanprover-community/mathlib commit 48085f140e684306f9e7da907cd5932056d1aded +! 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.Functor.FullyFaithful +import Mathbin.Logic.Equiv.Basic + +/-! +# The category `Type`. + +In this section we set up the theory so that Lean's types and functions between them +can be viewed as a `large_category` in our framework. + +Lean can not transparently view a function as a morphism in this category, and needs a hint in +order to be able to type check. We provide the abbreviation `as_hom f` to guide type checking, +as well as a corresponding notation `↾ f`. (Entered as `\upr `.) The notation is enabled using +`open_locale category_theory.Type`. + +We provide various simplification lemmas for functors and natural transformations valued in `Type`. + +We define `ulift_functor`, from `Type u` to `Type (max u v)`, and show that it is fully faithful +(but not, of course, essentially surjective). + +We prove some basic facts about the category `Type`: +* epimorphisms are surjections and monomorphisms are injections, +* `iso` is both `iso` and `equiv` to `equiv` (at least within a fixed universe), +* every type level `is_lawful_functor` gives a categorical functor `Type ⥤ Type` + (the corresponding fact about monads is in `src/category_theory/monad/types.lean`). +-/ + + +namespace CategoryTheory + +-- morphism levels before object levels. See note [category_theory universes]. +universe v v' w u u' + +/- The `@[to_additive]` attribute is just a hint that expressions involving this instance can + still be additivized. -/ +@[to_additive CategoryTheory.types] +instance types : LargeCategory (Type u) + where + Hom a b := a → b + id a := id + comp _ _ _ f g := g ∘ f +#align category_theory.types CategoryTheory.types +#align category_theory.types CategoryTheory.types + +theorem types_hom {α β : Type u} : (α ⟶ β) = (α → β) := + rfl +#align category_theory.types_hom CategoryTheory.types_hom + +theorem types_id (X : Type u) : 𝟙 X = id := + rfl +#align category_theory.types_id CategoryTheory.types_id + +theorem types_comp {X Y Z : Type u} (f : X ⟶ Y) (g : Y ⟶ Z) : f ≫ g = g ∘ f := + rfl +#align category_theory.types_comp CategoryTheory.types_comp + +@[simp] +theorem types_id_apply (X : Type u) (x : X) : (𝟙 X : X → X) x = x := + rfl +#align category_theory.types_id_apply CategoryTheory.types_id_apply + +@[simp] +theorem types_comp_apply {X Y Z : Type u} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := + rfl +#align category_theory.types_comp_apply CategoryTheory.types_comp_apply + +@[simp] +theorem hom_inv_id_apply {X Y : Type u} (f : X ≅ Y) (x : X) : f.inv (f.Hom x) = x := + congr_fun f.hom_inv_id x +#align category_theory.hom_inv_id_apply CategoryTheory.hom_inv_id_apply + +@[simp] +theorem inv_hom_id_apply {X Y : Type u} (f : X ≅ Y) (y : Y) : f.Hom (f.inv y) = y := + congr_fun f.inv_hom_id y +#align category_theory.inv_hom_id_apply CategoryTheory.inv_hom_id_apply + +-- Unfortunately without this wrapper we can't use `category_theory` idioms, such as `is_iso f`. +/-- `as_hom f` helps Lean type check a function as a morphism in the category `Type`. -/ +abbrev asHom {α β : Type u} (f : α → β) : α ⟶ β := + f +#align category_theory.as_hom CategoryTheory.asHom + +-- mathport name: category_theory.as_hom +-- If you don't mind some notation you can use fewer keystrokes: +scoped[CategoryTheory.Type] notation "↾" f:200 => CategoryTheory.asHom f + +-- type as \upr in VScode +section + +-- We verify the expected type checking behaviour of `as_hom`. +variable (α β γ : Type u) (f : α → β) (g : β → γ) + +example : α → γ := + ↾f ≫ ↾g + +example [IsIso (↾f)] : Mono (↾f) := by infer_instance + +example [IsIso (↾f)] : ↾f ≫ inv (↾f) = 𝟙 α := by simp + +end + +namespace Functor + +variable {J : Type u} [Category.{v} J] + +/-- The sections of a functor `J ⥤ Type` are +the choices of a point `u j : F.obj j` for each `j`, +such that `F.map f (u j) = u j` for every morphism `f : j ⟶ j'`. + +We later use these to define limits in `Type` and in many concrete categories. +-/ +def sections (F : J ⥤ Type w) : Set (∀ j, F.obj j) := + { u | ∀ {j j'} (f : j ⟶ j'), F.map f (u j) = u j' } +#align category_theory.functor.sections CategoryTheory.Functor.sections + +end Functor + +namespace FunctorToTypes + +variable {C : Type u} [Category.{v} C] (F G H : C ⥤ Type w) {X Y Z : C} + +variable (σ : F ⟶ G) (τ : G ⟶ H) + +@[simp] +theorem map_comp_apply (f : X ⟶ Y) (g : Y ⟶ Z) (a : F.obj X) : + (F.map (f ≫ g)) a = (F.map g) ((F.map f) a) := by simp [types_comp] +#align category_theory.functor_to_types.map_comp_apply CategoryTheory.FunctorToTypes.map_comp_apply + +@[simp] +theorem map_id_apply (a : F.obj X) : (F.map (𝟙 X)) a = a := by simp [types_id] +#align category_theory.functor_to_types.map_id_apply CategoryTheory.FunctorToTypes.map_id_apply + +theorem naturality (f : X ⟶ Y) (x : F.obj X) : σ.app Y ((F.map f) x) = (G.map f) (σ.app X x) := + congr_fun (σ.naturality f) x +#align category_theory.functor_to_types.naturality CategoryTheory.FunctorToTypes.naturality + +@[simp] +theorem comp (x : F.obj X) : (σ ≫ τ).app X x = τ.app X (σ.app X x) := + rfl +#align category_theory.functor_to_types.comp CategoryTheory.FunctorToTypes.comp + +variable {D : Type u'} [𝒟 : Category.{u'} D] (I J : D ⥤ C) (ρ : I ⟶ J) {W : D} + +@[simp] +theorem hcomp (x : (I ⋙ F).obj W) : (ρ ◫ σ).app W x = (G.map (ρ.app W)) (σ.app (I.obj W) x) := + rfl +#align category_theory.functor_to_types.hcomp CategoryTheory.FunctorToTypes.hcomp + +@[simp] +theorem map_inv_map_hom_apply (f : X ≅ Y) (x : F.obj X) : F.map f.inv (F.map f.Hom x) = x := + congr_fun (F.mapIso f).hom_inv_id x +#align category_theory.functor_to_types.map_inv_map_hom_apply CategoryTheory.FunctorToTypes.map_inv_map_hom_apply + +@[simp] +theorem map_hom_map_inv_apply (f : X ≅ Y) (y : F.obj Y) : F.map f.Hom (F.map f.inv y) = y := + congr_fun (F.mapIso f).inv_hom_id y +#align category_theory.functor_to_types.map_hom_map_inv_apply CategoryTheory.FunctorToTypes.map_hom_map_inv_apply + +@[simp] +theorem hom_inv_id_app_apply (α : F ≅ G) (X) (x) : α.inv.app X (α.Hom.app X x) = x := + congr_fun (α.hom_inv_id_app X) x +#align category_theory.functor_to_types.hom_inv_id_app_apply CategoryTheory.FunctorToTypes.hom_inv_id_app_apply + +@[simp] +theorem inv_hom_id_app_apply (α : F ≅ G) (X) (x) : α.Hom.app X (α.inv.app X x) = x := + congr_fun (α.inv_hom_id_app X) x +#align category_theory.functor_to_types.inv_hom_id_app_apply CategoryTheory.FunctorToTypes.inv_hom_id_app_apply + +end FunctorToTypes + +/-- The isomorphism between a `Type` which has been `ulift`ed to the same universe, +and the original type. +-/ +def uliftTrivial (V : Type u) : ULift.{u} V ≅ V := by tidy +#align category_theory.ulift_trivial CategoryTheory.uliftTrivial + +/-- The functor embedding `Type u` into `Type (max u v)`. +Write this as `ulift_functor.{5 2}` to get `Type 2 ⥤ Type 5`. +-/ +def uliftFunctor : Type u ⥤ Type max u v + where + obj X := ULift.{v} X + map X Y f := fun x : ULift.{v} X => ULift.up (f x.down) +#align category_theory.ulift_functor CategoryTheory.uliftFunctor + +@[simp] +theorem uliftFunctor_map {X Y : Type u} (f : X ⟶ Y) (x : ULift.{v} X) : + uliftFunctor.map f x = ULift.up (f x.down) := + rfl +#align category_theory.ulift_functor_map CategoryTheory.uliftFunctor_map + +instance uliftFunctorFull : Full.{u} uliftFunctor where preimage X Y f x := (f (ULift.up x)).down +#align category_theory.ulift_functor_full CategoryTheory.uliftFunctorFull + +instance uliftFunctor_faithful : Faithful uliftFunctor + where map_injective' X Y f g p := + funext fun x => + congr_arg ULift.down (congr_fun p (ULift.up x) : ULift.up (f x) = ULift.up (g x)) +#align category_theory.ulift_functor_faithful CategoryTheory.uliftFunctor_faithful + +/-- The functor embedding `Type u` into `Type u` via `ulift` is isomorphic to the identity functor. + -/ +def uliftFunctorTrivial : uliftFunctor.{u, u} ≅ 𝟭 _ := + NatIso.ofComponents uliftTrivial (by tidy) +#align category_theory.ulift_functor_trivial CategoryTheory.uliftFunctorTrivial + +-- TODO We should connect this to a general story about concrete categories +-- whose forgetful functor is representable. +/-- Any term `x` of a type `X` corresponds to a morphism `punit ⟶ X`. -/ +def homOfElement {X : Type u} (x : X) : PUnit ⟶ X := fun _ => x +#align category_theory.hom_of_element CategoryTheory.homOfElement + +theorem homOfElement_eq_iff {X : Type u} (x y : X) : homOfElement x = homOfElement y ↔ x = y := + ⟨fun H => congr_fun H PUnit.unit, by cc⟩ +#align category_theory.hom_of_element_eq_iff CategoryTheory.homOfElement_eq_iff + +/-- A morphism in `Type` is a monomorphism if and only if it is injective. + +See . +-/ +theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.Injective f := + by + constructor + · intro H x x' h + skip + rw [← hom_of_element_eq_iff] at h⊢ + exact (cancel_mono f).mp h + · exact fun H => ⟨fun Z => H.compLeft⟩ +#align category_theory.mono_iff_injective CategoryTheory.mono_iff_injective + +theorem injective_of_mono {X Y : Type u} (f : X ⟶ Y) [hf : Mono f] : Function.Injective f := + (mono_iff_injective f).1 hf +#align category_theory.injective_of_mono CategoryTheory.injective_of_mono + +/-- A morphism in `Type` is an epimorphism if and only if it is surjective. + +See . +-/ +theorem epi_iff_surjective {X Y : Type u} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := + by + constructor + · rintro ⟨H⟩ + refine' Function.surjective_of_right_cancellable_Prop fun g₁ g₂ hg => _ + rw [← equiv.ulift.symm.injective.comp_left.eq_iff] + apply H + change ULift.up ∘ g₁ ∘ f = ULift.up ∘ g₂ ∘ f + rw [hg] + · exact fun H => ⟨fun Z => H.injective_comp_right⟩ +#align category_theory.epi_iff_surjective CategoryTheory.epi_iff_surjective + +theorem surjective_of_epi {X Y : Type u} (f : X ⟶ Y) [hf : Epi f] : Function.Surjective f := + (epi_iff_surjective f).1 hf +#align category_theory.surjective_of_epi CategoryTheory.surjective_of_epi + +section + +/-- `of_type_functor m` converts from Lean's `Type`-based `category` to `category_theory`. This +allows us to use these functors in category theory. -/ +def ofTypeFunctor (m : Type u → Type v) [Functor m] [LawfulFunctor m] : Type u ⥤ Type v + where + obj := m + map α β := Functor.map + map_id' α := Functor.map_id + map_comp' α β γ f g := funext fun a => LawfulFunctor.comp_map f g _ +#align category_theory.of_type_functor CategoryTheory.ofTypeFunctor + +variable (m : Type u → Type v) [Functor m] [LawfulFunctor m] + +@[simp] +theorem ofTypeFunctor_obj : (ofTypeFunctor m).obj = m := + rfl +#align category_theory.of_type_functor_obj CategoryTheory.ofTypeFunctor_obj + +@[simp] +theorem ofTypeFunctor_map {α β} (f : α → β) : + (ofTypeFunctor m).map f = (Functor.map f : m α → m β) := + rfl +#align category_theory.of_type_functor_map CategoryTheory.ofTypeFunctor_map + +end + +end CategoryTheory + +-- Isomorphisms in Type and equivalences. +namespace Equiv + +universe u + +variable {X Y : Type u} + +/-- Any equivalence between types in the same universe gives +a categorical isomorphism between those types. +-/ +def toIso (e : X ≃ Y) : X ≅ Y where + Hom := e.toFun + inv := e.invFun + hom_inv_id' := funext e.left_inv + inv_hom_id' := funext e.right_inv +#align equiv.to_iso Equiv.toIso + +@[simp] +theorem toIso_hom {e : X ≃ Y} : e.toIso.Hom = e := + rfl +#align equiv.to_iso_hom Equiv.toIso_hom + +@[simp] +theorem toIso_inv {e : X ≃ Y} : e.toIso.inv = e.symm := + rfl +#align equiv.to_iso_inv Equiv.toIso_inv + +end Equiv + +universe u + +namespace CategoryTheory.Iso + +open CategoryTheory + +variable {X Y : Type u} + +/-- Any isomorphism between types gives an equivalence. +-/ +def toEquiv (i : X ≅ Y) : X ≃ Y where + toFun := i.Hom + invFun := i.inv + left_inv x := congr_fun i.hom_inv_id x + right_inv y := congr_fun i.inv_hom_id y +#align category_theory.iso.to_equiv CategoryTheory.Iso.toEquiv + +@[simp] +theorem toEquiv_fun (i : X ≅ Y) : (i.toEquiv : X → Y) = i.Hom := + rfl +#align category_theory.iso.to_equiv_fun CategoryTheory.Iso.toEquiv_fun + +@[simp] +theorem toEquiv_symm_fun (i : X ≅ Y) : (i.toEquiv.symm : Y → X) = i.inv := + rfl +#align category_theory.iso.to_equiv_symm_fun CategoryTheory.Iso.toEquiv_symm_fun + +@[simp] +theorem toEquiv_id (X : Type u) : (Iso.refl X).toEquiv = Equiv.refl X := + rfl +#align category_theory.iso.to_equiv_id CategoryTheory.Iso.toEquiv_id + +@[simp] +theorem toEquiv_comp {X Y Z : Type u} (f : X ≅ Y) (g : Y ≅ Z) : + (f ≪≫ g).toEquiv = f.toEquiv.trans g.toEquiv := + rfl +#align category_theory.iso.to_equiv_comp CategoryTheory.Iso.toEquiv_comp + +end CategoryTheory.Iso + +namespace CategoryTheory + +/-- A morphism in `Type u` is an isomorphism if and only if it is bijective. -/ +theorem isIso_iff_bijective {X Y : Type u} (f : X ⟶ Y) : IsIso f ↔ Function.Bijective f := + Iff.intro (fun i => (as_iso f : X ≅ Y).toEquiv.Bijective) fun b => + IsIso.of_iso (Equiv.ofBijective f b).toIso +#align category_theory.is_iso_iff_bijective CategoryTheory.isIso_iff_bijective + +instance : SplitEpiCategory (Type u) + where isSplitEpi_of_epi X Y f hf := + IsSplitEpi.mk' + { section_ := Function.surjInv <| (epi_iff_surjective f).1 hf + id' := funext <| Function.rightInverse_surjInv <| (epi_iff_surjective f).1 hf } + +end CategoryTheory + +-- We prove `equiv_iso_iso` and then use that to sneakily construct `equiv_equiv_iso`. +-- (In this order the proofs are handled by `obviously`.) +/-- Equivalences (between types in the same universe) are the same as (isomorphic to) isomorphisms +of types. -/ +@[simps] +def equivIsoIso {X Y : Type u} : X ≃ Y ≅ X ≅ Y + where + Hom e := e.toIso + inv i := i.toEquiv +#align equiv_iso_iso equivIsoIso + +/-- Equivalences (between types in the same universe) are the same as (equivalent to) isomorphisms +of types. -/ +def equivEquivIso {X Y : Type u} : X ≃ Y ≃ (X ≅ Y) := + equivIsoIso.toEquiv +#align equiv_equiv_iso equivEquivIso + +@[simp] +theorem equivEquivIso_hom {X Y : Type u} (e : X ≃ Y) : equivEquivIso e = e.toIso := + rfl +#align equiv_equiv_iso_hom equivEquivIso_hom + +@[simp] +theorem equivEquivIso_inv {X Y : Type u} (e : X ≅ Y) : equivEquivIso.symm e = e.toEquiv := + rfl +#align equiv_equiv_iso_inv equivEquivIso_inv + From b6d5b6ed4538f6b81b1a24460e3d045d6998c7f4 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 21:13:52 -0500 Subject: [PATCH 056/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Types.lean | 12 +++++------- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index a82728786eb75..ee07708944f1f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -245,6 +245,7 @@ import Mathlib.CategoryTheory.Products.Basic import Mathlib.CategoryTheory.Products.Bifunctor import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.Thin +import Mathlib.CategoryTheory.Types import Mathlib.CategoryTheory.Whiskering import Mathlib.Combinatorics.Additive.Energy import Mathlib.Combinatorics.Additive.RuzsaCovering diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean index a09c499a70887..1a924e24806ef 100644 --- a/Mathlib/CategoryTheory/Types.lean +++ b/Mathlib/CategoryTheory/Types.lean @@ -8,9 +8,9 @@ Authors: Stephen Morgan, Scott Morrison, Johannes Hölzl ! 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.Functor.FullyFaithful -import Mathbin.Logic.Equiv.Basic +import Mathlib.CategoryTheory.EpiMono +import Mathlib.CategoryTheory.Functor.FullyFaithful +import Mathlib.Logic.Equiv.Basic /-! # The category `Type`. @@ -228,8 +228,7 @@ theorem homOfElement_eq_iff {X : Type u} (x y : X) : homOfElement x = homOfEleme See . -/ -theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.Injective f := - by +theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.Injective f := by constructor · intro H x x' h skip @@ -246,8 +245,7 @@ theorem injective_of_mono {X Y : Type u} (f : X ⟶ Y) [hf : Mono f] : Function. See . -/ -theorem epi_iff_surjective {X Y : Type u} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := - by +theorem epi_iff_surjective {X Y : Type u} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := by constructor · rintro ⟨H⟩ refine' Function.surjective_of_right_cancellable_Prop fun g₁ g₂ hg => _ From 2b370d2ab4d5b6d9da8cb7eb43203e3acc8f1601 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:32:42 -0500 Subject: [PATCH 057/126] fix errors; lint; add porting note --- Mathlib/CategoryTheory/Types.lean | 106 +++++++++++++++--------------- 1 file changed, 54 insertions(+), 52 deletions(-) diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean index 1a924e24806ef..63f7c04605930 100644 --- a/Mathlib/CategoryTheory/Types.lean +++ b/Mathlib/CategoryTheory/Types.lean @@ -16,23 +16,22 @@ import Mathlib.Logic.Equiv.Basic # The category `Type`. In this section we set up the theory so that Lean's types and functions between them -can be viewed as a `large_category` in our framework. +can be viewed as a `LargeCategory` in our framework. Lean can not transparently view a function as a morphism in this category, and needs a hint in -order to be able to type check. We provide the abbreviation `as_hom f` to guide type checking, -as well as a corresponding notation `↾ f`. (Entered as `\upr `.) The notation is enabled using -`open_locale category_theory.Type`. +order to be able to type check. We provide the abbreviation `asHom f` to guide type checking, +as well as a corresponding notation `↾ f`. (Entered as `\upr `.) We provide various simplification lemmas for functors and natural transformations valued in `Type`. -We define `ulift_functor`, from `Type u` to `Type (max u v)`, and show that it is fully faithful +We define `uliftFunctor`, from `Type u` to `Type (max u v)`, and show that it is fully faithful (but not, of course, essentially surjective). We prove some basic facts about the category `Type`: * epimorphisms are surjections and monomorphisms are injections, -* `iso` is both `iso` and `equiv` to `equiv` (at least within a fixed universe), -* every type level `is_lawful_functor` gives a categorical functor `Type ⥤ Type` - (the corresponding fact about monads is in `src/category_theory/monad/types.lean`). +* `Iso` is both `Iso` and `Equiv` to `Equiv` (at least within a fixed universe), +* every type level `IsLawfulFunctor` gives a categorical functor `Type ⥤ Type` + (the corresponding fact about monads is in `Mathlib/CategoryTheory/Monad/Types.lean`). -/ @@ -48,8 +47,7 @@ instance types : LargeCategory (Type u) where Hom a b := a → b id a := id - comp _ _ _ f g := g ∘ f -#align category_theory.types CategoryTheory.types + comp f g := g ∘ f #align category_theory.types CategoryTheory.types theorem types_hom {α β : Type u} : (α ⟶ β) = (α → β) := @@ -75,29 +73,28 @@ theorem types_comp_apply {X Y Z : Type u} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : #align category_theory.types_comp_apply CategoryTheory.types_comp_apply @[simp] -theorem hom_inv_id_apply {X Y : Type u} (f : X ≅ Y) (x : X) : f.inv (f.Hom x) = x := +theorem hom_inv_id_apply {X Y : Type u} (f : X ≅ Y) (x : X) : f.inv (f.hom x) = x := congr_fun f.hom_inv_id x #align category_theory.hom_inv_id_apply CategoryTheory.hom_inv_id_apply @[simp] -theorem inv_hom_id_apply {X Y : Type u} (f : X ≅ Y) (y : Y) : f.Hom (f.inv y) = y := +theorem inv_hom_id_apply {X Y : Type u} (f : X ≅ Y) (y : Y) : f.hom (f.inv y) = y := congr_fun f.inv_hom_id y #align category_theory.inv_hom_id_apply CategoryTheory.inv_hom_id_apply --- Unfortunately without this wrapper we can't use `category_theory` idioms, such as `is_iso f`. -/-- `as_hom f` helps Lean type check a function as a morphism in the category `Type`. -/ +-- Unfortunately without this wrapper we can't use `CategoryTheory` idioms, such as `IsIso f`. +/-- `asHom f` helps Lean type check a function as a morphism in the category `Type`. -/ abbrev asHom {α β : Type u} (f : α → β) : α ⟶ β := f #align category_theory.as_hom CategoryTheory.asHom --- mathport name: category_theory.as_hom -- If you don't mind some notation you can use fewer keystrokes: -scoped[CategoryTheory.Type] notation "↾" f:200 => CategoryTheory.asHom f +/-- Type as \upr in VScode -/ +scoped notation "↾" f:200 => CategoryTheory.asHom f --- type as \upr in VScode section --- We verify the expected type checking behaviour of `as_hom`. +-- We verify the expected type checking behaviour of `asHom` variable (α β γ : Type u) (f : α → β) (g : β → γ) example : α → γ := @@ -157,22 +154,22 @@ theorem hcomp (x : (I ⋙ F).obj W) : (ρ ◫ σ).app W x = (G.map (ρ.app W)) ( #align category_theory.functor_to_types.hcomp CategoryTheory.FunctorToTypes.hcomp @[simp] -theorem map_inv_map_hom_apply (f : X ≅ Y) (x : F.obj X) : F.map f.inv (F.map f.Hom x) = x := +theorem map_inv_map_hom_apply (f : X ≅ Y) (x : F.obj X) : F.map f.inv (F.map f.hom x) = x := congr_fun (F.mapIso f).hom_inv_id x #align category_theory.functor_to_types.map_inv_map_hom_apply CategoryTheory.FunctorToTypes.map_inv_map_hom_apply @[simp] -theorem map_hom_map_inv_apply (f : X ≅ Y) (y : F.obj Y) : F.map f.Hom (F.map f.inv y) = y := +theorem map_hom_map_inv_apply (f : X ≅ Y) (y : F.obj Y) : F.map f.hom (F.map f.inv y) = y := congr_fun (F.mapIso f).inv_hom_id y #align category_theory.functor_to_types.map_hom_map_inv_apply CategoryTheory.FunctorToTypes.map_hom_map_inv_apply @[simp] -theorem hom_inv_id_app_apply (α : F ≅ G) (X) (x) : α.inv.app X (α.Hom.app X x) = x := +theorem hom_inv_id_app_apply (α : F ≅ G) (X) (x) : α.inv.app X (α.hom.app X x) = x := congr_fun (α.hom_inv_id_app X) x #align category_theory.functor_to_types.hom_inv_id_app_apply CategoryTheory.FunctorToTypes.hom_inv_id_app_apply @[simp] -theorem inv_hom_id_app_apply (α : F ≅ G) (X) (x) : α.Hom.app X (α.inv.app X x) = x := +theorem inv_hom_id_app_apply (α : F ≅ G) (X) (x) : α.hom.app X (α.inv.app X x) = x := congr_fun (α.inv_hom_id_app X) x #align category_theory.functor_to_types.inv_hom_id_app_apply CategoryTheory.FunctorToTypes.inv_hom_id_app_apply @@ -181,16 +178,20 @@ end FunctorToTypes /-- The isomorphism between a `Type` which has been `ulift`ed to the same universe, and the original type. -/ -def uliftTrivial (V : Type u) : ULift.{u} V ≅ V := by tidy +def uliftTrivial (V : Type u) : ULift.{u} V ≅ V where + hom a := a.1 + inv a := .up a + hom_inv_id := by aesop_cat + inv_hom_id := by aesop_cat #align category_theory.ulift_trivial CategoryTheory.uliftTrivial /-- The functor embedding `Type u` into `Type (max u v)`. -Write this as `ulift_functor.{5 2}` to get `Type 2 ⥤ Type 5`. +Write this as `uliftFunctor.{5 2}` to get `Type 2 ⥤ Type 5`. -/ def uliftFunctor : Type u ⥤ Type max u v where obj X := ULift.{v} X - map X Y f := fun x : ULift.{v} X => ULift.up (f x.down) + map {X} {Y} f := fun x : ULift.{v} X => ULift.up (f x.down) #align category_theory.ulift_functor CategoryTheory.uliftFunctor @[simp] @@ -199,11 +200,11 @@ theorem uliftFunctor_map {X Y : Type u} (f : X ⟶ Y) (x : ULift.{v} X) : rfl #align category_theory.ulift_functor_map CategoryTheory.uliftFunctor_map -instance uliftFunctorFull : Full.{u} uliftFunctor where preimage X Y f x := (f (ULift.up x)).down +instance uliftFunctorFull : Full.{u} uliftFunctor where preimage f x := (f (ULift.up x)).down #align category_theory.ulift_functor_full CategoryTheory.uliftFunctorFull instance uliftFunctor_faithful : Faithful uliftFunctor - where map_injective' X Y f g p := + where map_injective {_X} {_Y} f g p := funext fun x => congr_arg ULift.down (congr_fun p (ULift.up x) : ULift.up (f x) = ULift.up (g x)) #align category_theory.ulift_functor_faithful CategoryTheory.uliftFunctor_faithful @@ -211,7 +212,7 @@ instance uliftFunctor_faithful : Faithful uliftFunctor /-- The functor embedding `Type u` into `Type u` via `ulift` is isomorphic to the identity functor. -/ def uliftFunctorTrivial : uliftFunctor.{u, u} ≅ 𝟭 _ := - NatIso.ofComponents uliftTrivial (by tidy) + NatIso.ofComponents uliftTrivial (by aesop_cat) #align category_theory.ulift_functor_trivial CategoryTheory.uliftFunctorTrivial -- TODO We should connect this to a general story about concrete categories @@ -221,7 +222,7 @@ def homOfElement {X : Type u} (x : X) : PUnit ⟶ X := fun _ => x #align category_theory.hom_of_element CategoryTheory.homOfElement theorem homOfElement_eq_iff {X : Type u} (x y : X) : homOfElement x = homOfElement y ↔ x = y := - ⟨fun H => congr_fun H PUnit.unit, by cc⟩ + ⟨fun H => congr_fun H PUnit.unit, by aesop⟩ #align category_theory.hom_of_element_eq_iff CategoryTheory.homOfElement_eq_iff /-- A morphism in `Type` is a monomorphism if and only if it is injective. @@ -232,9 +233,9 @@ theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.In constructor · intro H x x' h skip - rw [← hom_of_element_eq_iff] at h⊢ + rw [← homOfElement_eq_iff] at h⊢ exact (cancel_mono f).mp h - · exact fun H => ⟨fun Z => H.compLeft⟩ + · exact fun H => ⟨fun g g' h => H.comp_left h⟩ #align category_theory.mono_iff_injective CategoryTheory.mono_iff_injective theorem injective_of_mono {X Y : Type u} (f : X ⟶ Y) [hf : Mono f] : Function.Injective f := @@ -249,11 +250,11 @@ theorem epi_iff_surjective {X Y : Type u} (f : X ⟶ Y) : Epi f ↔ Function.Sur constructor · rintro ⟨H⟩ refine' Function.surjective_of_right_cancellable_Prop fun g₁ g₂ hg => _ - rw [← equiv.ulift.symm.injective.comp_left.eq_iff] + rw [← Equiv.ulift.symm.injective.comp_left.eq_iff] apply H change ULift.up ∘ g₁ ∘ f = ULift.up ∘ g₂ ∘ f rw [hg] - · exact fun H => ⟨fun Z => H.injective_comp_right⟩ + · exact fun H => ⟨fun g g' h => H.injective_comp_right h⟩ #align category_theory.epi_iff_surjective CategoryTheory.epi_iff_surjective theorem surjective_of_epi {X Y : Type u} (f : X ⟶ Y) [hf : Epi f] : Function.Surjective f := @@ -262,17 +263,19 @@ theorem surjective_of_epi {X Y : Type u} (f : X ⟶ Y) [hf : Epi f] : Function.S section -/-- `of_type_functor m` converts from Lean's `Type`-based `category` to `category_theory`. This +/-- `ofTypeFunctor m` converts from Lean's `Type`-based `Category` to `CategoryTheory`. This allows us to use these functors in category theory. -/ -def ofTypeFunctor (m : Type u → Type v) [Functor m] [LawfulFunctor m] : Type u ⥤ Type v +def ofTypeFunctor (m : Type u → Type v) [_root_.Functor m] [LawfulFunctor m] : Type u ⥤ Type v where obj := m - map α β := Functor.map - map_id' α := Functor.map_id - map_comp' α β γ f g := funext fun a => LawfulFunctor.comp_map f g _ + map f := Functor.map f + map_id := fun α => by funext X; apply id_map /- Porting note: original proof is via + `fun α => _root_.Functor.map_id` but I cannot get Lean to find this. Reproduced its + original proof -/ + map_comp f g := funext fun a => LawfulFunctor.comp_map f g _ #align category_theory.of_type_functor CategoryTheory.ofTypeFunctor -variable (m : Type u → Type v) [Functor m] [LawfulFunctor m] +variable (m : Type u → Type v) [_root_.Functor m] [LawfulFunctor m] @[simp] theorem ofTypeFunctor_obj : (ofTypeFunctor m).obj = m := @@ -300,14 +303,14 @@ variable {X Y : Type u} a categorical isomorphism between those types. -/ def toIso (e : X ≃ Y) : X ≅ Y where - Hom := e.toFun + hom := e.toFun inv := e.invFun - hom_inv_id' := funext e.left_inv - inv_hom_id' := funext e.right_inv + hom_inv_id := funext e.left_inv + inv_hom_id := funext e.right_inv #align equiv.to_iso Equiv.toIso @[simp] -theorem toIso_hom {e : X ≃ Y} : e.toIso.Hom = e := +theorem toIso_hom {e : X ≃ Y} : e.toIso.hom = e := rfl #align equiv.to_iso_hom Equiv.toIso_hom @@ -329,14 +332,14 @@ variable {X Y : Type u} /-- Any isomorphism between types gives an equivalence. -/ def toEquiv (i : X ≅ Y) : X ≃ Y where - toFun := i.Hom + toFun := i.hom invFun := i.inv left_inv x := congr_fun i.hom_inv_id x right_inv y := congr_fun i.inv_hom_id y #align category_theory.iso.to_equiv CategoryTheory.Iso.toEquiv @[simp] -theorem toEquiv_fun (i : X ≅ Y) : (i.toEquiv : X → Y) = i.Hom := +theorem toEquiv_fun (i : X ≅ Y) : (i.toEquiv : X → Y) = i.hom := rfl #align category_theory.iso.to_equiv_fun CategoryTheory.Iso.toEquiv_fun @@ -362,26 +365,26 @@ namespace CategoryTheory /-- A morphism in `Type u` is an isomorphism if and only if it is bijective. -/ theorem isIso_iff_bijective {X Y : Type u} (f : X ⟶ Y) : IsIso f ↔ Function.Bijective f := - Iff.intro (fun i => (as_iso f : X ≅ Y).toEquiv.Bijective) fun b => + Iff.intro (fun _ => (asIso f : X ≅ Y).toEquiv.bijective) fun b => IsIso.of_iso (Equiv.ofBijective f b).toIso #align category_theory.is_iso_iff_bijective CategoryTheory.isIso_iff_bijective instance : SplitEpiCategory (Type u) - where isSplitEpi_of_epi X Y f hf := - IsSplitEpi.mk' + where isSplitEpi_of_epi f hf := + IsSplitEpi.mk' <| { section_ := Function.surjInv <| (epi_iff_surjective f).1 hf - id' := funext <| Function.rightInverse_surjInv <| (epi_iff_surjective f).1 hf } + id := funext <| Function.rightInverse_surjInv <| (epi_iff_surjective f).1 hf } end CategoryTheory --- We prove `equiv_iso_iso` and then use that to sneakily construct `equiv_equiv_iso`. +-- We prove `equivIsoIso` and then use that to sneakily construct `equivEquivIso`. -- (In this order the proofs are handled by `obviously`.) /-- Equivalences (between types in the same universe) are the same as (isomorphic to) isomorphisms of types. -/ @[simps] def equivIsoIso {X Y : Type u} : X ≃ Y ≅ X ≅ Y where - Hom e := e.toIso + hom e := e.toIso inv i := i.toEquiv #align equiv_iso_iso equivIsoIso @@ -400,4 +403,3 @@ theorem equivEquivIso_hom {X Y : Type u} (e : X ≃ Y) : equivEquivIso e = e.toI theorem equivEquivIso_inv {X Y : Type u} (e : X ≅ Y) : equivEquivIso.symm e = e.toEquiv := rfl #align equiv_equiv_iso_inv equivEquivIso_inv - From 7fe5c64ff3e38f6ac33c4a2577f5e45946be5413 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:41:16 -0500 Subject: [PATCH 058/126] feat: port CategoryTheory.Functor.Hom From d5ec61a73bba5f5e0180c17b677153d734f48b7f Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:41:16 -0500 Subject: [PATCH 059/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Functor/Hom.lean | 38 +++++++++++++++++++++++++ 1 file changed, 38 insertions(+) create mode 100644 Mathlib/CategoryTheory/Functor/Hom.lean diff --git a/Mathlib/CategoryTheory/Functor/Hom.lean b/Mathlib/CategoryTheory/Functor/Hom.lean new file mode 100644 index 0000000000000..53ea4854caf56 --- /dev/null +++ b/Mathlib/CategoryTheory/Functor/Hom.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2018 Reid Barton. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Reid Barton, Scott Morrison + +! This file was ported from Lean 3 source module category_theory.functor.hom +! leanprover-community/mathlib commit 369525b73f229ccd76a6ec0e0e0bf2be57599768 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Products.Basic +import Mathbin.CategoryTheory.Types + +/-! +The hom functor, sending `(X, Y)` to the type `X ⟶ Y`. +-/ + + +universe v u + +open Opposite + +open CategoryTheory + +namespace CategoryTheory.Functor + +variable (C : Type u) [Category.{v} C] + +/-- `functor.hom` is the hom-pairing, sending `(X, Y)` to `X ⟶ Y`, contravariant in `X` and +covariant in `Y`. -/ +@[simps] +def hom : Cᵒᵖ × C ⥤ Type v where + obj p := unop p.1 ⟶ p.2 + map X Y f h := f.1.unop ≫ h ≫ f.2 +#align category_theory.functor.hom CategoryTheory.Functor.hom + +end CategoryTheory.Functor + From b163ac1b57aeb7a2e3ca7ae1b18160e5d1adb736 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:41:16 -0500 Subject: [PATCH 060/126] 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/Hom.lean | 4 ++-- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index a82728786eb75..8b1a5e063e263 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -235,6 +235,7 @@ import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.Functor.Default import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial +import Mathlib.CategoryTheory.Functor.Hom import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.NatIso diff --git a/Mathlib/CategoryTheory/Functor/Hom.lean b/Mathlib/CategoryTheory/Functor/Hom.lean index 53ea4854caf56..3490c417c50c3 100644 --- a/Mathlib/CategoryTheory/Functor/Hom.lean +++ b/Mathlib/CategoryTheory/Functor/Hom.lean @@ -8,8 +8,8 @@ Authors: Reid Barton, Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Products.Basic -import Mathbin.CategoryTheory.Types +import Mathlib.CategoryTheory.Products.Basic +import Mathlib.CategoryTheory.Types /-! The hom functor, sending `(X, Y)` to the type `X ⟶ Y`. From 278e1103f53fa6da1a06fa9333c0092ed95618f1 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:50:46 -0500 Subject: [PATCH 061/126] fix error --- Mathlib/CategoryTheory/Functor/Hom.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Functor/Hom.lean b/Mathlib/CategoryTheory/Functor/Hom.lean index 3490c417c50c3..1db6fb49476f3 100644 --- a/Mathlib/CategoryTheory/Functor/Hom.lean +++ b/Mathlib/CategoryTheory/Functor/Hom.lean @@ -26,13 +26,14 @@ namespace CategoryTheory.Functor variable (C : Type u) [Category.{v} C] -/-- `functor.hom` is the hom-pairing, sending `(X, Y)` to `X ⟶ Y`, contravariant in `X` and +/-- `Functor.hom` is the hom-pairing, sending `(X, Y)` to `X ⟶ Y`, contravariant in `X` and covariant in `Y`. -/ @[simps] def hom : Cᵒᵖ × C ⥤ Type v where obj p := unop p.1 ⟶ p.2 - map X Y f h := f.1.unop ≫ h ≫ f.2 + map f h := f.1.unop ≫ h ≫ f.2 + map_id := by aesop_cat + map_comp := fun {X} {Y} {Z} f g => by funext b; simp #align category_theory.functor.hom CategoryTheory.Functor.hom end CategoryTheory.Functor - From 2e8b88faa086edb0f3b26c1ca2e0c7e7c32d3856 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:56:08 -0500 Subject: [PATCH 062/126] feat: port CategoryTheory.Functor.Currying From 2b3c437f0d93d63e52c27c162919ed29b4e0767a Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:56:08 -0500 Subject: [PATCH 063/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Functor/Currying.lean | 121 +++++++++++++++++++ 1 file changed, 121 insertions(+) create mode 100644 Mathlib/CategoryTheory/Functor/Currying.lean diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean new file mode 100644 index 0000000000000..ee6c5edabe051 --- /dev/null +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison + +! This file was ported from Lean 3 source module category_theory.functor.currying +! leanprover-community/mathlib commit 369525b73f229ccd76a6ec0e0e0bf2be57599768 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Products.Bifunctor + +/-! +# Curry and uncurry, as functors. + +We define `curry : ((C × D) ⥤ E) ⥤ (C ⥤ (D ⥤ E))` and `uncurry : (C ⥤ (D ⥤ E)) ⥤ ((C × D) ⥤ E)`, +and verify that they provide an equivalence of categories +`currying : (C ⥤ (D ⥤ E)) ≌ ((C × D) ⥤ E)`. + +-/ + + +namespace CategoryTheory + +universe v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ + +variable {B : Type u₁} [Category.{v₁} B] {C : Type u₂} [Category.{v₂} C] {D : Type u₃} + [Category.{v₃} D] {E : Type u₄} [Category.{v₄} E] + +/-- The uncurrying functor, taking a functor `C ⥤ (D ⥤ E)` and producing a functor `(C × D) ⥤ E`. +-/ +@[simps] +def uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E + where + obj F := + { obj := fun X => (F.obj X.1).obj X.2 + map := fun X Y f => (F.map f.1).app X.2 ≫ (F.obj Y.1).map f.2 + map_comp' := fun X Y Z f g => + by + simp only [prod_comp_fst, prod_comp_snd, functor.map_comp, nat_trans.comp_app, + category.assoc] + slice_lhs 2 3 => rw [← nat_trans.naturality] + rw [category.assoc] } + map F G T := + { app := fun X => (T.app X.1).app X.2 + naturality' := fun X Y f => + by + simp only [prod_comp_fst, prod_comp_snd, category.comp_id, category.assoc, Functor.map_id, + functor.map_comp, nat_trans.id_app, nat_trans.comp_app] + slice_lhs 2 3 => rw [nat_trans.naturality] + slice_lhs 1 2 => rw [← nat_trans.comp_app, nat_trans.naturality, nat_trans.comp_app] + rw [category.assoc] } +#align category_theory.uncurry CategoryTheory.uncurry + +/-- The object level part of the currying functor. (See `curry` for the functorial version.) +-/ +def curryObj (F : C × D ⥤ E) : C ⥤ D ⥤ E + where + obj X := + { obj := fun Y => F.obj (X, Y) + map := fun Y Y' g => F.map (𝟙 X, g) } + map X X' f := { app := fun Y => F.map (f, 𝟙 Y) } +#align category_theory.curry_obj CategoryTheory.curryObj + +/-- The currying functor, taking a functor `(C × D) ⥤ E` and producing a functor `C ⥤ (D ⥤ E)`. +-/ +@[simps obj_obj_obj obj_obj_map obj_map_app map_app_app] +def curry : (C × D ⥤ E) ⥤ C ⥤ D ⥤ E where + obj F := curryObj F + map F G T := + { app := fun X => + { app := fun Y => T.app (X, Y) + naturality' := fun Y Y' g => by + dsimp [curry_obj] + rw [nat_trans.naturality] } + naturality' := fun X X' f => by + ext; dsimp [curry_obj] + rw [nat_trans.naturality] } +#align category_theory.curry CategoryTheory.curry + +-- create projection simp lemmas even though this isn't a `{ .. }`. +/-- The equivalence of functor categories given by currying/uncurrying. +-/ +@[simps] +def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E := + Equivalence.mk uncurry curry + (NatIso.ofComponents + (fun F => + NatIso.ofComponents (fun X => NatIso.ofComponents (fun Y => Iso.refl _) (by tidy)) + (by tidy)) + (by tidy)) + (NatIso.ofComponents (fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) (by tidy)) + (by tidy)) +#align category_theory.currying CategoryTheory.currying + +/-- `F.flip` is isomorphic to uncurrying `F`, swapping the variables, and currying. -/ +@[simps] +def flipIsoCurrySwapUncurry (F : C ⥤ D ⥤ E) : F.flip ≅ curry.obj (Prod.swap _ _ ⋙ uncurry.obj F) := + NatIso.ofComponents (fun d => NatIso.ofComponents (fun c => Iso.refl _) (by tidy)) (by tidy) +#align category_theory.flip_iso_curry_swap_uncurry CategoryTheory.flipIsoCurrySwapUncurry + +/-- The uncurrying of `F.flip` is isomorphic to +swapping the factors followed by the uncurrying of `F`. -/ +@[simps] +def uncurryObjFlip (F : C ⥤ D ⥤ E) : uncurry.obj F.flip ≅ Prod.swap _ _ ⋙ uncurry.obj F := + NatIso.ofComponents (fun p => Iso.refl _) (by tidy) +#align category_theory.uncurry_obj_flip CategoryTheory.uncurryObjFlip + +variable (B C D E) + +/-- A version of `category_theory.whiskering_right` for bifunctors, obtained by uncurrying, +applying `whiskering_right` and currying back +-/ +@[simps] +def whiskeringRight₂ : (C ⥤ D ⥤ E) ⥤ (B ⥤ C) ⥤ (B ⥤ D) ⥤ B ⥤ E := + uncurry ⋙ + whiskeringRight _ _ _ ⋙ (whiskeringLeft _ _ _).obj (prodFunctorToFunctorProd _ _ _) ⋙ curry +#align category_theory.whiskering_right₂ CategoryTheory.whiskeringRight₂ + +end CategoryTheory + From f147507b8876443f930b7e9de0fed695e3980e69 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 22:56:08 -0500 Subject: [PATCH 064/126] 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/Currying.lean | 2 +- 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index a82728786eb75..1a28a72227162 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -232,6 +232,7 @@ import Mathlib.CategoryTheory.FullSubcategory import Mathlib.CategoryTheory.Functor.Basic import Mathlib.CategoryTheory.Functor.Category import Mathlib.CategoryTheory.Functor.Const +import Mathlib.CategoryTheory.Functor.Currying import Mathlib.CategoryTheory.Functor.Default import Mathlib.CategoryTheory.Functor.FullyFaithful import Mathlib.CategoryTheory.Functor.Functorial diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index ee6c5edabe051..4fe23b1ae6cdd 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -8,7 +8,7 @@ Authors: Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Products.Bifunctor +import Mathlib.CategoryTheory.Products.Bifunctor /-! # Curry and uncurry, as functors. From d4926192e3d1dff291e56f2e1ba3d3127b62e77e Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 14 Feb 2023 23:34:11 -0500 Subject: [PATCH 065/126] fix all but one decl --- Mathlib/CategoryTheory/Functor/Currying.lean | 64 +++++++++++--------- 1 file changed, 35 insertions(+), 29 deletions(-) diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index 4fe23b1ae6cdd..6a1a283e770b1 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -34,22 +34,22 @@ def uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E where obj F := { obj := fun X => (F.obj X.1).obj X.2 - map := fun X Y f => (F.map f.1).app X.2 ≫ (F.obj Y.1).map f.2 - map_comp' := fun X Y Z f g => + map := fun {X} {Y} f => (F.map f.1).app X.2 ≫ (F.obj Y.1).map f.2 + map_comp := fun f g => by - simp only [prod_comp_fst, prod_comp_snd, functor.map_comp, nat_trans.comp_app, - category.assoc] - slice_lhs 2 3 => rw [← nat_trans.naturality] - rw [category.assoc] } - map F G T := + simp only [prod_comp_fst, prod_comp_snd, Functor.map_comp, NatTrans.comp_app, + Category.assoc] + slice_lhs 2 3 => rw [← NatTrans.naturality] + rw [Category.assoc] } + map T := { app := fun X => (T.app X.1).app X.2 - naturality' := fun X Y f => + naturality := fun X Y f => by - simp only [prod_comp_fst, prod_comp_snd, category.comp_id, category.assoc, Functor.map_id, - functor.map_comp, nat_trans.id_app, nat_trans.comp_app] - slice_lhs 2 3 => rw [nat_trans.naturality] - slice_lhs 1 2 => rw [← nat_trans.comp_app, nat_trans.naturality, nat_trans.comp_app] - rw [category.assoc] } + simp only [prod_comp_fst, prod_comp_snd, Category.comp_id, Category.assoc, Functor.map_id, + Functor.map_comp, NatTrans.id_app, NatTrans.comp_app] + slice_lhs 2 3 => rw [NatTrans.naturality] + slice_lhs 1 2 => rw [← NatTrans.comp_app, NatTrans.naturality, NatTrans.comp_app] + rw [Category.assoc] } #align category_theory.uncurry CategoryTheory.uncurry /-- The object level part of the currying functor. (See `curry` for the functorial version.) @@ -58,8 +58,14 @@ def curryObj (F : C × D ⥤ E) : C ⥤ D ⥤ E where obj X := { obj := fun Y => F.obj (X, Y) - map := fun Y Y' g => F.map (𝟙 X, g) } - map X X' f := { app := fun Y => F.map (f, 𝟙 Y) } + map := fun g => F.map (𝟙 X, g) + map_id := fun Y => by simp only [F.map_id]; rw [←prod_id]; exact F.map_id ⟨X,Y⟩ + map_comp := fun f g => by simp; rw [prod_comp]; } + map f := + { app := fun Y => F.map (f, 𝟙 Y) + naturality := sorry } + map_id := sorry + map_comp := sorry #align category_theory.curry_obj CategoryTheory.curryObj /-- The currying functor, taking a functor `(C × D) ⥤ E` and producing a functor `C ⥤ (D ⥤ E)`. @@ -67,15 +73,15 @@ def curryObj (F : C × D ⥤ E) : C ⥤ D ⥤ E @[simps obj_obj_obj obj_obj_map obj_map_app map_app_app] def curry : (C × D ⥤ E) ⥤ C ⥤ D ⥤ E where obj F := curryObj F - map F G T := + map T := { app := fun X => { app := fun Y => T.app (X, Y) - naturality' := fun Y Y' g => by - dsimp [curry_obj] - rw [nat_trans.naturality] } - naturality' := fun X X' f => by - ext; dsimp [curry_obj] - rw [nat_trans.naturality] } + naturality := fun Y Y' g => by + dsimp [curryObj] + rw [NatTrans.naturality] } + naturality := fun X X' f => by + ext; dsimp [curryObj] + rw [NatTrans.naturality] } #align category_theory.curry CategoryTheory.curry -- create projection simp lemmas even though this isn't a `{ .. }`. @@ -86,24 +92,24 @@ def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E := Equivalence.mk uncurry curry (NatIso.ofComponents (fun F => - NatIso.ofComponents (fun X => NatIso.ofComponents (fun Y => Iso.refl _) (by tidy)) - (by tidy)) - (by tidy)) - (NatIso.ofComponents (fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) (by tidy)) - (by tidy)) + NatIso.ofComponents (fun X => NatIso.ofComponents (fun Y => Iso.refl _) (by aesop_cat)) + (by aesop_cat)) + (by aesop_cat)) + (NatIso.ofComponents (fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) (by aesop_cat)) + (by aesop_cat)) #align category_theory.currying CategoryTheory.currying /-- `F.flip` is isomorphic to uncurrying `F`, swapping the variables, and currying. -/ @[simps] def flipIsoCurrySwapUncurry (F : C ⥤ D ⥤ E) : F.flip ≅ curry.obj (Prod.swap _ _ ⋙ uncurry.obj F) := - NatIso.ofComponents (fun d => NatIso.ofComponents (fun c => Iso.refl _) (by tidy)) (by tidy) + NatIso.ofComponents (fun d => NatIso.ofComponents (fun c => Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.flip_iso_curry_swap_uncurry CategoryTheory.flipIsoCurrySwapUncurry /-- The uncurrying of `F.flip` is isomorphic to swapping the factors followed by the uncurrying of `F`. -/ @[simps] def uncurryObjFlip (F : C ⥤ D ⥤ E) : uncurry.obj F.flip ≅ Prod.swap _ _ ⋙ uncurry.obj F := - NatIso.ofComponents (fun p => Iso.refl _) (by tidy) + NatIso.ofComponents (fun p => Iso.refl _) (by aesop_cat) #align category_theory.uncurry_obj_flip CategoryTheory.uncurryObjFlip variable (B C D E) From 82e25124b8f38e2494e781bf4c1fd0ab432effc0 Mon Sep 17 00:00:00 2001 From: adamtopaz Date: Tue, 14 Feb 2023 22:06:54 -0700 Subject: [PATCH 066/126] feat: port CategoryTheory.DiscreteCategory From 10bf4bb690318c385e5fa84948e4d987ad115bd9 Mon Sep 17 00:00:00 2001 From: adamtopaz Date: Tue, 14 Feb 2023 22:06:54 -0700 Subject: [PATCH 067/126] Initial file copy from mathport --- Mathlib/CategoryTheory/DiscreteCategory.lean | 342 +++++++++++++++++++ 1 file changed, 342 insertions(+) create mode 100644 Mathlib/CategoryTheory/DiscreteCategory.lean diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean new file mode 100644 index 0000000000000..73f5bc502005e --- /dev/null +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -0,0 +1,342 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Stephen Morgan, Scott Morrison, Floris van Doorn + +! This file was ported from Lean 3 source module category_theory.discrete_category +! leanprover-community/mathlib commit 369525b73f229ccd76a6ec0e0e0bf2be57599768 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.EqToHom +import Mathbin.Data.Ulift + +/-! +# Discrete categories + +We define `discrete α` as a structure containing a term `a : α` for any type `α`, +and use this type alias to provide a `small_category` instance +whose only morphisms are the identities. + +There is an annoying technical difficulty that it has turned out to be inconvenient +to allow categories with morphisms living in `Prop`, +so instead of defining `X ⟶ Y` in `discrete α` as `X = Y`, +one might define it as `plift (X = Y)`. +In fact, to allow `discrete α` to be a `small_category` +(i.e. with morphisms in the same universe as the objects), +we actually define the hom type `X ⟶ Y` as `ulift (plift (X = Y))`. + +`discrete.functor` promotes a function `f : I → C` (for any category `C`) to a functor +`discrete.functor f : discrete I ⥤ C`. + +Similarly, `discrete.nat_trans` and `discrete.nat_iso` promote `I`-indexed families of morphisms, +or `I`-indexed families of isomorphisms to natural transformations or natural isomorphism. + +We show equivalences of types are the same as (categorical) equivalences of the corresponding +discrete categories. +-/ + + +namespace CategoryTheory + +-- morphism levels before object levels. See note [category_theory universes]. +universe v₁ v₂ v₃ u₁ u₁' u₂ u₃ + +-- This is intentionally a structure rather than a type synonym +-- to enforce using `discrete_equiv` (or `discrete.mk` and `discrete.as`) to move between +-- `discrete α` and `α`. Otherwise there is too much API leakage. +/-- A wrapper for promoting any type to a category, +with the only morphisms being equalities. +-/ +@[ext] +structure Discrete (α : Type u₁) where + as : α +#align category_theory.discrete CategoryTheory.Discrete + +@[simp] +theorem Discrete.mk_as {α : Type u₁} (X : Discrete α) : Discrete.mk X.as = X := + by + ext + rfl +#align category_theory.discrete.mk_as CategoryTheory.Discrete.mk_as + +/-- `discrete α` is equivalent to the original type `α`.-/ +@[simps] +def discreteEquiv {α : Type u₁} : Discrete α ≃ α + where + toFun := Discrete.as + invFun := Discrete.mk + left_inv := by tidy + right_inv := by tidy +#align category_theory.discrete_equiv CategoryTheory.discreteEquiv + +instance {α : Type u₁} [DecidableEq α] : DecidableEq (Discrete α) := + discreteEquiv.DecidableEq + +/-- The "discrete" category on a type, whose morphisms are equalities. + +Because we do not allow morphisms in `Prop` (only in `Type`), +somewhat annoyingly we have to define `X ⟶ Y` as `ulift (plift (X = Y))`. + +See +-/ +instance discreteCategory (α : Type u₁) : SmallCategory (Discrete α) + where + Hom X Y := ULift (PLift (X.as = Y.as)) + id X := ULift.up (PLift.up rfl) + comp X Y Z g f := by + cases X + cases Y + cases Z + rcases f with ⟨⟨⟨⟩⟩⟩ + exact g +#align category_theory.discrete_category CategoryTheory.discreteCategory + +namespace Discrete + +variable {α : Type u₁} + +instance [Inhabited α] : Inhabited (Discrete α) := + ⟨⟨default⟩⟩ + +instance [Subsingleton α] : Subsingleton (Discrete α) := + ⟨by + intros + ext + apply Subsingleton.elim⟩ + +/- ./././Mathport/Syntax/Translate/Expr.lean:334:4: warning: unsupported (TODO): `[tacs] -/ +/-- A simple tactic to run `cases` on any `discrete α` hypotheses. -/ +unsafe def _root_.tactic.discrete_cases : tactic Unit := + sorry +#align tactic.discrete_cases tactic.discrete_cases + +run_cmd + add_interactive [`` tactic.discrete_cases] + +attribute [local tidy] tactic.discrete_cases + +instance [Unique α] : Unique (Discrete α) := + Unique.mk' (Discrete α) + +/-- Extract the equation from a morphism in a discrete category. -/ +theorem eq_of_hom {X Y : Discrete α} (i : X ⟶ Y) : X.as = Y.as := + i.down.down +#align category_theory.discrete.eq_of_hom CategoryTheory.Discrete.eq_of_hom + +/-- Promote an equation between the wrapped terms in `X Y : discrete α` to a morphism `X ⟶ Y` +in the discrete category. -/ +abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := + eqToHom + (by + ext + exact h) +#align category_theory.discrete.eq_to_hom CategoryTheory.Discrete.eqToHom + +/-- Promote an equation between the wrapped terms in `X Y : discrete α` to an isomorphism `X ≅ Y` +in the discrete category. -/ +abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := + eqToIso + (by + ext + exact h) +#align category_theory.discrete.eq_to_iso CategoryTheory.Discrete.eqToIso + +/-- A variant of `eq_to_hom` that lifts terms to the discrete category. -/ +abbrev eqToHom' {a b : α} (h : a = b) : Discrete.mk a ⟶ Discrete.mk b := + eqToHom h +#align category_theory.discrete.eq_to_hom' CategoryTheory.Discrete.eqToHom' + +/-- A variant of `eq_to_iso` that lifts terms to the discrete category. -/ +abbrev eqToIso' {a b : α} (h : a = b) : Discrete.mk a ≅ Discrete.mk b := + eqToIso h +#align category_theory.discrete.eq_to_iso' CategoryTheory.Discrete.eqToIso' + +@[simp] +theorem id_def (X : Discrete α) : ULift.up (PLift.up (Eq.refl X.as)) = 𝟙 X := + rfl +#align category_theory.discrete.id_def CategoryTheory.Discrete.id_def + +variable {C : Type u₂} [Category.{v₂} C] + +instance {I : Type u₁} {i j : Discrete I} (f : i ⟶ j) : IsIso f := + ⟨⟨eqToHom (eq_of_hom f).symm, by tidy⟩⟩ + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- Any function `I → C` gives a functor `discrete I ⥤ C`. +-/ +def functor {I : Type u₁} (F : I → C) : Discrete I ⥤ C + where + obj := F ∘ Discrete.as + map X Y f := + by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + cases f + exact 𝟙 (F X) +#align category_theory.discrete.functor CategoryTheory.Discrete.functor + +@[simp] +theorem functor_obj {I : Type u₁} (F : I → C) (i : I) : + (Discrete.functor F).obj (Discrete.mk i) = F i := + rfl +#align category_theory.discrete.functor_obj CategoryTheory.Discrete.functor_obj + +theorem functor_map {I : Type u₁} (F : I → C) {i : Discrete I} (f : i ⟶ i) : + (Discrete.functor F).map f = 𝟙 (F i.as) := by tidy +#align category_theory.discrete.functor_map CategoryTheory.Discrete.functor_map + +/-- The discrete functor induced by a composition of maps can be written as a +composition of two discrete functors. +-/ +@[simps] +def functorComp {I : Type u₁} {J : Type u₁'} (f : J → C) (g : I → J) : + Discrete.functor (f ∘ g) ≅ Discrete.functor (Discrete.mk ∘ g) ⋙ Discrete.functor f := + NatIso.ofComponents (fun X => Iso.refl _) (by tidy) +#align category_theory.discrete.functor_comp CategoryTheory.Discrete.functorComp + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- For functors out of a discrete category, +a natural transformation is just a collection of maps, +as the naturality squares are trivial. +-/ +@[simps] +def natTrans {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ⟶ G.obj i) : F ⟶ G + where + app := f + naturality' X Y g := + by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + cases g + simp +#align category_theory.discrete.nat_trans CategoryTheory.Discrete.natTrans + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- For functors out of a discrete category, +a natural isomorphism is just a collection of isomorphisms, +as the naturality squares are trivial. +-/ +@[simps] +def natIso {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) : F ≅ G := + NatIso.ofComponents f fun X Y g => + by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + cases g + simp +#align category_theory.discrete.nat_iso CategoryTheory.Discrete.natIso + +@[simp] +theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) + (i : Discrete I) : (Discrete.natIso f).app i = f i := by tidy +#align category_theory.discrete.nat_iso_app CategoryTheory.Discrete.natIso_app + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- Every functor `F` from a discrete category is naturally isomorphic (actually, equal) to + `discrete.functor (F.obj)`. -/ +@[simp] +def natIsoFunctor {I : Type u₁} {F : Discrete I ⥤ C} : F ≅ Discrete.functor (F.obj ∘ Discrete.mk) := + natIso fun i => + by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + rfl +#align category_theory.discrete.nat_iso_functor CategoryTheory.Discrete.natIsoFunctor + +/-- Composing `discrete.functor F` with another functor `G` amounts to composing `F` with `G.obj` -/ +@[simp] +def compNatIsoDiscrete {I : Type u₁} {D : Type u₃} [Category.{v₃} D] (F : I → C) (G : C ⥤ D) : + Discrete.functor F ⋙ G ≅ Discrete.functor (G.obj ∘ F) := + natIso fun i => Iso.refl _ +#align category_theory.discrete.comp_nat_iso_discrete CategoryTheory.Discrete.compNatIsoDiscrete + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- We can promote a type-level `equiv` to +an equivalence between the corresponding `discrete` categories. +-/ +@[simps] +def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ Discrete J + where + Functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) + inverse := Discrete.functor (Discrete.mk ∘ (e.symm : J → I)) + unitIso := + Discrete.natIso fun i => + eqToIso + (by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + simp) + counitIso := + Discrete.natIso fun j => + eqToIso + (by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + simp) +#align category_theory.discrete.equivalence CategoryTheory.Discrete.equivalence + +/-- We can convert an equivalence of `discrete` categories to a type-level `equiv`. -/ +@[simps] +def equivOfEquivalence {α : Type u₁} {β : Type u₂} (h : Discrete α ≌ Discrete β) : α ≃ β + where + toFun := Discrete.as ∘ h.Functor.obj ∘ Discrete.mk + invFun := Discrete.as ∘ h.inverse.obj ∘ Discrete.mk + left_inv a := by simpa using eq_of_hom (h.unit_iso.app (discrete.mk a)).2 + right_inv a := by simpa using eq_of_hom (h.counit_iso.app (discrete.mk a)).1 +#align category_theory.discrete.equiv_of_equivalence CategoryTheory.Discrete.equivOfEquivalence + +end Discrete + +namespace Discrete + +variable {J : Type v₁} + +open Opposite + +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:72:18: unsupported non-interactive tactic tactic.op_induction' -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/-- A discrete category is equivalent to its opposite category. -/ +@[simps functor_obj_as inverse_obj] +protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := + by + let F : Discrete α ⥤ (Discrete α)ᵒᵖ := Discrete.functor fun x => op (Discrete.mk x) + refine' + equivalence.mk (functor.left_op F) F _ + (discrete.nat_iso fun X => + by + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + simp [F]) + refine' + nat_iso.of_components + (fun X => + by + run_tac + tactic.op_induction' + trace + "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + simp [F]) + _ + tidy +#align category_theory.discrete.opposite CategoryTheory.Discrete.opposite + +variable {C : Type u₂} [Category.{v₂} C] + +@[simp] +theorem functor_map_id (F : Discrete J ⥤ C) {j : Discrete J} (f : j ⟶ j) : F.map f = 𝟙 (F.obj j) := + by + have h : f = 𝟙 j := by + cases f + cases f + ext + rw [h] + simp +#align category_theory.discrete.functor_map_id CategoryTheory.Discrete.functor_map_id + +end Discrete + +end CategoryTheory + From 22034856ebc97b161fcc7b9d476049ebeeb8daca Mon Sep 17 00:00:00 2001 From: adamtopaz Date: Tue, 14 Feb 2023 22:06:54 -0700 Subject: [PATCH 068/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/DiscreteCategory.lean | 10 ++++------ 2 files changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 6aa0550bef09c..532844fa3acc7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -225,6 +225,7 @@ import Mathlib.CategoryTheory.Category.KleisliCat import Mathlib.CategoryTheory.Category.RelCat import Mathlib.CategoryTheory.Comma import Mathlib.CategoryTheory.ConcreteCategory.Bundled +import Mathlib.CategoryTheory.DiscreteCategory import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Equivalence import Mathlib.CategoryTheory.EssentialImage diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 73f5bc502005e..871c80e12f731 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -8,8 +8,8 @@ Authors: 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.EqToHom -import Mathbin.Data.Ulift +import Mathlib.CategoryTheory.EqToHom +import Mathlib.Data.Ulift /-! # Discrete categories @@ -54,8 +54,7 @@ structure Discrete (α : Type u₁) where #align category_theory.discrete CategoryTheory.Discrete @[simp] -theorem Discrete.mk_as {α : Type u₁} (X : Discrete α) : Discrete.mk X.as = X := - by +theorem Discrete.mk_as {α : Type u₁} (X : Discrete α) : Discrete.mk X.as = X := by ext rfl #align category_theory.discrete.mk_as CategoryTheory.Discrete.mk_as @@ -300,8 +299,7 @@ open Opposite /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ /-- A discrete category is equivalent to its opposite category. -/ @[simps functor_obj_as inverse_obj] -protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := - by +protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := by let F : Discrete α ⥤ (Discrete α)ᵒᵖ := Discrete.functor fun x => op (Discrete.mk x) refine' equivalence.mk (functor.left_op F) F _ From 92feef62cbf017eefc2d5fad5646e5bd62463437 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 00:36:52 -0500 Subject: [PATCH 069/126] fix last errors --- Mathlib/CategoryTheory/Functor/Currying.lean | 25 ++++++++++---------- 1 file changed, 13 insertions(+), 12 deletions(-) diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index 6a1a283e770b1..3eec08d7ff1f3 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -60,17 +60,17 @@ def curryObj (F : C × D ⥤ E) : C ⥤ D ⥤ E { obj := fun Y => F.obj (X, Y) map := fun g => F.map (𝟙 X, g) map_id := fun Y => by simp only [F.map_id]; rw [←prod_id]; exact F.map_id ⟨X,Y⟩ - map_comp := fun f g => by simp; rw [prod_comp]; } + map_comp := fun f g => by simp [←F.map_comp]} map f := { app := fun Y => F.map (f, 𝟙 Y) - naturality := sorry } - map_id := sorry - map_comp := sorry + naturality := fun {Y} {Y'} g => by simp [←F.map_comp] } + map_id := fun X => by ext Y; exact F.map_id _ + map_comp := fun f g => by ext Y; dsimp; simp [←F.map_comp] #align category_theory.curry_obj CategoryTheory.curryObj /-- The currying functor, taking a functor `(C × D) ⥤ E` and producing a functor `C ⥤ (D ⥤ E)`. -/ -@[simps obj_obj_obj obj_obj_map obj_map_app map_app_app] +@[simps! obj_obj_obj obj_obj_map obj_map_app map_app_app] def curry : (C × D ⥤ E) ⥤ C ⥤ D ⥤ E where obj F := curryObj F map T := @@ -87,7 +87,7 @@ def curry : (C × D ⥤ E) ⥤ C ⥤ D ⥤ E where -- create projection simp lemmas even though this isn't a `{ .. }`. /-- The equivalence of functor categories given by currying/uncurrying. -/ -@[simps] +@[simps!] def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E := Equivalence.mk uncurry curry (NatIso.ofComponents @@ -95,19 +95,21 @@ def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E := NatIso.ofComponents (fun X => NatIso.ofComponents (fun Y => Iso.refl _) (by aesop_cat)) (by aesop_cat)) (by aesop_cat)) - (NatIso.ofComponents (fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) (by aesop_cat)) + (NatIso.ofComponents (fun F => NatIso.ofComponents (fun X => eqToIso (by simp)) + (by intros X Y f; cases X; cases Y; cases f; dsimp at *; rw [←F.map_comp]; simp )) (by aesop_cat)) #align category_theory.currying CategoryTheory.currying /-- `F.flip` is isomorphic to uncurrying `F`, swapping the variables, and currying. -/ -@[simps] +@[simps!] def flipIsoCurrySwapUncurry (F : C ⥤ D ⥤ E) : F.flip ≅ curry.obj (Prod.swap _ _ ⋙ uncurry.obj F) := - NatIso.ofComponents (fun d => NatIso.ofComponents (fun c => Iso.refl _) (by aesop_cat)) (by aesop_cat) + NatIso.ofComponents (fun d => NatIso.ofComponents (fun c => Iso.refl _) + (by aesop_cat)) (by aesop_cat) #align category_theory.flip_iso_curry_swap_uncurry CategoryTheory.flipIsoCurrySwapUncurry /-- The uncurrying of `F.flip` is isomorphic to swapping the factors followed by the uncurrying of `F`. -/ -@[simps] +@[simps!] def uncurryObjFlip (F : C ⥤ D ⥤ E) : uncurry.obj F.flip ≅ Prod.swap _ _ ⋙ uncurry.obj F := NatIso.ofComponents (fun p => Iso.refl _) (by aesop_cat) #align category_theory.uncurry_obj_flip CategoryTheory.uncurryObjFlip @@ -117,11 +119,10 @@ variable (B C D E) /-- A version of `category_theory.whiskering_right` for bifunctors, obtained by uncurrying, applying `whiskering_right` and currying back -/ -@[simps] +@[simps!] def whiskeringRight₂ : (C ⥤ D ⥤ E) ⥤ (B ⥤ C) ⥤ (B ⥤ D) ⥤ B ⥤ E := uncurry ⋙ whiskeringRight _ _ _ ⋙ (whiskeringLeft _ _ _).obj (prodFunctorToFunctorProd _ _ _) ⋙ curry #align category_theory.whiskering_right₂ CategoryTheory.whiskeringRight₂ end CategoryTheory - From 26291ab73ffbe15322093e80cbaccf07be7002d1 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 00:40:41 -0500 Subject: [PATCH 070/126] feat: port CategoryTheory.Yoneda From 5a1eb5b7785c20cddeb69e8262ce57740cd40546 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 00:40:41 -0500 Subject: [PATCH 071/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Yoneda.lean | 467 +++++++++++++++++++++++++++++ 1 file changed, 467 insertions(+) create mode 100644 Mathlib/CategoryTheory/Yoneda.lean diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean new file mode 100644 index 0000000000000..d49f765aeb8a9 --- /dev/null +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -0,0 +1,467 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison + +! This file was ported from Lean 3 source module category_theory.yoneda +! leanprover-community/mathlib commit 369525b73f229ccd76a6ec0e0e0bf2be57599768 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Functor.Hom +import Mathbin.CategoryTheory.Functor.Currying +import Mathbin.CategoryTheory.Products.Basic + +/-! +# The Yoneda embedding + +The Yoneda embedding as a functor `yoneda : C ⥤ (Cᵒᵖ ⥤ Type v₁)`, +along with an instance that it is `fully_faithful`. + +Also the Yoneda lemma, `yoneda_lemma : (yoneda_pairing C) ≅ (yoneda_evaluation C)`. + +## References +* [Stacks: Opposite Categories and the Yoneda Lemma](https://stacks.math.columbia.edu/tag/001L) +-/ + + +namespace CategoryTheory + +open Opposite + +universe v₁ u₁ u₂ + +-- morphism levels before object levels. See note [category_theory universes]. +variable {C : Type u₁} [Category.{v₁} C] + +/-- The Yoneda embedding, as a functor from `C` into presheaves on `C`. + +See . +-/ +@[simps] +def yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁ + where + obj X := + { obj := fun Y => unop Y ⟶ X + map := fun Y Y' f g => f.unop ≫ g + map_comp' := fun _ _ _ f g => by ext; dsimp; erw [category.assoc] + map_id' := fun Y => by ext; dsimp; erw [category.id_comp] } + map X X' f := { app := fun Y g => g ≫ f } +#align category_theory.yoneda CategoryTheory.yoneda + +/-- The co-Yoneda embedding, as a functor from `Cᵒᵖ` into co-presheaves on `C`. +-/ +@[simps] +def coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁ + where + obj X := + { obj := fun Y => unop X ⟶ Y + map := fun Y Y' f g => g ≫ f } + map X X' f := { app := fun Y g => f.unop ≫ g } +#align category_theory.coyoneda CategoryTheory.coyoneda + +namespace Yoneda + +theorem obj_map_id {X Y : C} (f : op X ⟶ op Y) : + (yoneda.obj X).map f (𝟙 X) = (yoneda.map f.unop).app (op Y) (𝟙 Y) := + by + dsimp + simp +#align category_theory.yoneda.obj_map_id CategoryTheory.yoneda.obj_map_id + +@[simp] +theorem naturality {X Y : C} (α : yoneda.obj X ⟶ yoneda.obj Y) {Z Z' : C} (f : Z ⟶ Z') + (h : Z' ⟶ X) : f ≫ α.app (op Z') h = α.app (op Z) (f ≫ h) := + (FunctorToTypes.naturality _ _ α f.op h).symm +#align category_theory.yoneda.naturality CategoryTheory.yoneda.naturality + +/-- The Yoneda embedding is full. + +See . +-/ +instance yonedaFull : Full (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) where preimage X Y f := f.app (op X) (𝟙 X) +#align category_theory.yoneda.yoneda_full CategoryTheory.yoneda.yonedaFull + +/-- The Yoneda embedding is faithful. + +See . +-/ +instance yoneda_faithful : Faithful (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) + where map_injective' X Y f g p := by + convert congr_fun (congr_app p (op X)) (𝟙 X) <;> dsimp <;> simp +#align category_theory.yoneda.yoneda_faithful CategoryTheory.yoneda.yoneda_faithful + +/-- Extensionality via Yoneda. The typical usage would be +``` +-- Goal is `X ≅ Y` +apply yoneda.ext, +-- Goals are now functions `(Z ⟶ X) → (Z ⟶ Y)`, `(Z ⟶ Y) → (Z ⟶ X)`, and the fact that these +functions are inverses and natural in `Z`. +``` +-/ +def ext (X Y : C) (p : ∀ {Z : C}, (Z ⟶ X) → (Z ⟶ Y)) (q : ∀ {Z : C}, (Z ⟶ Y) → (Z ⟶ X)) + (h₁ : ∀ {Z : C} (f : Z ⟶ X), q (p f) = f) (h₂ : ∀ {Z : C} (f : Z ⟶ Y), p (q f) = f) + (n : ∀ {Z Z' : C} (f : Z' ⟶ Z) (g : Z ⟶ X), p (f ≫ g) = f ≫ p g) : X ≅ Y := + yoneda.preimageIso + (NatIso.ofComponents + (fun Z => + { Hom := p + inv := q }) + (by tidy)) +#align category_theory.yoneda.ext CategoryTheory.yoneda.ext + +/-- If `yoneda.map f` is an isomorphism, so was `f`. +-/ +theorem isIso {X Y : C} (f : X ⟶ Y) [IsIso (yoneda.map f)] : IsIso f := + isIso_of_fully_faithful yoneda f +#align category_theory.yoneda.is_iso CategoryTheory.yoneda.isIso + +end Yoneda + +namespace Coyoneda + +@[simp] +theorem naturality {X Y : Cᵒᵖ} (α : coyoneda.obj X ⟶ coyoneda.obj Y) {Z Z' : C} (f : Z' ⟶ Z) + (h : unop X ⟶ Z') : α.app Z' h ≫ f = α.app Z (h ≫ f) := + (FunctorToTypes.naturality _ _ α f h).symm +#align category_theory.coyoneda.naturality CategoryTheory.coyoneda.naturality + +instance coyonedaFull : Full (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) + where preimage X Y f := (f.app _ (𝟙 X.unop)).op +#align category_theory.coyoneda.coyoneda_full CategoryTheory.coyoneda.coyonedaFull + +instance coyoneda_faithful : Faithful (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) + where map_injective' X Y f g p := + by + have t := congr_fun (congr_app p X.unop) (𝟙 _) + simpa using congr_arg Quiver.Hom.op t +#align category_theory.coyoneda.coyoneda_faithful CategoryTheory.coyoneda.coyoneda_faithful + +/-- If `coyoneda.map f` is an isomorphism, so was `f`. +-/ +theorem isIso {X Y : Cᵒᵖ} (f : X ⟶ Y) [IsIso (coyoneda.map f)] : IsIso f := + isIso_of_fully_faithful coyoneda f +#align category_theory.coyoneda.is_iso CategoryTheory.coyoneda.isIso + +/-- The identity functor on `Type` is isomorphic to the coyoneda functor coming from `punit`. -/ +def punitIso : coyoneda.obj (Opposite.op PUnit) ≅ 𝟭 (Type v₁) := + NatIso.ofComponents + (fun X => + { Hom := fun f => f ⟨⟩ + inv := fun x _ => x }) + (by tidy) +#align category_theory.coyoneda.punit_iso CategoryTheory.coyoneda.punitIso + +/-- Taking the `unop` of morphisms is a natural isomorphism. -/ +@[simps] +def objOpOp (X : C) : coyoneda.obj (op (op X)) ≅ yoneda.obj X := + NatIso.ofComponents (fun Y => (opEquiv _ _).toIso) fun X Y f => rfl +#align category_theory.coyoneda.obj_op_op CategoryTheory.coyoneda.objOpOp + +end Coyoneda + +namespace Functor + +/-- A functor `F : Cᵒᵖ ⥤ Type v₁` is representable if there is object `X` so `F ≅ yoneda.obj X`. + +See . +-/ +class Representable (F : Cᵒᵖ ⥤ Type v₁) : Prop where + has_representation : ∃ (X : _)(f : yoneda.obj X ⟶ F), IsIso f +#align category_theory.functor.representable CategoryTheory.Functor.Representable + +instance {X : C} : Representable (yoneda.obj X) where has_representation := ⟨X, 𝟙 _, inferInstance⟩ + +/-- A functor `F : C ⥤ Type v₁` is corepresentable if there is object `X` so `F ≅ coyoneda.obj X`. + +See . +-/ +class Corepresentable (F : C ⥤ Type v₁) : Prop where + has_corepresentation : ∃ (X : _)(f : coyoneda.obj X ⟶ F), IsIso f +#align category_theory.functor.corepresentable CategoryTheory.Functor.Corepresentable + +instance {X : Cᵒᵖ} : Corepresentable (coyoneda.obj X) + where has_corepresentation := ⟨X, 𝟙 _, inferInstance⟩ + +-- instance : corepresentable (𝟭 (Type v₁)) := +-- corepresentable_of_nat_iso (op punit) coyoneda.punit_iso +section Representable + +variable (F : Cᵒᵖ ⥤ Type v₁) + +variable [F.Representable] + +/-- The representing object for the representable functor `F`. -/ +noncomputable def reprX : C := + (Representable.has_representation : ∃ (X : _)(f : _ ⟶ F), _).some +#align category_theory.functor.repr_X CategoryTheory.Functor.reprX + +/-- The (forward direction of the) isomorphism witnessing `F` is representable. -/ +noncomputable def reprF : yoneda.obj F.reprX ⟶ F := + Representable.has_representation.choose_spec.some +#align category_theory.functor.repr_f CategoryTheory.Functor.reprF + +/- warning: category_theory.functor.repr_x clashes with category_theory.functor.repr_X -> CategoryTheory.Functor.reprX +warning: category_theory.functor.repr_x -> CategoryTheory.Functor.reprX is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1} F (Opposite.op.{succ u2} C (CategoryTheory.Functor.reprX.{u1, u2} C _inst_1 F _inst_2)) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], C +Case conversion may be inaccurate. Consider using '#align category_theory.functor.repr_x CategoryTheory.Functor.reprXₓ'. -/ +/-- The representing element for the representable functor `F`, sometimes called the universal +element of the functor. +-/ +noncomputable def reprX : F.obj (op F.reprX) := + F.reprF.app (op F.reprX) (𝟙 F.reprX) +#align category_theory.functor.repr_x CategoryTheory.Functor.reprX + +instance : IsIso F.reprF := + Representable.has_representation.choose_spec.choose_spec + +/-- An isomorphism between `F` and a functor of the form `C(-, F.repr_X)`. Note the components +`F.repr_w.app X` definitionally have type `(X.unop ⟶ F.repr_X) ≅ F.obj X`. +-/ +noncomputable def reprW : yoneda.obj F.reprX ≅ F := + asIso F.reprF +#align category_theory.functor.repr_w CategoryTheory.Functor.reprW + +@[simp] +theorem reprW_hom : F.reprW.Hom = F.reprF := + rfl +#align category_theory.functor.repr_w_hom CategoryTheory.Functor.reprW_hom + +theorem reprW_app_hom (X : Cᵒᵖ) (f : unop X ⟶ F.reprX) : + (F.reprW.app X).Hom f = F.map f.op F.reprX := + by + change F.repr_f.app X f = (F.repr_f.app (op F.repr_X) ≫ F.map f.op) (𝟙 F.repr_X) + rw [← F.repr_f.naturality] + dsimp + simp +#align category_theory.functor.repr_w_app_hom CategoryTheory.Functor.reprW_app_hom + +end Representable + +section Corepresentable + +variable (F : C ⥤ Type v₁) + +variable [F.Corepresentable] + +/-- The representing object for the corepresentable functor `F`. -/ +noncomputable def coreprX : C := + (Corepresentable.has_corepresentation : ∃ (X : _)(f : _ ⟶ F), _).some.unop +#align category_theory.functor.corepr_X CategoryTheory.Functor.coreprX + +/-- The (forward direction of the) isomorphism witnessing `F` is corepresentable. -/ +noncomputable def coreprF : coyoneda.obj (op F.coreprX) ⟶ F := + Corepresentable.has_corepresentation.choose_spec.some +#align category_theory.functor.corepr_f CategoryTheory.Functor.coreprF + +/- warning: category_theory.functor.corepr_x clashes with category_theory.functor.corepr_X -> CategoryTheory.Functor.coreprX +warning: category_theory.functor.corepr_x -> CategoryTheory.Functor.coreprX is a dubious translation: +lean 3 declaration is + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1} F (CategoryTheory.Functor.coreprX.{u1, u2} C _inst_1 F _inst_2) +but is expected to have type + forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], C +Case conversion may be inaccurate. Consider using '#align category_theory.functor.corepr_x CategoryTheory.Functor.coreprXₓ'. -/ +/-- The representing element for the corepresentable functor `F`, sometimes called the universal +element of the functor. +-/ +noncomputable def coreprX : F.obj F.coreprX := + F.coreprF.app F.coreprX (𝟙 F.coreprX) +#align category_theory.functor.corepr_x CategoryTheory.Functor.coreprX + +instance : IsIso F.coreprF := + Corepresentable.has_corepresentation.choose_spec.choose_spec + +/-- An isomorphism between `F` and a functor of the form `C(F.corepr X, -)`. Note the components +`F.corepr_w.app X` definitionally have type `F.corepr_X ⟶ X ≅ F.obj X`. +-/ +noncomputable def coreprW : coyoneda.obj (op F.coreprX) ≅ F := + asIso F.coreprF +#align category_theory.functor.corepr_w CategoryTheory.Functor.coreprW + +theorem coreprW_app_hom (X : C) (f : F.coreprX ⟶ X) : (F.coreprW.app X).Hom f = F.map f F.coreprX := + by + change F.corepr_f.app X f = (F.corepr_f.app F.corepr_X ≫ F.map f) (𝟙 F.corepr_X) + rw [← F.corepr_f.naturality] + dsimp + simp +#align category_theory.functor.corepr_w_app_hom CategoryTheory.Functor.coreprW_app_hom + +end Corepresentable + +end Functor + +theorem representable_of_nat_iso (F : Cᵒᵖ ⥤ Type v₁) {G} (i : F ≅ G) [F.Representable] : + G.Representable := + { has_representation := ⟨F.reprX, F.reprF ≫ i.Hom, inferInstance⟩ } +#align category_theory.representable_of_nat_iso CategoryTheory.representable_of_nat_iso + +theorem corepresentable_of_nat_iso (F : C ⥤ Type v₁) {G} (i : F ≅ G) [F.Corepresentable] : + G.Corepresentable := + { has_corepresentation := ⟨op F.coreprX, F.coreprF ≫ i.Hom, inferInstance⟩ } +#align category_theory.corepresentable_of_nat_iso CategoryTheory.corepresentable_of_nat_iso + +instance : Functor.Corepresentable (𝟭 (Type v₁)) := + corepresentable_of_nat_iso (coyoneda.obj (op PUnit)) coyoneda.punitIso + +open Opposite + +variable (C) + +-- We need to help typeclass inference with some awkward universe levels here. +instance prodCategoryInstance1 : Category ((Cᵒᵖ ⥤ Type v₁) × Cᵒᵖ) := + CategoryTheory.prod.{max u₁ v₁, v₁} (Cᵒᵖ ⥤ Type v₁) Cᵒᵖ +#align category_theory.prod_category_instance_1 CategoryTheory.prodCategoryInstance1 + +instance prodCategoryInstance2 : Category (Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)) := + CategoryTheory.prod.{v₁, max u₁ v₁} Cᵒᵖ (Cᵒᵖ ⥤ Type v₁) +#align category_theory.prod_category_instance_2 CategoryTheory.prodCategoryInstance2 + +open Yoneda + +/-- The "Yoneda evaluation" functor, which sends `X : Cᵒᵖ` and `F : Cᵒᵖ ⥤ Type` +to `F.obj X`, functorially in both `X` and `F`. +-/ +def yonedaEvaluation : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁) ⥤ Type max u₁ v₁ := + evaluationUncurried Cᵒᵖ (Type v₁) ⋙ uliftFunctor.{u₁} +#align category_theory.yoneda_evaluation CategoryTheory.yonedaEvaluation + +@[simp] +theorem yonedaEvaluation_map_down (P Q : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)) (α : P ⟶ Q) + (x : (yonedaEvaluation C).obj P) : + ((yonedaEvaluation C).map α x).down = α.2.app Q.1 (P.2.map α.1 x.down) := + rfl +#align category_theory.yoneda_evaluation_map_down CategoryTheory.yonedaEvaluation_map_down + +/-- The "Yoneda pairing" functor, which sends `X : Cᵒᵖ` and `F : Cᵒᵖ ⥤ Type` +to `yoneda.op.obj X ⟶ F`, functorially in both `X` and `F`. +-/ +def yonedaPairing : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁) ⥤ Type max u₁ v₁ := + Functor.prod yoneda.op (𝟭 (Cᵒᵖ ⥤ Type v₁)) ⋙ Functor.hom (Cᵒᵖ ⥤ Type v₁) +#align category_theory.yoneda_pairing CategoryTheory.yonedaPairing + +@[simp] +theorem yonedaPairing_map (P Q : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)) (α : P ⟶ Q) (β : (yonedaPairing C).obj P) : + (yonedaPairing C).map α β = yoneda.map α.1.unop ≫ β ≫ α.2 := + rfl +#align category_theory.yoneda_pairing_map CategoryTheory.yonedaPairing_map + +/-- The Yoneda lemma asserts that that the Yoneda pairing +`(X : Cᵒᵖ, F : Cᵒᵖ ⥤ Type) ↦ (yoneda.obj (unop X) ⟶ F)` +is naturally isomorphic to the evaluation `(X, F) ↦ F.obj X`. + +See . +-/ +def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C + where + Hom := + { app := fun F x => ULift.up ((x.app F.1) (𝟙 (unop F.1))) + naturality' := by + intro X Y f; ext; dsimp + erw [category.id_comp, ← functor_to_types.naturality] + simp only [category.comp_id, yoneda_obj_map] } + inv := + { app := fun F x => + { app := fun X a => (F.2.map a.op) x.down + naturality' := by + intro X Y f; ext; dsimp + rw [functor_to_types.map_comp_apply] } + naturality' := by + intro X Y f; ext; dsimp + rw [← functor_to_types.naturality, functor_to_types.map_comp_apply] } + hom_inv_id' := by + ext; dsimp + erw [← functor_to_types.naturality, obj_map_id] + simp only [yoneda_map_app, Quiver.Hom.unop_op] + erw [category.id_comp] + inv_hom_id' := by + ext; dsimp + rw [functor_to_types.map_id_apply] +#align category_theory.yoneda_lemma CategoryTheory.yonedaLemma + +variable {C} + +/-- The isomorphism between `yoneda.obj X ⟶ F` and `F.obj (op X)` +(we need to insert a `ulift` to get the universes right!) +given by the Yoneda lemma. +-/ +@[simps] +def yonedaSections (X : C) (F : Cᵒᵖ ⥤ Type v₁) : (yoneda.obj X ⟶ F) ≅ ULift.{u₁} (F.obj (op X)) := + (yonedaLemma C).app (op X, F) +#align category_theory.yoneda_sections CategoryTheory.yonedaSections + +/-- We have a type-level equivalence between natural transformations from the yoneda embedding +and elements of `F.obj X`, without any universe switching. +-/ +def yonedaEquiv {X : C} {F : Cᵒᵖ ⥤ Type v₁} : (yoneda.obj X ⟶ F) ≃ F.obj (op X) := + (yonedaSections X F).toEquiv.trans Equiv.ulift +#align category_theory.yoneda_equiv CategoryTheory.yonedaEquiv + +@[simp] +theorem yonedaEquiv_apply {X : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) : + yonedaEquiv f = f.app (op X) (𝟙 X) := + rfl +#align category_theory.yoneda_equiv_apply CategoryTheory.yonedaEquiv_apply + +@[simp] +theorem yonedaEquiv_symm_app_apply {X : C} {F : Cᵒᵖ ⥤ Type v₁} (x : F.obj (op X)) (Y : Cᵒᵖ) + (f : Y.unop ⟶ X) : (yonedaEquiv.symm x).app Y f = F.map f.op x := + rfl +#align category_theory.yoneda_equiv_symm_app_apply CategoryTheory.yonedaEquiv_symm_app_apply + +theorem yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) (g : Y ⟶ X) : + F.map g.op (yonedaEquiv f) = yonedaEquiv (yoneda.map g ≫ f) := + by + change (f.app (op X) ≫ F.map g.op) (𝟙 X) = f.app (op Y) (𝟙 Y ≫ g) + rw [← f.naturality] + dsimp + simp +#align category_theory.yoneda_equiv_naturality CategoryTheory.yonedaEquiv_naturality + +/-- When `C` is a small category, we can restate the isomorphism from `yoneda_sections` +without having to change universes. +-/ +def yonedaSectionsSmall {C : Type u₁} [SmallCategory C] (X : C) (F : Cᵒᵖ ⥤ Type u₁) : + (yoneda.obj X ⟶ F) ≅ F.obj (op X) := + yonedaSections X F ≪≫ uliftTrivial _ +#align category_theory.yoneda_sections_small CategoryTheory.yonedaSectionsSmall + +@[simp] +theorem yonedaSectionsSmall_hom {C : Type u₁} [SmallCategory C] (X : C) (F : Cᵒᵖ ⥤ Type u₁) + (f : yoneda.obj X ⟶ F) : (yonedaSectionsSmall X F).Hom f = f.app _ (𝟙 _) := + rfl +#align category_theory.yoneda_sections_small_hom CategoryTheory.yonedaSectionsSmall_hom + +@[simp] +theorem yonedaSectionsSmall_inv_app_apply {C : Type u₁} [SmallCategory C] (X : C) + (F : Cᵒᵖ ⥤ Type u₁) (t : F.obj (op X)) (Y : Cᵒᵖ) (f : Y.unop ⟶ X) : + ((yonedaSectionsSmall X F).inv t).app Y f = F.map f.op t := + rfl +#align category_theory.yoneda_sections_small_inv_app_apply CategoryTheory.yonedaSectionsSmall_inv_app_apply + +attribute [local ext] Functor.ext + +/-- The curried version of yoneda lemma when `C` is small. -/ +def curriedYonedaLemma {C : Type u₁} [SmallCategory C] : + (yoneda.op ⋙ coyoneda : Cᵒᵖ ⥤ (Cᵒᵖ ⥤ Type u₁) ⥤ Type u₁) ≅ evaluation Cᵒᵖ (Type u₁) := + eqToIso (by tidy) ≪≫ + curry.mapIso + (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial) ≪≫ + eqToIso (by tidy) +#align category_theory.curried_yoneda_lemma CategoryTheory.curriedYonedaLemma + +/-- The curried version of yoneda lemma when `C` is small. -/ +def curriedYonedaLemma' {C : Type u₁} [SmallCategory C] : + yoneda ⋙ (whiskeringLeft Cᵒᵖ (Cᵒᵖ ⥤ Type u₁)ᵒᵖ (Type u₁)).obj yoneda.op ≅ 𝟭 (Cᵒᵖ ⥤ Type u₁) := + eqToIso (by tidy) ≪≫ + curry.mapIso + (isoWhiskerLeft (Prod.swap _ _) + (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial : + _)) ≪≫ + eqToIso (by tidy) +#align category_theory.curried_yoneda_lemma' CategoryTheory.curriedYonedaLemma' + +end CategoryTheory + From 3a22ebfc6b3d40a01f30fded3482451ee5dd93f6 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 00:40:41 -0500 Subject: [PATCH 072/126] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Yoneda.lean | 15 ++++++--------- 2 files changed, 7 insertions(+), 9 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 6aa0550bef09c..5283dd4c6f379 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -247,6 +247,7 @@ import Mathlib.CategoryTheory.Products.Bifunctor import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.Thin import Mathlib.CategoryTheory.Whiskering +import Mathlib.CategoryTheory.Yoneda import Mathlib.Combinatorics.Additive.Energy import Mathlib.Combinatorics.Additive.RuzsaCovering import Mathlib.Combinatorics.Colex diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index d49f765aeb8a9..00f6e3535237d 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -8,9 +8,9 @@ Authors: Scott Morrison ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.CategoryTheory.Functor.Hom -import Mathbin.CategoryTheory.Functor.Currying -import Mathbin.CategoryTheory.Products.Basic +import Mathlib.CategoryTheory.Functor.Hom +import Mathlib.CategoryTheory.Functor.Currying +import Mathlib.CategoryTheory.Products.Basic /-! # The Yoneda embedding @@ -63,8 +63,7 @@ def coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁ namespace Yoneda theorem obj_map_id {X Y : C} (f : op X ⟶ op Y) : - (yoneda.obj X).map f (𝟙 X) = (yoneda.map f.unop).app (op Y) (𝟙 Y) := - by + (yoneda.obj X).map f (𝟙 X) = (yoneda.map f.unop).app (op Y) (𝟙 Y) := by dsimp simp #align category_theory.yoneda.obj_map_id CategoryTheory.yoneda.obj_map_id @@ -231,8 +230,7 @@ theorem reprW_hom : F.reprW.Hom = F.reprF := #align category_theory.functor.repr_w_hom CategoryTheory.Functor.reprW_hom theorem reprW_app_hom (X : Cᵒᵖ) (f : unop X ⟶ F.reprX) : - (F.reprW.app X).Hom f = F.map f.op F.reprX := - by + (F.reprW.app X).Hom f = F.map f.op F.reprX := by change F.repr_f.app X f = (F.repr_f.app (op F.repr_X) ≫ F.map f.op) (𝟙 F.repr_X) rw [← F.repr_f.naturality] dsimp @@ -412,8 +410,7 @@ theorem yonedaEquiv_symm_app_apply {X : C} {F : Cᵒᵖ ⥤ Type v₁} (x : F.ob #align category_theory.yoneda_equiv_symm_app_apply CategoryTheory.yonedaEquiv_symm_app_apply theorem yonedaEquiv_naturality {X Y : C} {F : Cᵒᵖ ⥤ Type v₁} (f : yoneda.obj X ⟶ F) (g : Y ⟶ X) : - F.map g.op (yonedaEquiv f) = yonedaEquiv (yoneda.map g ≫ f) := - by + F.map g.op (yonedaEquiv f) = yonedaEquiv (yoneda.map g ≫ f) := by change (f.app (op X) ≫ F.map g.op) (𝟙 X) = f.app (op Y) (𝟙 Y ≫ g) rw [← f.naturality] dsimp From 8a683dd0861f9755b384525d3707a3e9af5278b9 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 01:13:21 -0500 Subject: [PATCH 073/126] fix some errors --- Mathlib/CategoryTheory/Yoneda.lean | 154 ++++++++++++++++------------- 1 file changed, 84 insertions(+), 70 deletions(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 00f6e3535237d..77f23681c5a97 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -16,7 +16,7 @@ import Mathlib.CategoryTheory.Products.Basic # The Yoneda embedding The Yoneda embedding as a functor `yoneda : C ⥤ (Cᵒᵖ ⥤ Type v₁)`, -along with an instance that it is `fully_faithful`. +along with an instance that it is `FullyFaithful`. Also the Yoneda lemma, `yoneda_lemma : (yoneda_pairing C) ≅ (yoneda_evaluation C)`. @@ -43,10 +43,14 @@ def yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁ where obj X := { obj := fun Y => unop Y ⟶ X - map := fun Y Y' f g => f.unop ≫ g - map_comp' := fun _ _ _ f g => by ext; dsimp; erw [category.assoc] - map_id' := fun Y => by ext; dsimp; erw [category.id_comp] } - map X X' f := { app := fun Y g => g ≫ f } + map := fun f g => f.unop ≫ g + map_comp := fun f g => by funext; dsimp; erw [Category.assoc] + map_id := fun Y => by funext; dsimp; erw [Category.id_comp] } + map f := + { app := fun Y g => g ≫ f + naturality := sorry } + map_id := by aesop_cat + map_comp := sorry #align category_theory.yoneda CategoryTheory.yoneda /-- The co-Yoneda embedding, as a functor from `Cᵒᵖ` into co-presheaves on `C`. @@ -56,8 +60,12 @@ def coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁ where obj X := { obj := fun Y => unop X ⟶ Y - map := fun Y Y' f g => g ≫ f } - map X X' f := { app := fun Y g => f.unop ≫ g } + map := fun f g => g ≫ f + map_comp := sorry + map_id := sorry } + map f := + { app := fun Y g => f.unop ≫ g + naturality := sorry } #align category_theory.coyoneda CategoryTheory.coyoneda namespace Yoneda @@ -66,29 +74,32 @@ theorem obj_map_id {X Y : C} (f : op X ⟶ op Y) : (yoneda.obj X).map f (𝟙 X) = (yoneda.map f.unop).app (op Y) (𝟙 Y) := by dsimp simp -#align category_theory.yoneda.obj_map_id CategoryTheory.yoneda.obj_map_id +#align category_theory.yoneda.obj_map_id CategoryTheory.Yoneda.obj_map_id @[simp] theorem naturality {X Y : C} (α : yoneda.obj X ⟶ yoneda.obj Y) {Z Z' : C} (f : Z ⟶ Z') (h : Z' ⟶ X) : f ≫ α.app (op Z') h = α.app (op Z) (f ≫ h) := (FunctorToTypes.naturality _ _ α f.op h).symm -#align category_theory.yoneda.naturality CategoryTheory.yoneda.naturality +#align category_theory.yoneda.naturality CategoryTheory.Yoneda.naturality /-- The Yoneda embedding is full. See . -/ -instance yonedaFull : Full (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) where preimage X Y f := f.app (op X) (𝟙 X) -#align category_theory.yoneda.yoneda_full CategoryTheory.yoneda.yonedaFull +instance yonedaFull : Full (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) where + preimage {X} {Y} f := f.app (op X) (𝟙 X) + witness := sorry +#align category_theory.yoneda.yoneda_full CategoryTheory.Yoneda.yonedaFull /-- The Yoneda embedding is faithful. See . -/ instance yoneda_faithful : Faithful (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) - where map_injective' X Y f g p := by - convert congr_fun (congr_app p (op X)) (𝟙 X) <;> dsimp <;> simp -#align category_theory.yoneda.yoneda_faithful CategoryTheory.yoneda.yoneda_faithful + where + map_injective {X} {Y} f g p := by + convert congr_fun (congr_app p (op X)) (𝟙 X) <;> dsimp <;> simp +#align category_theory.yoneda.yoneda_faithful CategoryTheory.Yoneda.yoneda_faithful /-- Extensionality via Yoneda. The typical usage would be ``` @@ -104,16 +115,16 @@ def ext (X Y : C) (p : ∀ {Z : C}, (Z ⟶ X) → (Z ⟶ Y)) (q : ∀ {Z : C}, ( yoneda.preimageIso (NatIso.ofComponents (fun Z => - { Hom := p + { hom := p inv := q }) - (by tidy)) -#align category_theory.yoneda.ext CategoryTheory.yoneda.ext + (by aesop_cat)) +#align category_theory.yoneda.ext CategoryTheory.Yoneda.ext /-- If `yoneda.map f` is an isomorphism, so was `f`. -/ theorem isIso {X Y : C} (f : X ⟶ Y) [IsIso (yoneda.map f)] : IsIso f := isIso_of_fully_faithful yoneda f -#align category_theory.yoneda.is_iso CategoryTheory.yoneda.isIso +#align category_theory.yoneda.is_iso CategoryTheory.Yoneda.isIso end Yoneda @@ -123,39 +134,40 @@ namespace Coyoneda theorem naturality {X Y : Cᵒᵖ} (α : coyoneda.obj X ⟶ coyoneda.obj Y) {Z Z' : C} (f : Z' ⟶ Z) (h : unop X ⟶ Z') : α.app Z' h ≫ f = α.app Z (h ≫ f) := (FunctorToTypes.naturality _ _ α f h).symm -#align category_theory.coyoneda.naturality CategoryTheory.coyoneda.naturality +#align category_theory.coyoneda.naturality CategoryTheory.Coyoneda.naturality instance coyonedaFull : Full (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) - where preimage X Y f := (f.app _ (𝟙 X.unop)).op -#align category_theory.coyoneda.coyoneda_full CategoryTheory.coyoneda.coyonedaFull + where + preimage {X} {Y} f := (f.app _ (𝟙 X.unop)).op + witness := sorry +#align category_theory.coyoneda.coyoneda_full CategoryTheory.Coyoneda.coyonedaFull -instance coyoneda_faithful : Faithful (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) - where map_injective' X Y f g p := - by +instance coyoneda_faithful : Faithful (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) where + map_injective g p := by have t := congr_fun (congr_app p X.unop) (𝟙 _) simpa using congr_arg Quiver.Hom.op t -#align category_theory.coyoneda.coyoneda_faithful CategoryTheory.coyoneda.coyoneda_faithful +#align category_theory.coyoneda.coyoneda_faithful CategoryTheory.Coyoneda.coyoneda_faithful /-- If `coyoneda.map f` is an isomorphism, so was `f`. -/ theorem isIso {X Y : Cᵒᵖ} (f : X ⟶ Y) [IsIso (coyoneda.map f)] : IsIso f := isIso_of_fully_faithful coyoneda f -#align category_theory.coyoneda.is_iso CategoryTheory.coyoneda.isIso +#align category_theory.coyoneda.is_iso CategoryTheory.Coyoneda.isIso /-- The identity functor on `Type` is isomorphic to the coyoneda functor coming from `punit`. -/ def punitIso : coyoneda.obj (Opposite.op PUnit) ≅ 𝟭 (Type v₁) := NatIso.ofComponents (fun X => - { Hom := fun f => f ⟨⟩ + { hom := fun f => f ⟨⟩ inv := fun x _ => x }) - (by tidy) -#align category_theory.coyoneda.punit_iso CategoryTheory.coyoneda.punitIso + (by aesop_cat) +#align category_theory.coyoneda.punit_iso CategoryTheory.Coyoneda.punitIso /-- Taking the `unop` of morphisms is a natural isomorphism. -/ @[simps] def objOpOp (X : C) : coyoneda.obj (op (op X)) ≅ yoneda.obj X := - NatIso.ofComponents (fun Y => (opEquiv _ _).toIso) fun X Y f => rfl -#align category_theory.coyoneda.obj_op_op CategoryTheory.coyoneda.objOpOp + NatIso.ofComponents (fun _ => (opEquiv _ _).toIso) fun _ {_Y} _f => rfl +#align category_theory.coyoneda.obj_op_op CategoryTheory.Coyoneda.objOpOp end Coyoneda @@ -192,12 +204,13 @@ variable [F.Representable] /-- The representing object for the representable functor `F`. -/ noncomputable def reprX : C := - (Representable.has_representation : ∃ (X : _)(f : _ ⟶ F), _).some + (Representable.has_representation : ∃ (_ : _)(_ : _ ⟶ F), _).choose +set_option linter.uppercaseLean3 false #align category_theory.functor.repr_X CategoryTheory.Functor.reprX /-- The (forward direction of the) isomorphism witnessing `F` is representable. -/ noncomputable def reprF : yoneda.obj F.reprX ⟶ F := - Representable.has_representation.choose_spec.some + Representable.has_representation.choose_spec.choose #align category_theory.functor.repr_f CategoryTheory.Functor.reprF /- warning: category_theory.functor.repr_x clashes with category_theory.functor.repr_X -> CategoryTheory.Functor.reprX @@ -210,9 +223,9 @@ Case conversion may be inaccurate. Consider using '#align category_theory.functo /-- The representing element for the representable functor `F`, sometimes called the universal element of the functor. -/ -noncomputable def reprX : F.obj (op F.reprX) := +noncomputable def reprx : F.obj (op F.reprX) := F.reprF.app (op F.reprX) (𝟙 F.reprX) -#align category_theory.functor.repr_x CategoryTheory.Functor.reprX +#align category_theory.functor.repr_x CategoryTheory.Functor.reprx instance : IsIso F.reprF := Representable.has_representation.choose_spec.choose_spec @@ -225,14 +238,14 @@ noncomputable def reprW : yoneda.obj F.reprX ≅ F := #align category_theory.functor.repr_w CategoryTheory.Functor.reprW @[simp] -theorem reprW_hom : F.reprW.Hom = F.reprF := +theorem reprW_hom : F.reprW.hom = F.reprF := rfl #align category_theory.functor.repr_w_hom CategoryTheory.Functor.reprW_hom theorem reprW_app_hom (X : Cᵒᵖ) (f : unop X ⟶ F.reprX) : - (F.reprW.app X).Hom f = F.map f.op F.reprX := by - change F.repr_f.app X f = (F.repr_f.app (op F.repr_X) ≫ F.map f.op) (𝟙 F.repr_X) - rw [← F.repr_f.naturality] + (F.reprW.app X).hom f = F.map f.op F.reprx := by + change F.reprF.app X f = (F.reprF.app (op F.reprX) ≫ F.map f.op) (𝟙 F.reprX) + rw [← F.reprF.naturality] dsimp simp #align category_theory.functor.repr_w_app_hom CategoryTheory.Functor.reprW_app_hom @@ -247,12 +260,13 @@ variable [F.Corepresentable] /-- The representing object for the corepresentable functor `F`. -/ noncomputable def coreprX : C := - (Corepresentable.has_corepresentation : ∃ (X : _)(f : _ ⟶ F), _).some.unop + (Corepresentable.has_corepresentation : ∃ (_ : _)(_ : _ ⟶ F), _).choose.unop +set_option linter.uppercaseLean3 false #align category_theory.functor.corepr_X CategoryTheory.Functor.coreprX /-- The (forward direction of the) isomorphism witnessing `F` is corepresentable. -/ noncomputable def coreprF : coyoneda.obj (op F.coreprX) ⟶ F := - Corepresentable.has_corepresentation.choose_spec.some + Corepresentable.has_corepresentation.choose_spec.choose #align category_theory.functor.corepr_f CategoryTheory.Functor.coreprF /- warning: category_theory.functor.corepr_x clashes with category_theory.functor.corepr_X -> CategoryTheory.Functor.coreprX @@ -265,9 +279,9 @@ Case conversion may be inaccurate. Consider using '#align category_theory.functo /-- The representing element for the corepresentable functor `F`, sometimes called the universal element of the functor. -/ -noncomputable def coreprX : F.obj F.coreprX := +noncomputable def coreprx : F.obj F.coreprX := F.coreprF.app F.coreprX (𝟙 F.coreprX) -#align category_theory.functor.corepr_x CategoryTheory.Functor.coreprX +#align category_theory.functor.corepr_x CategoryTheory.Functor.coreprx instance : IsIso F.coreprF := Corepresentable.has_corepresentation.choose_spec.choose_spec @@ -279,10 +293,10 @@ noncomputable def coreprW : coyoneda.obj (op F.coreprX) ≅ F := asIso F.coreprF #align category_theory.functor.corepr_w CategoryTheory.Functor.coreprW -theorem coreprW_app_hom (X : C) (f : F.coreprX ⟶ X) : (F.coreprW.app X).Hom f = F.map f F.coreprX := +theorem coreprW_app_hom (X : C) (f : F.coreprX ⟶ X) : (F.coreprW.app X).hom f = F.map f F.coreprx := by - change F.corepr_f.app X f = (F.corepr_f.app F.corepr_X ≫ F.map f) (𝟙 F.corepr_X) - rw [← F.corepr_f.naturality] + change F.coreprF.app X f = (F.coreprF.app F.coreprX ≫ F.map f) (𝟙 F.coreprX) + rw [← F.coreprF.naturality] dsimp simp #align category_theory.functor.corepr_w_app_hom CategoryTheory.Functor.coreprW_app_hom @@ -293,16 +307,16 @@ end Functor theorem representable_of_nat_iso (F : Cᵒᵖ ⥤ Type v₁) {G} (i : F ≅ G) [F.Representable] : G.Representable := - { has_representation := ⟨F.reprX, F.reprF ≫ i.Hom, inferInstance⟩ } + { has_representation := ⟨F.reprX, F.reprF ≫ i.hom, inferInstance⟩ } #align category_theory.representable_of_nat_iso CategoryTheory.representable_of_nat_iso theorem corepresentable_of_nat_iso (F : C ⥤ Type v₁) {G} (i : F ≅ G) [F.Corepresentable] : G.Corepresentable := - { has_corepresentation := ⟨op F.coreprX, F.coreprF ≫ i.Hom, inferInstance⟩ } + { has_corepresentation := ⟨op F.coreprX, F.coreprF ≫ i.hom, inferInstance⟩ } #align category_theory.corepresentable_of_nat_iso CategoryTheory.corepresentable_of_nat_iso instance : Functor.Corepresentable (𝟭 (Type v₁)) := - corepresentable_of_nat_iso (coyoneda.obj (op PUnit)) coyoneda.punitIso + corepresentable_of_nat_iso (coyoneda.obj (op PUnit)) Coyoneda.punitIso open Opposite @@ -354,29 +368,29 @@ See . -/ def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C where - Hom := + hom := { app := fun F x => ULift.up ((x.app F.1) (𝟙 (unop F.1))) - naturality' := by - intro X Y f; ext; dsimp - erw [category.id_comp, ← functor_to_types.naturality] - simp only [category.comp_id, yoneda_obj_map] } + naturality := by + intro X Y f; funext; dsimp + erw [Category.id_comp, ← FunctorToTypes.naturality] + simp only [Category.comp_id, yoneda_obj_map] } inv := { app := fun F x => { app := fun X a => (F.2.map a.op) x.down - naturality' := by - intro X Y f; ext; dsimp - rw [functor_to_types.map_comp_apply] } - naturality' := by - intro X Y f; ext; dsimp - rw [← functor_to_types.naturality, functor_to_types.map_comp_apply] } - hom_inv_id' := by + naturality := by + intro X Y f; funext; dsimp + rw [FunctorToTypes.map_comp_apply] } + naturality := by + intro X Y f; funext; dsimp + rw [← FunctorToTypes.naturality, FunctorToTypes.map_comp_apply] } + hom_inv_id := by ext; dsimp - erw [← functor_to_types.naturality, obj_map_id] + erw [← FunctorToTypes.naturality, obj_map_id] simp only [yoneda_map_app, Quiver.Hom.unop_op] - erw [category.id_comp] - inv_hom_id' := by + erw [Category.id_comp] + inv_hom_id := by ext; dsimp - rw [functor_to_types.map_id_apply] + rw [FunctorToTypes.map_id_apply] #align category_theory.yoneda_lemma CategoryTheory.yonedaLemma variable {C} @@ -427,7 +441,7 @@ def yonedaSectionsSmall {C : Type u₁} [SmallCategory C] (X : C) (F : Cᵒᵖ @[simp] theorem yonedaSectionsSmall_hom {C : Type u₁} [SmallCategory C] (X : C) (F : Cᵒᵖ ⥤ Type u₁) - (f : yoneda.obj X ⟶ F) : (yonedaSectionsSmall X F).Hom f = f.app _ (𝟙 _) := + (f : yoneda.obj X ⟶ F) : (yonedaSectionsSmall X F).hom f = f.app _ (𝟙 _) := rfl #align category_theory.yoneda_sections_small_hom CategoryTheory.yonedaSectionsSmall_hom @@ -443,21 +457,21 @@ attribute [local ext] Functor.ext /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma {C : Type u₁} [SmallCategory C] : (yoneda.op ⋙ coyoneda : Cᵒᵖ ⥤ (Cᵒᵖ ⥤ Type u₁) ⥤ Type u₁) ≅ evaluation Cᵒᵖ (Type u₁) := - eqToIso (by tidy) ≪≫ + eqToIso (by aesop_cat) ≪≫ curry.mapIso (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial) ≪≫ - eqToIso (by tidy) + eqToIso (by aesop_cat) #align category_theory.curried_yoneda_lemma CategoryTheory.curriedYonedaLemma /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma' {C : Type u₁} [SmallCategory C] : yoneda ⋙ (whiskeringLeft Cᵒᵖ (Cᵒᵖ ⥤ Type u₁)ᵒᵖ (Type u₁)).obj yoneda.op ≅ 𝟭 (Cᵒᵖ ⥤ Type u₁) := - eqToIso (by tidy) ≪≫ + eqToIso (by aesop_cat) ≪≫ curry.mapIso (isoWhiskerLeft (Prod.swap _ _) (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial : _)) ≪≫ - eqToIso (by tidy) + eqToIso (by aesop_cat) #align category_theory.curried_yoneda_lemma' CategoryTheory.curriedYonedaLemma' end CategoryTheory From dadd6da033ff3197bceb38937a855f7aff064df6 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 09:40:49 -0500 Subject: [PATCH 074/126] minor fixes --- Mathlib/CategoryTheory/Yoneda.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 77f23681c5a97..1b1059ae066e9 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -31,7 +31,7 @@ open Opposite universe v₁ u₁ u₂ --- morphism levels before object levels. See note [category_theory universes]. +-- morphism levels before object levels. See note [CategoryTheory universes]. variable {C : Type u₁} [Category.{v₁} C] /-- The Yoneda embedding, as a functor from `C` into presheaves on `C`. @@ -48,9 +48,9 @@ def yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁ map_id := fun Y => by funext; dsimp; erw [Category.id_comp] } map f := { app := fun Y g => g ≫ f - naturality := sorry } + naturality := fun Y Y' g => by funext Z; aesop_cat } map_id := by aesop_cat - map_comp := sorry + map_comp := fun f g => by ext Y; dsimp; funext f; simp #align category_theory.yoneda CategoryTheory.yoneda /-- The co-Yoneda embedding, as a functor from `Cᵒᵖ` into co-presheaves on `C`. @@ -61,11 +61,11 @@ def coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁ obj X := { obj := fun Y => unop X ⟶ Y map := fun f g => g ≫ f - map_comp := sorry - map_id := sorry } + map_comp := fun f g => by funext; dsimp; erw [Category.assoc] + map_id := fun Y => by funext; dsimp; erw [Category.comp_id] } map f := { app := fun Y g => f.unop ≫ g - naturality := sorry } + naturality := fun Y Y' g => by funext Z; aesop_cat } #align category_theory.coyoneda CategoryTheory.coyoneda namespace Yoneda @@ -88,7 +88,7 @@ See . -/ instance yonedaFull : Full (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) where preimage {X} {Y} f := f.app (op X) (𝟙 X) - witness := sorry + witness {X} {Y} f := by dsimp; sorry #align category_theory.yoneda.yoneda_full CategoryTheory.Yoneda.yonedaFull /-- The Yoneda embedding is faithful. From 8981ee1d62f837168758c37b572af2f5d04fe6f2 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Wed, 15 Feb 2023 12:29:50 -0500 Subject: [PATCH 075/126] fix all but four --- Mathlib/CategoryTheory/Yoneda.lean | 43 ++++++++++++++++-------------- 1 file changed, 23 insertions(+), 20 deletions(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 1b1059ae066e9..dfb808624746b 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -18,7 +18,7 @@ import Mathlib.CategoryTheory.Products.Basic The Yoneda embedding as a functor `yoneda : C ⥤ (Cᵒᵖ ⥤ Type v₁)`, along with an instance that it is `FullyFaithful`. -Also the Yoneda lemma, `yoneda_lemma : (yoneda_pairing C) ≅ (yoneda_evaluation C)`. +Also the Yoneda lemma, `yonedaLemma : (yoneda_pairing C) ≅ (yoneda_evaluation C)`. ## References * [Stacks: Opposite Categories and the Yoneda Lemma](https://stacks.math.columbia.edu/tag/001L) @@ -88,7 +88,7 @@ See . -/ instance yonedaFull : Full (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁) where preimage {X} {Y} f := f.app (op X) (𝟙 X) - witness {X} {Y} f := by dsimp; sorry + witness {X} {Y} f := by simp only [yoneda]; aesop_cat #align category_theory.yoneda.yoneda_full CategoryTheory.Yoneda.yonedaFull /-- The Yoneda embedding is faithful. @@ -116,8 +116,10 @@ def ext (X Y : C) (p : ∀ {Z : C}, (Z ⟶ X) → (Z ⟶ Y)) (q : ∀ {Z : C}, ( (NatIso.ofComponents (fun Z => { hom := p - inv := q }) - (by aesop_cat)) + inv := q + hom_inv_id := by simp only [yoneda]; funext; apply h₁ + inv_hom_id := by simp only [yoneda]; funext; apply h₂ }) + (fun f => by simp only [yoneda]; funext; apply n)) #align category_theory.yoneda.ext CategoryTheory.Yoneda.ext /-- If `yoneda.map f` is an isomorphism, so was `f`. @@ -138,12 +140,12 @@ theorem naturality {X Y : Cᵒᵖ} (α : coyoneda.obj X ⟶ coyoneda.obj Y) {Z Z instance coyonedaFull : Full (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) where - preimage {X} {Y} f := (f.app _ (𝟙 X.unop)).op - witness := sorry + preimage {X} _ f := (f.app _ (𝟙 X.unop)).op + witness {X} {Y} f := by simp only [coyoneda]; aesop_cat #align category_theory.coyoneda.coyoneda_full CategoryTheory.Coyoneda.coyonedaFull instance coyoneda_faithful : Faithful (coyoneda : Cᵒᵖ ⥤ C ⥤ Type v₁) where - map_injective g p := by + map_injective {X} _ _ _ p := by have t := congr_fun (congr_app p X.unop) (𝟙 _) simpa using congr_arg Quiver.Hom.op t #align category_theory.coyoneda.coyoneda_faithful CategoryTheory.Coyoneda.coyoneda_faithful @@ -164,7 +166,7 @@ def punitIso : coyoneda.obj (Opposite.op PUnit) ≅ 𝟭 (Type v₁) := #align category_theory.coyoneda.punit_iso CategoryTheory.Coyoneda.punitIso /-- Taking the `unop` of morphisms is a natural isomorphism. -/ -@[simps] +@[simps!] def objOpOp (X : C) : coyoneda.obj (op (op X)) ≅ yoneda.obj X := NatIso.ofComponents (fun _ => (opEquiv _ _).toIso) fun _ {_Y} _f => rfl #align category_theory.coyoneda.obj_op_op CategoryTheory.Coyoneda.objOpOp @@ -371,8 +373,8 @@ def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C hom := { app := fun F x => ULift.up ((x.app F.1) (𝟙 (unop F.1))) naturality := by - intro X Y f; funext; dsimp - erw [Category.id_comp, ← FunctorToTypes.naturality] + intro X Y f; simp only [yonedaEvaluation]; funext; dsimp + erw [Category.id_comp, ←FunctorToTypes.naturality] simp only [Category.comp_id, yoneda_obj_map] } inv := { app := fun F x => @@ -381,16 +383,17 @@ def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C intro X Y f; funext; dsimp rw [FunctorToTypes.map_comp_apply] } naturality := by - intro X Y f; funext; dsimp - rw [← FunctorToTypes.naturality, FunctorToTypes.map_comp_apply] } + intro X Y f; simp only [yoneda]; funext; apply NatTrans.ext; funext; funext; dsimp + rw [←FunctorToTypes.naturality X.snd Y.snd f.snd, FunctorToTypes.map_comp_apply] } hom_inv_id := by - ext; dsimp + apply NatTrans.ext; funext; funext; + apply NatTrans.ext; funext; funext; dsimp erw [← FunctorToTypes.naturality, obj_map_id] simp only [yoneda_map_app, Quiver.Hom.unop_op] erw [Category.id_comp] inv_hom_id := by - ext; dsimp - rw [FunctorToTypes.map_id_apply] + apply NatTrans.ext; funext; funext; dsimp + rw [FunctorToTypes.map_id_apply, ULift.up_down] #align category_theory.yoneda_lemma CategoryTheory.yonedaLemma variable {C} @@ -399,7 +402,7 @@ variable {C} (we need to insert a `ulift` to get the universes right!) given by the Yoneda lemma. -/ -@[simps] +@[simps!] def yonedaSections (X : C) (F : Cᵒᵖ ⥤ Type v₁) : (yoneda.obj X ⟶ F) ≅ ULift.{u₁} (F.obj (op X)) := (yonedaLemma C).app (op X, F) #align category_theory.yoneda_sections CategoryTheory.yonedaSections @@ -457,21 +460,21 @@ attribute [local ext] Functor.ext /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma {C : Type u₁} [SmallCategory C] : (yoneda.op ⋙ coyoneda : Cᵒᵖ ⥤ (Cᵒᵖ ⥤ Type u₁) ⥤ Type u₁) ≅ evaluation Cᵒᵖ (Type u₁) := - eqToIso (by aesop_cat) ≪≫ + eqToIso (by apply Functor.ext; intro X Y f; sorry; sorry ) ≪≫ curry.mapIso (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial) ≪≫ - eqToIso (by aesop_cat) + eqToIso (by ext X; intro Y f; sorry; sorry; sorry;) #align category_theory.curried_yoneda_lemma CategoryTheory.curriedYonedaLemma /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma' {C : Type u₁} [SmallCategory C] : yoneda ⋙ (whiskeringLeft Cᵒᵖ (Cᵒᵖ ⥤ Type u₁)ᵒᵖ (Type u₁)).obj yoneda.op ≅ 𝟭 (Cᵒᵖ ⥤ Type u₁) := - eqToIso (by aesop_cat) ≪≫ + eqToIso (by sorry) ≪≫ curry.mapIso (isoWhiskerLeft (Prod.swap _ _) (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial : _)) ≪≫ - eqToIso (by aesop_cat) + eqToIso (by sorry) #align category_theory.curried_yoneda_lemma' CategoryTheory.curriedYonedaLemma' end CategoryTheory From 5b11f3d9912b56e6f758fc83ca9cfa78f48770f3 Mon Sep 17 00:00:00 2001 From: mpenciak Date: Wed, 15 Feb 2023 14:29:42 -0500 Subject: [PATCH 076/126] 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 077/126] 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 078/126] 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 079/126] 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 080/126] 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 081/126] 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 082/126] 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 083/126] 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 084/126] 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 085/126] 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 086/126] 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 087/126] 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 088/126] 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 089/126] 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 090/126] 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 091/126] 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 092/126] 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 093/126] 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 8860e0dbee2c5b58a3e890ba3447af0443052316 Mon Sep 17 00:00:00 2001 From: adamtopaz Date: Thu, 16 Feb 2023 17:33:00 -0700 Subject: [PATCH 094/126] get file to build --- Mathlib/CategoryTheory/DiscreteCategory.lean | 148 ++++++++++--------- 1 file changed, 81 insertions(+), 67 deletions(-) diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 871c80e12f731..eec8f2d961faf 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -9,7 +9,10 @@ Authors: Stephen Morgan, Scott Morrison, Floris van Doorn ! if you have ported upstream changes. -/ import Mathlib.CategoryTheory.EqToHom -import Mathlib.Data.Ulift +import Mathlib.Data.ULift +import Mathlib.Tactic.CasesM + +open Lean Meta /-! # Discrete categories @@ -65,12 +68,12 @@ def discreteEquiv {α : Type u₁} : Discrete α ≃ α where toFun := Discrete.as invFun := Discrete.mk - left_inv := by tidy - right_inv := by tidy + left_inv := by aesop_cat + right_inv := by aesop_cat #align category_theory.discrete_equiv CategoryTheory.discreteEquiv instance {α : Type u₁} [DecidableEq α] : DecidableEq (Discrete α) := - discreteEquiv.DecidableEq + discreteEquiv.decidableEq /-- The "discrete" category on a type, whose morphisms are equalities. @@ -83,7 +86,7 @@ instance discreteCategory (α : Type u₁) : SmallCategory (Discrete α) where Hom X Y := ULift (PLift (X.as = Y.as)) id X := ULift.up (PLift.up rfl) - comp X Y Z g f := by + comp {X Y Z} g f := by cases X cases Y cases Z @@ -106,14 +109,22 @@ instance [Subsingleton α] : Subsingleton (Discrete α) := /- ./././Mathport/Syntax/Translate/Expr.lean:334:4: warning: unsupported (TODO): `[tacs] -/ /-- A simple tactic to run `cases` on any `discrete α` hypotheses. -/ -unsafe def _root_.tactic.discrete_cases : tactic Unit := - sorry -#align tactic.discrete_cases tactic.discrete_cases +--unsafe def _root_.tactic.discrete_cases : tactic Unit := +-- sorry +--#align tactic.discrete_cases tactic.discrete_cases + + +--run_cmd +-- add_interactive [`` tactic.discrete_cases] -run_cmd - add_interactive [`` tactic.discrete_cases] +--attribute [local tidy] tactic.discrete_cases +--`[cases_matching* [discrete _, (_ : discrete _) ⟶ (_ : discrete _), plift _]] -attribute [local tidy] tactic.discrete_cases +--macro (name := discrete_cases) "discrete_cases" : tactic => +-- `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) + +macro "discrete_cases": tactic => + `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) instance [Unique α] : Unique (Discrete α) := Unique.mk' (Discrete α) @@ -125,7 +136,7 @@ theorem eq_of_hom {X Y : Discrete α} (i : X ⟶ Y) : X.as = Y.as := /-- Promote an equation between the wrapped terms in `X Y : discrete α` to a morphism `X ⟶ Y` in the discrete category. -/ -abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := +protected abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := eqToHom (by ext @@ -134,7 +145,7 @@ abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := /-- Promote an equation between the wrapped terms in `X Y : discrete α` to an isomorphism `X ≅ Y` in the discrete category. -/ -abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := +protected abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := eqToIso (by ext @@ -143,12 +154,12 @@ abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := /-- A variant of `eq_to_hom` that lifts terms to the discrete category. -/ abbrev eqToHom' {a b : α} (h : a = b) : Discrete.mk a ⟶ Discrete.mk b := - eqToHom h + Discrete.eqToHom h #align category_theory.discrete.eq_to_hom' CategoryTheory.Discrete.eqToHom' /-- A variant of `eq_to_iso` that lifts terms to the discrete category. -/ abbrev eqToIso' {a b : α} (h : a = b) : Discrete.mk a ≅ Discrete.mk b := - eqToIso h + Discrete.eqToIso h #align category_theory.discrete.eq_to_iso' CategoryTheory.Discrete.eqToIso' @[simp] @@ -159,7 +170,7 @@ theorem id_def (X : Discrete α) : ULift.up (PLift.up (Eq.refl X.as)) = 𝟙 X : variable {C : Type u₂} [Category.{v₂} C] instance {I : Type u₁} {i j : Discrete I} (f : i ⟶ j) : IsIso f := - ⟨⟨eqToHom (eq_of_hom f).symm, by tidy⟩⟩ + ⟨⟨Discrete.eqToHom (eq_of_hom f).symm, by aesop_cat⟩⟩ /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ /-- Any function `I → C` gives a functor `discrete I ⥤ C`. @@ -167,12 +178,14 @@ instance {I : Type u₁} {i j : Discrete I} (f : i ⟶ j) : IsIso f := def functor {I : Type u₁} (F : I → C) : Discrete I ⥤ C where obj := F ∘ Discrete.as - map X Y f := - by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - cases f - exact 𝟙 (F X) + map {X Y} f := by + dsimp + rcases f with ⟨⟨h⟩⟩ + exact eqToHom (congrArg _ h) + map_id := by aesop_cat + map_comp := fun {X Y Z} f g => by + discrete_cases + aesop_cat #align category_theory.discrete.functor CategoryTheory.Discrete.functor @[simp] @@ -182,7 +195,7 @@ theorem functor_obj {I : Type u₁} (F : I → C) (i : I) : #align category_theory.discrete.functor_obj CategoryTheory.Discrete.functor_obj theorem functor_map {I : Type u₁} (F : I → C) {i : Discrete I} (f : i ⟶ i) : - (Discrete.functor F).map f = 𝟙 (F i.as) := by tidy + (Discrete.functor F).map f = 𝟙 (F i.as) := by aesop_cat #align category_theory.discrete.functor_map CategoryTheory.Discrete.functor_map /-- The discrete functor induced by a composition of maps can be written as a @@ -191,7 +204,7 @@ composition of two discrete functors. @[simps] def functorComp {I : Type u₁} {J : Type u₁'} (f : J → C) (g : I → J) : Discrete.functor (f ∘ g) ≅ Discrete.functor (Discrete.mk ∘ g) ⋙ Discrete.functor f := - NatIso.ofComponents (fun X => Iso.refl _) (by tidy) + NatIso.ofComponents (fun X => Iso.refl _) (by aesop_cat) #align category_theory.discrete.functor_comp CategoryTheory.Discrete.functorComp /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ @@ -203,11 +216,11 @@ as the naturality squares are trivial. def natTrans {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ⟶ G.obj i) : F ⟶ G where app := f - naturality' X Y g := - by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - cases g + naturality {X Y} g := by + rcases g with ⟨⟨g⟩⟩ + discrete_cases + rcases g + change F.map (𝟙 _) ≫ _ = _ ≫ G.map (𝟙 _) simp #align category_theory.discrete.nat_trans CategoryTheory.Discrete.natTrans @@ -218,17 +231,17 @@ as the naturality squares are trivial. -/ @[simps] def natIso {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) : F ≅ G := - NatIso.ofComponents f fun X Y g => - by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - cases g + NatIso.ofComponents f fun g => by + rcases g with ⟨⟨g⟩⟩ + discrete_cases + rcases g + change F.map (𝟙 _) ≫ _ = _ ≫ G.map (𝟙 _) simp #align category_theory.discrete.nat_iso CategoryTheory.Discrete.natIso @[simp] theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) - (i : Discrete I) : (Discrete.natIso f).app i = f i := by tidy + (i : Discrete I) : (Discrete.natIso f).app i = f i := by aesop_cat #align category_theory.discrete.nat_iso_app CategoryTheory.Discrete.natIso_app /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ @@ -236,11 +249,7 @@ theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discret `discrete.functor (F.obj)`. -/ @[simp] def natIsoFunctor {I : Type u₁} {F : Discrete I ⥤ C} : F ≅ Discrete.functor (F.obj ∘ Discrete.mk) := - natIso fun i => - by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - rfl + natIso fun _ => Iso.refl _ #align category_theory.discrete.nat_iso_functor CategoryTheory.Discrete.natIsoFunctor /-- Composing `discrete.functor F` with another functor `G` amounts to composing `F` with `G.obj` -/ @@ -258,21 +267,19 @@ an equivalence between the corresponding `discrete` categories. @[simps] def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ Discrete J where - Functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) + functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) inverse := Discrete.functor (Discrete.mk ∘ (e.symm : J → I)) unitIso := Discrete.natIso fun i => eqToIso (by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + discrete_cases simp) counitIso := Discrete.natIso fun j => eqToIso (by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" + discrete_cases simp) #align category_theory.discrete.equivalence CategoryTheory.Discrete.equivalence @@ -280,10 +287,10 @@ def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ D @[simps] def equivOfEquivalence {α : Type u₁} {β : Type u₂} (h : Discrete α ≌ Discrete β) : α ≃ β where - toFun := Discrete.as ∘ h.Functor.obj ∘ Discrete.mk + toFun := Discrete.as ∘ h.functor.obj ∘ Discrete.mk invFun := Discrete.as ∘ h.inverse.obj ∘ Discrete.mk - left_inv a := by simpa using eq_of_hom (h.unit_iso.app (discrete.mk a)).2 - right_inv a := by simpa using eq_of_hom (h.counit_iso.app (discrete.mk a)).1 + left_inv a := by simpa using eq_of_hom (h.unitIso.app (Discrete.mk a)).2 + right_inv a := by simpa using eq_of_hom (h.counitIso.app (Discrete.mk a)).1 #align category_theory.discrete.equiv_of_equivalence CategoryTheory.Discrete.equivOfEquivalence end Discrete @@ -298,27 +305,36 @@ open Opposite /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:72:18: unsupported non-interactive tactic tactic.op_induction' -/ /- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ /-- A discrete category is equivalent to its opposite category. -/ -@[simps functor_obj_as inverse_obj] -protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := by +@[simps! functor_obj_as inverse_obj] +protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := let F : Discrete α ⥤ (Discrete α)ᵒᵖ := Discrete.functor fun x => op (Discrete.mk x) + Equivalence.mk F.leftOp F + (NatIso.ofComponents (fun ⟨X⟩ => Iso.refl _) + <| fun {X Y} f => by + induction X using Opposite.rec + induction Y using Opposite.rec + rcases f with ⟨⟨f⟩⟩ + discrete_cases + rcases f + aesop_cat) + (Discrete.natIso <| fun ⟨X⟩ => Iso.refl _) +/- refine' - equivalence.mk (functor.left_op F) F _ - (discrete.nat_iso fun X => + Equivalence.mk (F.leftOp) F _ + (Discrete.natIso fun X => by - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - simp [F]) + discrete_cases + rfl) refine' - nat_iso.of_components + NatIso.ofComponents (fun X => by - run_tac - tactic.op_induction' - trace - "./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[]" - simp [F]) - _ - tidy + discrete_cases + induction X using Opposite.rec + discrete_cases + exact Iso.refl _) +-/ + #align category_theory.discrete.opposite CategoryTheory.Discrete.opposite variable {C : Type u₂} [Category.{v₂} C] @@ -327,9 +343,8 @@ variable {C : Type u₂} [Category.{v₂} C] theorem functor_map_id (F : Discrete J ⥤ C) {j : Discrete J} (f : j ⟶ j) : F.map f = 𝟙 (F.obj j) := by have h : f = 𝟙 j := by - cases f - cases f - ext + rcases f with ⟨⟨f⟩⟩ + rfl rw [h] simp #align category_theory.discrete.functor_map_id CategoryTheory.Discrete.functor_map_id @@ -337,4 +352,3 @@ theorem functor_map_id (F : Discrete J ⥤ C) {j : Discrete J} (f : j ⟶ j) : F end Discrete end CategoryTheory - From 118d7ed86fd869f52d3f1a07923a4ebcf2242be8 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 19:57:39 -0500 Subject: [PATCH 095/126] 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 096/126] 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 eb1464469358a6a9502a0b8ddaba399068e45293 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:14:21 -0500 Subject: [PATCH 097/126] lint --- Mathlib/CategoryTheory/DiscreteCategory.lean | 66 +++++++++++--------- 1 file changed, 38 insertions(+), 28 deletions(-) diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index eec8f2d961faf..c19ec36629f0c 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -17,22 +17,22 @@ open Lean Meta /-! # Discrete categories -We define `discrete α` as a structure containing a term `a : α` for any type `α`, -and use this type alias to provide a `small_category` instance +We define `Discrete α` as a structure containing a term `a : α` for any type `α`, +and use this type alias to provide a `SmallCategory` instance whose only morphisms are the identities. There is an annoying technical difficulty that it has turned out to be inconvenient to allow categories with morphisms living in `Prop`, -so instead of defining `X ⟶ Y` in `discrete α` as `X = Y`, -one might define it as `plift (X = Y)`. -In fact, to allow `discrete α` to be a `small_category` +so instead of defining `X ⟶ Y` in `Discrete α` as `X = Y`, +one might define it as `PLift (X = Y)`. +In fact, to allow `Discrete α` to be a `SmallCategory` (i.e. with morphisms in the same universe as the objects), -we actually define the hom type `X ⟶ Y` as `ulift (plift (X = Y))`. +we actually define the hom type `X ⟶ Y` as `ULift (PLift (X = Y))`. -`discrete.functor` promotes a function `f : I → C` (for any category `C`) to a functor -`discrete.functor f : discrete I ⥤ C`. +`Discrete.functor` promotes a function `f : I → C` (for any category `C`) to a functor +`Discrete.functor f : Discrete I ⥤ C`. -Similarly, `discrete.nat_trans` and `discrete.nat_iso` promote `I`-indexed families of morphisms, +Similarly, `Discrete.natTrans` and `Discrete.natIso` promote `I`-indexed families of morphisms, or `I`-indexed families of isomorphisms to natural transformations or natural isomorphism. We show equivalences of types are the same as (categorical) equivalences of the corresponding @@ -42,17 +42,20 @@ discrete categories. namespace CategoryTheory --- morphism levels before object levels. See note [category_theory universes]. +-- morphism levels before object levels. See note [CategoryTheory universes]. universe v₁ v₂ v₃ u₁ u₁' u₂ u₃ -- This is intentionally a structure rather than a type synonym --- to enforce using `discrete_equiv` (or `discrete.mk` and `discrete.as`) to move between +-- to enforce using `DiscreteEquiv` (or `Discrete.mk` and `Discrete.as`) to move between -- `discrete α` and `α`. Otherwise there is too much API leakage. /-- A wrapper for promoting any type to a category, with the only morphisms being equalities. -/ @[ext] structure Discrete (α : Type u₁) where + /-- A wrapper for promoting any type to a category, + with the only morphisms being equalities. + -/ as : α #align category_theory.discrete CategoryTheory.Discrete @@ -62,7 +65,7 @@ theorem Discrete.mk_as {α : Type u₁} (X : Discrete α) : Discrete.mk X.as = X rfl #align category_theory.discrete.mk_as CategoryTheory.Discrete.mk_as -/-- `discrete α` is equivalent to the original type `α`.-/ +/-- `Discrete α` is equivalent to the original type `α`.-/ @[simps] def discreteEquiv {α : Type u₁} : Discrete α ≃ α where @@ -75,10 +78,10 @@ def discreteEquiv {α : Type u₁} : Discrete α ≃ α instance {α : Type u₁} [DecidableEq α] : DecidableEq (Discrete α) := discreteEquiv.decidableEq -/-- The "discrete" category on a type, whose morphisms are equalities. +/-- The "Discrete" category on a type, whose morphisms are equalities. Because we do not allow morphisms in `Prop` (only in `Type`), -somewhat annoyingly we have to define `X ⟶ Y` as `ulift (plift (X = Y))`. +somewhat annoyingly we have to define `X ⟶ Y` as `ULift (PLift (X = Y))`. See -/ @@ -122,7 +125,7 @@ instance [Subsingleton α] : Subsingleton (Discrete α) := --macro (name := discrete_cases) "discrete_cases" : tactic => -- `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) - +/- Porting note: rewrote `discrete_cases` tactic -/ macro "discrete_cases": tactic => `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) @@ -134,7 +137,7 @@ theorem eq_of_hom {X Y : Discrete α} (i : X ⟶ Y) : X.as = Y.as := i.down.down #align category_theory.discrete.eq_of_hom CategoryTheory.Discrete.eq_of_hom -/-- Promote an equation between the wrapped terms in `X Y : discrete α` to a morphism `X ⟶ Y` +/-- Promote an equation between the wrapped terms in `X Y : Discrete α` to a morphism `X ⟶ Y` in the discrete category. -/ protected abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := eqToHom @@ -143,7 +146,7 @@ protected abbrev eqToHom {X Y : Discrete α} (h : X.as = Y.as) : X ⟶ Y := exact h) #align category_theory.discrete.eq_to_hom CategoryTheory.Discrete.eqToHom -/-- Promote an equation between the wrapped terms in `X Y : discrete α` to an isomorphism `X ≅ Y` +/-- Promote an equation between the wrapped terms in `X Y : Discrete α` to an isomorphism `X ≅ Y` in the discrete category. -/ protected abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := eqToIso @@ -157,7 +160,7 @@ abbrev eqToHom' {a b : α} (h : a = b) : Discrete.mk a ⟶ Discrete.mk b := Discrete.eqToHom h #align category_theory.discrete.eq_to_hom' CategoryTheory.Discrete.eqToHom' -/-- A variant of `eq_to_iso` that lifts terms to the discrete category. -/ +/-- A variant of `eqToIso` that lifts terms to the discrete category. -/ abbrev eqToIso' {a b : α} (h : a = b) : Discrete.mk a ≅ Discrete.mk b := Discrete.eqToIso h #align category_theory.discrete.eq_to_iso' CategoryTheory.Discrete.eqToIso' @@ -172,8 +175,9 @@ variable {C : Type u₂} [Category.{v₂} C] instance {I : Type u₁} {i j : Discrete I} (f : i ⟶ j) : IsIso f := ⟨⟨Discrete.eqToHom (eq_of_hom f).symm, by aesop_cat⟩⟩ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ -/-- Any function `I → C` gives a functor `discrete I ⥤ C`. +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ +/-- Any function `I → C` gives a functor `Discrete I ⥤ C`. -/ def functor {I : Type u₁} (F : I → C) : Discrete I ⥤ C where @@ -201,13 +205,14 @@ theorem functor_map {I : Type u₁} (F : I → C) {i : Discrete I} (f : i ⟶ i) /-- The discrete functor induced by a composition of maps can be written as a composition of two discrete functors. -/ -@[simps] +@[simps!] def functorComp {I : Type u₁} {J : Type u₁'} (f : J → C) (g : I → J) : Discrete.functor (f ∘ g) ≅ Discrete.functor (Discrete.mk ∘ g) ⋙ Discrete.functor f := NatIso.ofComponents (fun X => Iso.refl _) (by aesop_cat) #align category_theory.discrete.functor_comp CategoryTheory.Discrete.functorComp -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ /-- For functors out of a discrete category, a natural transformation is just a collection of maps, as the naturality squares are trivial. @@ -224,12 +229,13 @@ def natTrans {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F simp #align category_theory.discrete.nat_trans CategoryTheory.Discrete.natTrans -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ /-- For functors out of a discrete category, a natural isomorphism is just a collection of isomorphisms, as the naturality squares are trivial. -/ -@[simps] +@[simps!] def natIso {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) : F ≅ G := NatIso.ofComponents f fun g => by rcases g with ⟨⟨g⟩⟩ @@ -244,7 +250,8 @@ theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discret (i : Discrete I) : (Discrete.natIso f).app i = f i := by aesop_cat #align category_theory.discrete.nat_iso_app CategoryTheory.Discrete.natIso_app -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ /-- Every functor `F` from a discrete category is naturally isomorphic (actually, equal) to `discrete.functor (F.obj)`. -/ @[simp] @@ -256,11 +263,13 @@ def natIsoFunctor {I : Type u₁} {F : Discrete I ⥤ C} : F ≅ Discrete.functo @[simp] def compNatIsoDiscrete {I : Type u₁} {D : Type u₃} [Category.{v₃} D] (F : I → C) (G : C ⥤ D) : Discrete.functor F ⋙ G ≅ Discrete.functor (G.obj ∘ F) := - natIso fun i => Iso.refl _ + natIso fun _ => Iso.refl _ #align category_theory.discrete.comp_nat_iso_discrete CategoryTheory.Discrete.compNatIsoDiscrete -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ /-- We can promote a type-level `equiv` to an equivalence between the corresponding `discrete` categories. -/ @@ -352,3 +361,4 @@ theorem functor_map_id (F : Discrete J ⥤ C) {j : Discrete J} (f : j ⟶ j) : F end Discrete end CategoryTheory + From 38a89b87e6655d9e89c6b89032e22096d94d24b5 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:20:24 -0500 Subject: [PATCH 098/126] lint some more --- Mathlib/CategoryTheory/DiscreteCategory.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index c19ec36629f0c..cac60eb13ddee 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -12,8 +12,6 @@ import Mathlib.CategoryTheory.EqToHom import Mathlib.Data.ULift import Mathlib.Tactic.CasesM -open Lean Meta - /-! # Discrete categories @@ -39,7 +37,6 @@ We show equivalences of types are the same as (categorical) equivalences of the discrete categories. -/ - namespace CategoryTheory -- morphism levels before object levels. See note [CategoryTheory universes]. @@ -310,9 +307,12 @@ variable {J : Type v₁} open Opposite -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:72:18: unsupported non-interactive tactic tactic.op_induction' -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:72:18: unsupported +\ non-interactive tactic tactic.op_induction' -/ +/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported +\ tactic `discrete_cases #[] -/ /-- A discrete category is equivalent to its opposite category. -/ @[simps! functor_obj_as inverse_obj] protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := From 12c4a61f8b1e0a74354e0e36a022320f822b915c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 20:28:27 -0500 Subject: [PATCH 099/126] 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 100/126] 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 101/126] 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 102/126] 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 103/126] 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 ea56ec586f895b600d33d1af80a55e08296a696f Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> Date: Thu, 16 Feb 2023 20:53:18 -0500 Subject: [PATCH 104/126] 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 105/126] 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 106/126] 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 107/126] 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 108/126] 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 109/126] 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 110/126] 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 111/126] 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 112/126] 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 From 37804c731698bd81deec996e09179b50799ed7ce Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 22:48:37 -0500 Subject: [PATCH 113/126] fix last errors; lint --- Mathlib/CategoryTheory/Yoneda.lean | 101 +++++++++++++++++++++-------- 1 file changed, 75 insertions(+), 26 deletions(-) diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index dfb808624746b..7773489921444 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -180,6 +180,7 @@ namespace Functor See . -/ class Representable (F : Cᵒᵖ ⥤ Type v₁) : Prop where + /-- `Hom(-,X) ≅ F` via `f` -/ has_representation : ∃ (X : _)(f : yoneda.obj X ⟶ F), IsIso f #align category_theory.functor.representable CategoryTheory.Functor.Representable @@ -190,6 +191,7 @@ instance {X : C} : Representable (yoneda.obj X) where has_representation := ⟨X See . -/ class Corepresentable (F : C ⥤ Type v₁) : Prop where + /-- `Hom(X,-) ≅ F` via `f` -/ has_corepresentation : ∃ (X : _)(f : coyoneda.obj X ⟶ F), IsIso f #align category_theory.functor.corepresentable CategoryTheory.Functor.Corepresentable @@ -215,13 +217,24 @@ noncomputable def reprF : yoneda.obj F.reprX ⟶ F := Representable.has_representation.choose_spec.choose #align category_theory.functor.repr_f CategoryTheory.Functor.reprF -/- warning: category_theory.functor.repr_x clashes with category_theory.functor.repr_X -> CategoryTheory.Functor.reprX -warning: category_theory.functor.repr_x -> CategoryTheory.Functor.reprX is a dubious translation: -lean 3 declaration is - forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1} F (Opposite.op.{succ u2} C (CategoryTheory.Functor.reprX.{u1, u2} C _inst_1 F _inst_2)) -but is expected to have type - forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], C -Case conversion may be inaccurate. Consider using '#align category_theory.functor.repr_x CategoryTheory.Functor.reprXₓ'. -/ +/- warning: category_theory.functor.repr_x clashes with +\ category_theory.functor.repr_X -> CategoryTheory.Functor.reprX +warning: category_theory.functor.repr_x -> CategoryTheory.Functor.reprX is a +dubious translation: lean 3 declaration is forall {C : Type.{u2}} [_inst_1 : +CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, +succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C +_inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : +CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], +CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) +(CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} +CategoryTheory.types.{u1} F (Opposite.op.{succ u2} C +(CategoryTheory.Functor.reprX.{u1, u2} C _inst_1 F _inst_2)) but is expected to +have type forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] +(F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) +(CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} +CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, +u2} C _inst_1 F], C Case conversion may be inaccurate. Consider using '#align +category_theory.functor.repr_x CategoryTheory.Functor.reprXₓ'. -/ /-- The representing element for the representable functor `F`, sometimes called the universal element of the functor. -/ @@ -271,13 +284,21 @@ noncomputable def coreprF : coyoneda.obj (op F.coreprX) ⟶ F := Corepresentable.has_corepresentation.choose_spec.choose #align category_theory.functor.corepr_f CategoryTheory.Functor.coreprF -/- warning: category_theory.functor.corepr_x clashes with category_theory.functor.corepr_X -> CategoryTheory.Functor.coreprX -warning: category_theory.functor.corepr_x -> CategoryTheory.Functor.coreprX is a dubious translation: -lean 3 declaration is - forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1} F (CategoryTheory.Functor.coreprX.{u1, u2} C _inst_1 F _inst_2) -but is expected to have type - forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], C -Case conversion may be inaccurate. Consider using '#align category_theory.functor.corepr_x CategoryTheory.Functor.coreprXₓ'. -/ +/- warning: category_theory.functor.corepr_x clashes with +\ category_theory.functor.corepr_X -> CategoryTheory.Functor.coreprX +warning: category_theory.functor.corepr_x -> CategoryTheory.Functor.coreprX is +a dubious translation: lean 3 declaration is forall {C : Type.{u2}} [_inst_1 : +CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, +succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : +CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], +CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} +CategoryTheory.types.{u1} F (CategoryTheory.Functor.coreprX.{u1, u2} C _inst_1 +F _inst_2) but is expected to have type forall {C : Type.{u2}} [_inst_1 : +CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, +succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : +CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], C Case conversion +may be inaccurate. Consider using '#align category_theory.functor.corepr_x +CategoryTheory.Functor.coreprXₓ'. -/ /-- The representing element for the corepresentable functor `F`, sometimes called the universal element of the functor. -/ @@ -457,24 +478,52 @@ theorem yonedaSectionsSmall_inv_app_apply {C : Type u₁} [SmallCategory C] (X : attribute [local ext] Functor.ext +/- Porting note: this used to be two calls to `tidy` -/ /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma {C : Type u₁} [SmallCategory C] : - (yoneda.op ⋙ coyoneda : Cᵒᵖ ⥤ (Cᵒᵖ ⥤ Type u₁) ⥤ Type u₁) ≅ evaluation Cᵒᵖ (Type u₁) := - eqToIso (by apply Functor.ext; intro X Y f; sorry; sorry ) ≪≫ - curry.mapIso - (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial) ≪≫ - eqToIso (by ext X; intro Y f; sorry; sorry; sorry;) + (yoneda.op ⋙ coyoneda : Cᵒᵖ ⥤ (Cᵒᵖ ⥤ Type u₁) ⥤ Type u₁) ≅ evaluation Cᵒᵖ (Type u₁) := by + refine eqToIso ?_ ≪≫ curry.mapIso + (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial) ≪≫ + eqToIso ?_ + · apply Functor.ext + · intro X Y f + simp only [curry, yoneda, coyoneda, curryObj, yonedaPairing] + dsimp + apply NatTrans.ext + dsimp at * + funext F g + apply NatTrans.ext + simp + · intro X + simp only [curry, yoneda, coyoneda, curryObj, yonedaPairing] + aesop_cat + · apply Functor.ext + · intro X Y f + simp only [curry, yoneda, coyoneda, curryObj, yonedaPairing] + dsimp + apply NatTrans.ext + dsimp at * + funext F g + simp + · intro X + simp only [curry, yoneda, coyoneda, curryObj, yonedaPairing] + aesop_cat #align category_theory.curried_yoneda_lemma CategoryTheory.curriedYonedaLemma /-- The curried version of yoneda lemma when `C` is small. -/ def curriedYonedaLemma' {C : Type u₁} [SmallCategory C] : - yoneda ⋙ (whiskeringLeft Cᵒᵖ (Cᵒᵖ ⥤ Type u₁)ᵒᵖ (Type u₁)).obj yoneda.op ≅ 𝟭 (Cᵒᵖ ⥤ Type u₁) := - eqToIso (by sorry) ≪≫ - curry.mapIso - (isoWhiskerLeft (Prod.swap _ _) - (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial : - _)) ≪≫ - eqToIso (by sorry) + yoneda ⋙ (whiskeringLeft Cᵒᵖ (Cᵒᵖ ⥤ Type u₁)ᵒᵖ (Type u₁)).obj yoneda.op ≅ 𝟭 (Cᵒᵖ ⥤ Type u₁) + := by + refine eqToIso ?_ ≪≫ curry.mapIso (isoWhiskerLeft (Prod.swap _ _) + (yonedaLemma C ≪≫ isoWhiskerLeft (evaluationUncurried Cᵒᵖ (Type u₁)) uliftFunctorTrivial :_)) + ≪≫ eqToIso ?_ + · apply Functor.ext + · intro X Y f + simp only [curry, yoneda, coyoneda, curryObj, yonedaPairing] + aesop_cat + · apply Functor.ext + · intro X Y f + aesop_cat #align category_theory.curried_yoneda_lemma' CategoryTheory.curriedYonedaLemma' end CategoryTheory From 9e93b3f3c49ec2c492c1c5f4d06cf179e5e08aef Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 22:59:11 -0500 Subject: [PATCH 114/126] feat: port CategoryTheory.Limits.Cones From 8beaf77097eabe84bffb0d26a9c99ed176a75411 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 22:59:12 -0500 Subject: [PATCH 115/126] Initial file copy from mathport --- Mathlib/CategoryTheory/Limits/Cones.lean | 1097 ++++++++++++++++++++++ 1 file changed, 1097 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Cones.lean diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean new file mode 100644 index 0000000000000..c5c1bb53927c7 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -0,0 +1,1097 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Stephen Morgan, Scott Morrison, Floris van Doorn + +! This file was ported from Lean 3 source module category_theory.limits.cones +! leanprover-community/mathlib commit 740acc0e6f9adf4423f92a485d0456fc271482da +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.CategoryTheory.Functor.Const +import Mathbin.CategoryTheory.DiscreteCategory +import Mathbin.CategoryTheory.Yoneda +import Mathbin.CategoryTheory.Functor.ReflectsIsomorphisms + +/-! +# Cones and cocones + +We define `cone F`, a cone over a functor `F`, +and `F.cones : Cᵒᵖ ⥤ Type`, the functor associating to `X` the cones over `F` with cone point `X`. + +A cone `c` is defined by specifying its cone point `c.X` and a natural transformation `c.π` +from the constant `c.X` valued functor to `F`. + +We provide `c.w f : c.π.app j ≫ F.map f = c.π.app j'` for any `f : j ⟶ j'` +as a wrapper for `c.π.naturality f` avoiding unneeded identity morphisms. + +We define `c.extend f`, where `c : cone F` and `f : Y ⟶ c.X` for some other `Y`, +which replaces the cone point by `Y` and inserts `f` into each of the components of the cone. +Similarly we have `c.whisker F` producing a `cone (E ⋙ F)` + +We define morphisms of cones, and the category of cones. + +We define `cone.postcompose α : cone F ⥤ cone G` for `α` a natural transformation `F ⟶ G`. + +And, of course, we dualise all this to cocones as well. + +For more results about the category of cones, see `cone_category.lean`. +-/ + + +-- morphism levels before object levels. See note [category_theory universes]. +universe v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ + +open CategoryTheory + +variable {J : Type u₁} [Category.{v₁} J] + +variable {K : Type u₂} [Category.{v₂} K] + +variable {C : Type u₃} [Category.{v₃} C] + +variable {D : Type u₄} [Category.{v₄} D] + +open CategoryTheory + +open CategoryTheory.Category + +open CategoryTheory.Functor + +open Opposite + +namespace CategoryTheory + +namespace Functor + +variable {J C} (F : J ⥤ C) + +/-- `F.cones` is the functor assigning to an object `X` the type of +natural transformations from the constant functor with value `X` to `F`. +An object representing this functor is a limit of `F`. +-/ +@[simps] +def cones : Cᵒᵖ ⥤ Type max u₁ v₃ := + (const J).op ⋙ yoneda.obj F +#align category_theory.functor.cones CategoryTheory.Functor.cones + +/-- `F.cocones` is the functor assigning to an object `X` the type of +natural transformations from `F` to the constant functor with value `X`. +An object corepresenting this functor is a colimit of `F`. +-/ +@[simps] +def cocones : C ⥤ Type max u₁ v₃ := + const J ⋙ coyoneda.obj (op F) +#align category_theory.functor.cocones CategoryTheory.Functor.cocones + +end Functor + +section + +variable (J C) + +/-- Functorially associated to each functor `J ⥤ C`, we have the `C`-presheaf consisting of +cones with a given cone point. +-/ +@[simps] +def cones : (J ⥤ C) ⥤ Cᵒᵖ ⥤ Type max u₁ v₃ + where + obj := Functor.cones + map F G f := whiskerLeft (const J).op (yoneda.map f) +#align category_theory.cones CategoryTheory.cones + +/-- Contravariantly associated to each functor `J ⥤ C`, we have the `C`-copresheaf consisting of +cocones with a given cocone point. +-/ +@[simps] +def cocones : (J ⥤ C)ᵒᵖ ⥤ C ⥤ Type max u₁ v₃ + where + obj F := Functor.cocones (unop F) + map F G f := whiskerLeft (const J) (coyoneda.map f) +#align category_theory.cocones CategoryTheory.cocones + +end + +namespace Limits + +section + +attribute [local tidy] tactic.discrete_cases + +/-- A `c : cone F` is: +* an object `c.X` and +* a natural transformation `c.π : c.X ⟶ F` from the constant `c.X` functor to `F`. + +`cone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cones.obj X`. +-/ +structure Cone (F : J ⥤ C) where + x : C + π : (const J).obj X ⟶ F +#align category_theory.limits.cone CategoryTheory.Limits.Cone + +instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := + ⟨{ x := F.obj ⟨⟨⟩⟩ + π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ } }⟩ +#align category_theory.limits.inhabited_cone CategoryTheory.Limits.inhabitedCone + +@[simp, reassoc.1] +theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : + c.π.app j ≫ F.map f = c.π.app j' := + by + rw [← c.π.naturality f] + apply id_comp +#align category_theory.limits.cone.w CategoryTheory.Limits.Cone.w + +/-- A `c : cocone F` is +* an object `c.X` and +* a natural transformation `c.ι : F ⟶ c.X` from `F` to the constant `c.X` functor. + +`cocone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cocones.obj X`. +-/ +structure Cocone (F : J ⥤ C) where + x : C + ι : F ⟶ (const J).obj X +#align category_theory.limits.cocone CategoryTheory.Limits.Cocone + +instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := + ⟨{ x := F.obj ⟨⟨⟩⟩ + ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ } }⟩ +#align category_theory.limits.inhabited_cocone CategoryTheory.Limits.inhabitedCocone + +@[simp, reassoc.1] +theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : + F.map f ≫ c.ι.app j' = c.ι.app j := + by + rw [c.ι.naturality f] + apply comp_id +#align category_theory.limits.cocone.w CategoryTheory.Limits.Cocone.w + +end + +variable {F : J ⥤ C} + +namespace Cone + +/-- The isomorphism between a cone on `F` and an element of the functor `F.cones`. -/ +@[simps] +def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X + where + Hom c := ⟨op c.x, c.π⟩ + inv c := + { x := c.1.unop + π := c.2 } + hom_inv_id' := by + ext1 + cases x + rfl + inv_hom_id' := by + ext1 + cases x + rfl +#align category_theory.limits.cone.equiv CategoryTheory.Limits.Cone.equiv + +/-- A map to the vertex of a cone naturally induces a cone by composition. -/ +@[simps] +def extensions (c : Cone F) : yoneda.obj c.x ⋙ uliftFunctor.{u₁} ⟶ F.cones + where app X f := (const J).map f.down ≫ c.π +#align category_theory.limits.cone.extensions CategoryTheory.Limits.Cone.extensions + +/-- A map to the vertex of a cone induces a cone by composition. -/ +@[simps] +def extend (c : Cone F) {X : C} (f : X ⟶ c.x) : Cone F := + { x + π := c.extensions.app (op X) ⟨f⟩ } +#align category_theory.limits.cone.extend CategoryTheory.Limits.Cone.extend + +/-- Whisker a cone by precomposition of a functor. -/ +@[simps] +def whisker (E : K ⥤ J) (c : Cone F) : Cone (E ⋙ F) + where + x := c.x + π := whiskerLeft E c.π +#align category_theory.limits.cone.whisker CategoryTheory.Limits.Cone.whisker + +end Cone + +namespace Cocone + +/-- The isomorphism between a cocone on `F` and an element of the functor `F.cocones`. -/ +def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X + where + Hom c := ⟨c.x, c.ι⟩ + inv c := + { x := c.1 + ι := c.2 } + hom_inv_id' := by + ext1 + cases x + rfl + inv_hom_id' := by + ext1 + cases x + rfl +#align category_theory.limits.cocone.equiv CategoryTheory.Limits.Cocone.equiv + +/-- A map from the vertex of a cocone naturally induces a cocone by composition. -/ +@[simps] +def extensions (c : Cocone F) : coyoneda.obj (op c.x) ⋙ uliftFunctor.{u₁} ⟶ F.cocones + where app X f := c.ι ≫ (const J).map f.down +#align category_theory.limits.cocone.extensions CategoryTheory.Limits.Cocone.extensions + +/-- A map from the vertex of a cocone induces a cocone by composition. -/ +@[simps] +def extend (c : Cocone F) {X : C} (f : c.x ⟶ X) : Cocone F := + { x + ι := c.extensions.app X ⟨f⟩ } +#align category_theory.limits.cocone.extend CategoryTheory.Limits.Cocone.extend + +/-- Whisker a cocone by precomposition of a functor. See `whiskering` for a functorial +version. +-/ +@[simps] +def whisker (E : K ⥤ J) (c : Cocone F) : Cocone (E ⋙ F) + where + x := c.x + ι := whiskerLeft E c.ι +#align category_theory.limits.cocone.whisker CategoryTheory.Limits.Cocone.whisker + +end Cocone + +/-- A cone morphism between two cones for the same diagram is a morphism of the cone points which +commutes with the cone legs. -/ +@[ext] +structure ConeMorphism (A B : Cone F) where + Hom : A.x ⟶ B.x + w' : ∀ j : J, hom ≫ B.π.app j = A.π.app j := by obviously +#align category_theory.limits.cone_morphism CategoryTheory.Limits.ConeMorphism + +restate_axiom cone_morphism.w' + +attribute [simp, reassoc.1] cone_morphism.w + +instance inhabitedConeMorphism (A : Cone F) : Inhabited (ConeMorphism A A) := + ⟨{ Hom := 𝟙 _ }⟩ +#align category_theory.limits.inhabited_cone_morphism CategoryTheory.Limits.inhabitedConeMorphism + +/-- The category of cones on a given diagram. -/ +@[simps] +instance Cone.category : Category (Cone F) + where + Hom A B := ConeMorphism A B + comp X Y Z f g := { Hom := f.Hom ≫ g.Hom } + id B := { Hom := 𝟙 B.x } +#align category_theory.limits.cone.category CategoryTheory.Limits.Cone.category + +namespace Cones + +/-- To give an isomorphism between cones, it suffices to give an + isomorphism between their vertices which commutes with the cone + maps. -/ +@[ext, simps] +def ext {c c' : Cone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.π.app j = φ.Hom ≫ c'.π.app j) : c ≅ c' + where + Hom := { Hom := φ.Hom } + inv := + { Hom := φ.inv + w' := fun j => φ.inv_comp_eq.mpr (w j) } +#align category_theory.limits.cones.ext CategoryTheory.Limits.Cones.ext + +/-- Eta rule for cones. -/ +@[simps] +def eta (c : Cone F) : c ≅ ⟨c.x, c.π⟩ := + Cones.ext (Iso.refl _) (by tidy) +#align category_theory.limits.cones.eta CategoryTheory.Limits.Cones.eta + +/-- Given a cone morphism whose object part is an isomorphism, produce an +isomorphism of cones. +-/ +theorem cone_iso_of_hom_iso {K : J ⥤ C} {c d : Cone K} (f : c ⟶ d) [i : IsIso f.Hom] : IsIso f := + ⟨⟨{ Hom := inv f.Hom + w' := fun j => (asIso f.Hom).inv_comp_eq.2 (f.w j).symm }, by tidy⟩⟩ +#align category_theory.limits.cones.cone_iso_of_hom_iso CategoryTheory.Limits.Cones.cone_iso_of_hom_iso + +/-- +Functorially postcompose a cone for `F` by a natural transformation `F ⟶ G` to give a cone for `G`. +-/ +@[simps] +def postcompose {G : J ⥤ C} (α : F ⟶ G) : Cone F ⥤ Cone G + where + obj c := + { x := c.x + π := c.π ≫ α } + map c₁ c₂ f := { Hom := f.Hom } +#align category_theory.limits.cones.postcompose CategoryTheory.Limits.Cones.postcompose + +/-- Postcomposing a cone by the composite natural transformation `α ≫ β` is the same as +postcomposing by `α` and then by `β`. -/ +@[simps] +def postcomposeComp {G H : J ⥤ C} (α : F ⟶ G) (β : G ⟶ H) : + postcompose (α ≫ β) ≅ postcompose α ⋙ postcompose β := + NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cones.postcompose_comp CategoryTheory.Limits.Cones.postcomposeComp + +/-- Postcomposing by the identity does not change the cone up to isomorphism. -/ +@[simps] +def postcomposeId : postcompose (𝟙 F) ≅ 𝟭 (Cone F) := + NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cones.postcompose_id CategoryTheory.Limits.Cones.postcomposeId + +/-- If `F` and `G` are naturally isomorphic functors, then they have equivalent categories of +cones. +-/ +@[simps] +def postcomposeEquivalence {G : J ⥤ C} (α : F ≅ G) : Cone F ≌ Cone G + where + Functor := postcompose α.Hom + inverse := postcompose α.inv + unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + counitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cones.postcompose_equivalence CategoryTheory.Limits.Cones.postcomposeEquivalence + +/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cone F` to `cone (E ⋙ F)`. +-/ +@[simps] +def whiskering (E : K ⥤ J) : Cone F ⥤ Cone (E ⋙ F) + where + obj c := c.whisker E + map c c' f := { Hom := f.Hom } +#align category_theory.limits.cones.whiskering CategoryTheory.Limits.Cones.whiskering + +/-- Whiskering by an equivalence gives an equivalence between categories of cones. +-/ +@[simps] +def whiskeringEquivalence (e : K ≌ J) : Cone F ≌ Cone (e.Functor ⋙ F) + where + Functor := whiskering e.Functor + inverse := whiskering e.inverse ⋙ postcompose (e.invFunIdAssoc F).Hom + unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + counitIso := + NatIso.ofComponents + (fun s => + Cones.ext (Iso.refl _) + (by + intro k + dsimp + -- See library note [dsimp, simp] + simpa [e.counit_app_functor] using s.w (e.unit_inv.app k))) + (by tidy) +#align category_theory.limits.cones.whiskering_equivalence CategoryTheory.Limits.Cones.whiskeringEquivalence + +/-- The categories of cones over `F` and `G` are equivalent if `F` and `G` are naturally isomorphic +(possibly after changing the indexing category by an equivalence). +-/ +@[simps Functor inverse unitIso counitIso] +def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.Functor ⋙ F ≅ G) : Cone F ≌ Cone G := + (whiskeringEquivalence e).trans (postcomposeEquivalence α) +#align category_theory.limits.cones.equivalence_of_reindexing CategoryTheory.Limits.Cones.equivalenceOfReindexing + +section + +variable (F) + +/-- Forget the cone structure and obtain just the cone point. -/ +@[simps] +def forget : Cone F ⥤ C where + obj t := t.x + map s t f := f.Hom +#align category_theory.limits.cones.forget CategoryTheory.Limits.Cones.forget + +variable (G : C ⥤ D) + +/-- A functor `G : C ⥤ D` sends cones over `F` to cones over `F ⋙ G` functorially. -/ +@[simps] +def functoriality : Cone F ⥤ Cone (F ⋙ G) + where + obj A := + { x := G.obj A.x + π := + { app := fun j => G.map (A.π.app j) + naturality' := by intros <;> erw [← G.map_comp] <;> tidy } } + map X Y f := + { Hom := G.map f.Hom + w' := fun j => by simp [-cone_morphism.w, ← f.w j] } +#align category_theory.limits.cones.functoriality CategoryTheory.Limits.Cones.functoriality + +instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) + where preimage X Y t := + { Hom := G.preimage t.Hom + w' := fun j => G.map_injective (by simpa using t.w j) } +#align category_theory.limits.cones.functoriality_full CategoryTheory.Limits.Cones.functorialityFull + +instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G) + where map_injective' X Y f g e := by + ext1 + injection e + apply G.map_injective h_1 +#align category_theory.limits.cones.functoriality_faithful CategoryTheory.Limits.Cones.functoriality_faithful + +/-- If `e : C ≌ D` is an equivalence of categories, then `functoriality F e.functor` induces an +equivalence between cones over `F` and cones over `F ⋙ e.functor`. +-/ +@[simps] +def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.Functor) := + let f : (F ⋙ e.Functor) ⋙ e.inverse ≅ F := + Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ + { Functor := functoriality F e.Functor + inverse := functoriality (F ⋙ e.Functor) e.inverse ⋙ (postcomposeEquivalence f).Functor + unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by tidy)) (by tidy) + counitIso := NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by tidy)) (by tidy) } +#align category_theory.limits.cones.functoriality_equivalence CategoryTheory.Limits.Cones.functorialityEquivalence + +/-- If `F` reflects isomorphisms, then `cones.functoriality F` reflects isomorphisms +as well. +-/ +instance reflects_cone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : + ReflectsIsomorphisms (Cones.functoriality K F) := + by + constructor + intros + haveI : is_iso (F.map f.hom) := (cones.forget (K ⋙ F)).map_isIso ((cones.functoriality K F).map f) + haveI := reflects_isomorphisms.reflects F f.hom + apply cone_iso_of_hom_iso +#align category_theory.limits.cones.reflects_cone_isomorphism CategoryTheory.Limits.Cones.reflects_cone_isomorphism + +end + +end Cones + +/-- A cocone morphism between two cocones for the same diagram is a morphism of the cocone points +which commutes with the cocone legs. -/ +@[ext] +structure CoconeMorphism (A B : Cocone F) where + Hom : A.x ⟶ B.x + w' : ∀ j : J, A.ι.app j ≫ hom = B.ι.app j := by obviously +#align category_theory.limits.cocone_morphism CategoryTheory.Limits.CoconeMorphism + +instance inhabitedCoconeMorphism (A : Cocone F) : Inhabited (CoconeMorphism A A) := + ⟨{ Hom := 𝟙 _ }⟩ +#align category_theory.limits.inhabited_cocone_morphism CategoryTheory.Limits.inhabitedCoconeMorphism + +restate_axiom cocone_morphism.w' + +attribute [simp, reassoc.1] cocone_morphism.w + +@[simps] +instance Cocone.category : Category (Cocone F) + where + Hom A B := CoconeMorphism A B + comp _ _ _ f g := { Hom := f.Hom ≫ g.Hom } + id B := { Hom := 𝟙 B.x } +#align category_theory.limits.cocone.category CategoryTheory.Limits.Cocone.category + +namespace Cocones + +/-- To give an isomorphism between cocones, it suffices to give an + isomorphism between their vertices which commutes with the cocone + maps. -/ +@[ext, simps] +def ext {c c' : Cocone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.ι.app j ≫ φ.Hom = c'.ι.app j) : c ≅ c' + where + Hom := { Hom := φ.Hom } + inv := + { Hom := φ.inv + w' := fun j => φ.comp_inv_eq.mpr (w j).symm } +#align category_theory.limits.cocones.ext CategoryTheory.Limits.Cocones.ext + +/-- Eta rule for cocones. -/ +@[simps] +def eta (c : Cocone F) : c ≅ ⟨c.x, c.ι⟩ := + Cocones.ext (Iso.refl _) (by tidy) +#align category_theory.limits.cocones.eta CategoryTheory.Limits.Cocones.eta + +/-- Given a cocone morphism whose object part is an isomorphism, produce an +isomorphism of cocones. +-/ +theorem cocone_iso_of_hom_iso {K : J ⥤ C} {c d : Cocone K} (f : c ⟶ d) [i : IsIso f.Hom] : + IsIso f := + ⟨⟨{ Hom := inv f.Hom + w' := fun j => (asIso f.Hom).comp_inv_eq.2 (f.w j).symm }, by tidy⟩⟩ +#align category_theory.limits.cocones.cocone_iso_of_hom_iso CategoryTheory.Limits.Cocones.cocone_iso_of_hom_iso + +/-- Functorially precompose a cocone for `F` by a natural transformation `G ⟶ F` to give a cocone +for `G`. -/ +@[simps] +def precompose {G : J ⥤ C} (α : G ⟶ F) : Cocone F ⥤ Cocone G + where + obj c := + { x := c.x + ι := α ≫ c.ι } + map c₁ c₂ f := { Hom := f.Hom } +#align category_theory.limits.cocones.precompose CategoryTheory.Limits.Cocones.precompose + +/-- Precomposing a cocone by the composite natural transformation `α ≫ β` is the same as +precomposing by `β` and then by `α`. -/ +def precomposeComp {G H : J ⥤ C} (α : F ⟶ G) (β : G ⟶ H) : + precompose (α ≫ β) ≅ precompose β ⋙ precompose α := + NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cocones.precompose_comp CategoryTheory.Limits.Cocones.precomposeComp + +/-- Precomposing by the identity does not change the cocone up to isomorphism. -/ +def precomposeId : precompose (𝟙 F) ≅ 𝟭 (Cocone F) := + NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cocones.precompose_id CategoryTheory.Limits.Cocones.precomposeId + +/-- If `F` and `G` are naturally isomorphic functors, then they have equivalent categories of +cocones. +-/ +@[simps] +def precomposeEquivalence {G : J ⥤ C} (α : G ≅ F) : Cocone F ≌ Cocone G + where + Functor := precompose α.Hom + inverse := precompose α.inv + unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + counitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) +#align category_theory.limits.cocones.precompose_equivalence CategoryTheory.Limits.Cocones.precomposeEquivalence + +/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cocone F` to `cocone (E ⋙ F)`. +-/ +@[simps] +def whiskering (E : K ⥤ J) : Cocone F ⥤ Cocone (E ⋙ F) + where + obj c := c.whisker E + map c c' f := { Hom := f.Hom } +#align category_theory.limits.cocones.whiskering CategoryTheory.Limits.Cocones.whiskering + +/-- Whiskering by an equivalence gives an equivalence between categories of cones. +-/ +@[simps] +def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.Functor ⋙ F) + where + Functor := whiskering e.Functor + inverse := + whiskering e.inverse ⋙ + precompose + ((Functor.leftUnitor F).inv ≫ + whiskerRight e.counitIso.inv F ≫ (Functor.associator _ _ _).inv) + unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + counitIso := + NatIso.ofComponents + (fun s => + Cocones.ext (Iso.refl _) + (by + intro k + dsimp + simpa [e.counit_inv_app_functor k] using s.w (e.unit.app k))) + (by tidy) +#align category_theory.limits.cocones.whiskering_equivalence CategoryTheory.Limits.Cocones.whiskeringEquivalence + +/-- +The categories of cocones over `F` and `G` are equivalent if `F` and `G` are naturally isomorphic +(possibly after changing the indexing category by an equivalence). +-/ +@[simps functor_obj] +def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.Functor ⋙ F ≅ G) : Cocone F ≌ Cocone G := + (whiskeringEquivalence e).trans (precomposeEquivalence α.symm) +#align category_theory.limits.cocones.equivalence_of_reindexing CategoryTheory.Limits.Cocones.equivalenceOfReindexing + +section + +variable (F) + +/-- Forget the cocone structure and obtain just the cocone point. -/ +@[simps] +def forget : Cocone F ⥤ C where + obj t := t.x + map s t f := f.Hom +#align category_theory.limits.cocones.forget CategoryTheory.Limits.Cocones.forget + +variable (G : C ⥤ D) + +/-- A functor `G : C ⥤ D` sends cocones over `F` to cocones over `F ⋙ G` functorially. -/ +@[simps] +def functoriality : Cocone F ⥤ Cocone (F ⋙ G) + where + obj A := + { x := G.obj A.x + ι := + { app := fun j => G.map (A.ι.app j) + naturality' := by intros <;> erw [← G.map_comp] <;> tidy } } + map _ _ f := + { Hom := G.map f.Hom + w' := by intros <;> rw [← functor.map_comp, cocone_morphism.w] } +#align category_theory.limits.cocones.functoriality CategoryTheory.Limits.Cocones.functoriality + +instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) + where preimage X Y t := + { Hom := G.preimage t.Hom + w' := fun j => G.map_injective (by simpa using t.w j) } +#align category_theory.limits.cocones.functoriality_full CategoryTheory.Limits.Cocones.functorialityFull + +instance functoriality_faithful [Faithful G] : Faithful (functoriality F G) + where map_injective' X Y f g e := by + ext1 + injection e + apply G.map_injective h_1 +#align category_theory.limits.cocones.functoriality_faithful CategoryTheory.Limits.Cocones.functoriality_faithful + +/-- If `e : C ≌ D` is an equivalence of categories, then `functoriality F e.functor` induces an +equivalence between cocones over `F` and cocones over `F ⋙ e.functor`. +-/ +@[simps] +def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.Functor) := + let f : (F ⋙ e.Functor) ⋙ e.inverse ≅ F := + Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ + { Functor := functoriality F e.Functor + inverse := functoriality (F ⋙ e.Functor) e.inverse ⋙ (precomposeEquivalence f.symm).Functor + unitIso := NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by tidy)) (by tidy) + counitIso := + NatIso.ofComponents + (fun c => + Cocones.ext (e.counitIso.app _) + (by + -- Unfortunately this doesn't work by `tidy`. + -- In this configuration `simp` reaches a dead-end and needs help. + intro j + dsimp + simp only [← equivalence.counit_inv_app_functor, iso.inv_hom_id_app, map_comp, + equivalence.fun_inv_map, assoc, id_comp, iso.inv_hom_id_app_assoc] + dsimp; simp))-- See note [dsimp, simp]. + fun c c' f => by + ext + dsimp + simp + dsimp + simp } +#align category_theory.limits.cocones.functoriality_equivalence CategoryTheory.Limits.Cocones.functorialityEquivalence + +/-- If `F` reflects isomorphisms, then `cocones.functoriality F` reflects isomorphisms +as well. +-/ +instance reflects_cocone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : + ReflectsIsomorphisms (Cocones.functoriality K F) := + by + constructor + intros + haveI : is_iso (F.map f.hom) := + (cocones.forget (K ⋙ F)).map_isIso ((cocones.functoriality K F).map f) + haveI := reflects_isomorphisms.reflects F f.hom + apply cocone_iso_of_hom_iso +#align category_theory.limits.cocones.reflects_cocone_isomorphism CategoryTheory.Limits.Cocones.reflects_cocone_isomorphism + +end + +end Cocones + +end Limits + +namespace Functor + +variable {F : J ⥤ C} {G : J ⥤ C} (H : C ⥤ D) + +open CategoryTheory.Limits + +/-- The image of a cone in C under a functor G : C ⥤ D is a cone in D. -/ +@[simps] +def mapCone (c : Cone F) : Cone (F ⋙ H) := + (Cones.functoriality F H).obj c +#align category_theory.functor.map_cone CategoryTheory.Functor.mapCone + +/-- The image of a cocone in C under a functor G : C ⥤ D is a cocone in D. -/ +@[simps] +def mapCocone (c : Cocone F) : Cocone (F ⋙ H) := + (Cocones.functoriality F H).obj c +#align category_theory.functor.map_cocone CategoryTheory.Functor.mapCocone + +/-- Given a cone morphism `c ⟶ c'`, construct a cone morphism on the mapped cones functorially. -/ +def mapConeMorphism {c c' : Cone F} (f : c ⟶ c') : H.mapCone c ⟶ H.mapCone c' := + (Cones.functoriality F H).map f +#align category_theory.functor.map_cone_morphism CategoryTheory.Functor.mapConeMorphism + +/-- Given a cocone morphism `c ⟶ c'`, construct a cocone morphism on the mapped cocones +functorially. -/ +def mapCoconeMorphism {c c' : Cocone F} (f : c ⟶ c') : H.mapCocone c ⟶ H.mapCocone c' := + (Cocones.functoriality F H).map f +#align category_theory.functor.map_cocone_morphism CategoryTheory.Functor.mapCoconeMorphism + +/-- If `H` is an equivalence, we invert `H.map_cone` and get a cone for `F` from a cone +for `F ⋙ H`.-/ +def mapConeInv [IsEquivalence H] (c : Cone (F ⋙ H)) : Cone F := + (Limits.Cones.functorialityEquivalence F (asEquivalence H)).inverse.obj c +#align category_theory.functor.map_cone_inv CategoryTheory.Functor.mapConeInv + +/-- `map_cone` is the left inverse to `map_cone_inv`. -/ +def mapConeMapConeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone (F ⋙ H)) : + mapCone H (mapConeInv H c) ≅ c := + (Limits.Cones.functorialityEquivalence F (asEquivalence H)).counitIso.app c +#align category_theory.functor.map_cone_map_cone_inv CategoryTheory.Functor.mapConeMapConeInv + +/-- `map_cone` is the right inverse to `map_cone_inv`. -/ +def mapConeInvMapCone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone F) : + mapConeInv H (mapCone H c) ≅ c := + (Limits.Cones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c +#align category_theory.functor.map_cone_inv_map_cone CategoryTheory.Functor.mapConeInvMapCone + +/-- If `H` is an equivalence, we invert `H.map_cone` and get a cone for `F` from a cone +for `F ⋙ H`.-/ +def mapCoconeInv [IsEquivalence H] (c : Cocone (F ⋙ H)) : Cocone F := + (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).inverse.obj c +#align category_theory.functor.map_cocone_inv CategoryTheory.Functor.mapCoconeInv + +/-- `map_cocone` is the left inverse to `map_cocone_inv`. -/ +def mapCoconeMapCoconeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone (F ⋙ H)) : + mapCocone H (mapCoconeInv H c) ≅ c := + (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).counitIso.app c +#align category_theory.functor.map_cocone_map_cocone_inv CategoryTheory.Functor.mapCoconeMapCoconeInv + +/-- `map_cocone` is the right inverse to `map_cocone_inv`. -/ +def mapCoconeInvMapCocone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone F) : + mapCoconeInv H (mapCocone H c) ≅ c := + (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c +#align category_theory.functor.map_cocone_inv_map_cocone CategoryTheory.Functor.mapCoconeInvMapCocone + +/-- `functoriality F _ ⋙ postcompose (whisker_left F _)` simplifies to `functoriality F _`. -/ +@[simps] +def functorialityCompPostcompose {H H' : C ⥤ D} (α : H ≅ H') : + Cones.functoriality F H ⋙ Cones.postcompose (whiskerLeft F α.Hom) ≅ Cones.functoriality F H' := + NatIso.ofComponents (fun c => Cones.ext (α.app _) (by tidy)) (by tidy) +#align category_theory.functor.functoriality_comp_postcompose CategoryTheory.Functor.functorialityCompPostcompose + +/-- For `F : J ⥤ C`, given a cone `c : cone F`, and a natural isomorphism `α : H ≅ H'` for functors +`H H' : C ⥤ D`, the postcomposition of the cone `H.map_cone` using the isomorphism `α` is +isomorphic to the cone `H'.map_cone`. +-/ +@[simps] +def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) : + (Cones.postcompose (whiskerLeft F α.Hom : _)).obj (H.mapCone c) ≅ H'.mapCone c := + (functorialityCompPostcompose α).app c +#align category_theory.functor.postcompose_whisker_left_map_cone CategoryTheory.Functor.postcomposeWhiskerLeftMapCone + +/-- +`map_cone` commutes with `postcompose`. In particular, for `F : J ⥤ C`, given a cone `c : cone F`, a +natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious ways of producing +a cone over `G ⋙ H`, and they are both isomorphic. +-/ +@[simps] +def mapConePostcompose {α : F ⟶ G} {c} : + H.mapCone ((Cones.postcompose α).obj c) ≅ + (Cones.postcompose (whiskerRight α H : _)).obj (H.mapCone c) := + Cones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cone_postcompose CategoryTheory.Functor.mapConePostcompose + +/-- `map_cone` commutes with `postcompose_equivalence` +-/ +@[simps] +def mapConePostcomposeEquivalenceFunctor {α : F ≅ G} {c} : + H.mapCone ((Cones.postcomposeEquivalence α).Functor.obj c) ≅ + (Cones.postcomposeEquivalence (isoWhiskerRight α H : _)).Functor.obj (H.mapCone c) := + Cones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cone_postcompose_equivalence_functor CategoryTheory.Functor.mapConePostcomposeEquivalenceFunctor + +/-- `functoriality F _ ⋙ precompose (whisker_left F _)` simplifies to `functoriality F _`. -/ +@[simps] +def functorialityCompPrecompose {H H' : C ⥤ D} (α : H ≅ H') : + Cocones.functoriality F H ⋙ Cocones.precompose (whiskerLeft F α.inv) ≅ + Cocones.functoriality F H' := + NatIso.ofComponents (fun c => Cocones.ext (α.app _) (by tidy)) (by tidy) +#align category_theory.functor.functoriality_comp_precompose CategoryTheory.Functor.functorialityCompPrecompose + +/-- +For `F : J ⥤ C`, given a cocone `c : cocone F`, and a natural isomorphism `α : H ≅ H'` for functors +`H H' : C ⥤ D`, the precomposition of the cocone `H.map_cocone` using the isomorphism `α` is +isomorphic to the cocone `H'.map_cocone`. +-/ +@[simps] +def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone F) : + (Cocones.precompose (whiskerLeft F α.inv : _)).obj (H.mapCocone c) ≅ H'.mapCocone c := + (functorialityCompPrecompose α).app c +#align category_theory.functor.precompose_whisker_left_map_cocone CategoryTheory.Functor.precomposeWhiskerLeftMapCocone + +/-- `map_cocone` commutes with `precompose`. In particular, for `F : J ⥤ C`, given a cocone +`c : cocone F`, a natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious +ways of producing a cocone over `G ⋙ H`, and they are both isomorphic. +-/ +@[simps] +def mapCoconePrecompose {α : F ⟶ G} {c} : + H.mapCocone ((Cocones.precompose α).obj c) ≅ + (Cocones.precompose (whiskerRight α H : _)).obj (H.mapCocone c) := + Cocones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cocone_precompose CategoryTheory.Functor.mapCoconePrecompose + +/-- `map_cocone` commutes with `precompose_equivalence` +-/ +@[simps] +def mapCoconePrecomposeEquivalenceFunctor {α : F ≅ G} {c} : + H.mapCocone ((Cocones.precomposeEquivalence α).Functor.obj c) ≅ + (Cocones.precomposeEquivalence (isoWhiskerRight α H : _)).Functor.obj (H.mapCocone c) := + Cocones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cocone_precompose_equivalence_functor CategoryTheory.Functor.mapCoconePrecomposeEquivalenceFunctor + +/-- `map_cone` commutes with `whisker` +-/ +@[simps] +def mapConeWhisker {E : K ⥤ J} {c : Cone F} : H.mapCone (c.whisker E) ≅ (H.mapCone c).whisker E := + Cones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cone_whisker CategoryTheory.Functor.mapConeWhisker + +/-- `map_cocone` commutes with `whisker` +-/ +@[simps] +def mapCoconeWhisker {E : K ⥤ J} {c : Cocone F} : + H.mapCocone (c.whisker E) ≅ (H.mapCocone c).whisker E := + Cocones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cocone_whisker CategoryTheory.Functor.mapCoconeWhisker + +end Functor + +end CategoryTheory + +namespace CategoryTheory.Limits + +section + +variable {F : J ⥤ C} + +/-- Change a `cocone F` into a `cone F.op`. -/ +@[simps] +def Cocone.op (c : Cocone F) : Cone F.op where + x := op c.x + π := NatTrans.op c.ι +#align category_theory.limits.cocone.op CategoryTheory.Limits.Cocone.op + +/-- Change a `cone F` into a `cocone F.op`. -/ +@[simps] +def Cone.op (c : Cone F) : Cocone F.op where + x := op c.x + ι := NatTrans.op c.π +#align category_theory.limits.cone.op CategoryTheory.Limits.Cone.op + +/-- Change a `cocone F.op` into a `cone F`. -/ +@[simps] +def Cocone.unop (c : Cocone F.op) : Cone F + where + x := unop c.x + π := NatTrans.removeOp c.ι +#align category_theory.limits.cocone.unop CategoryTheory.Limits.Cocone.unop + +/-- Change a `cone F.op` into a `cocone F`. -/ +@[simps] +def Cone.unop (c : Cone F.op) : Cocone F + where + x := unop c.x + ι := NatTrans.removeOp c.π +#align category_theory.limits.cone.unop CategoryTheory.Limits.Cone.unop + +variable (F) + +/-- The category of cocones on `F` +is equivalent to the opposite category of +the category of cones on the opposite of `F`. +-/ +def coconeEquivalenceOpConeOp : Cocone F ≌ (Cone F.op)ᵒᵖ + where + Functor := + { obj := fun c => op (Cocone.op c) + map := fun X Y f => + Quiver.Hom.op + { Hom := f.Hom.op + w' := fun j => by + apply Quiver.Hom.unop_inj + dsimp + apply cocone_morphism.w } } + inverse := + { obj := fun c => Cone.unop (unop c) + map := fun X Y f => + { Hom := f.unop.Hom.unop + w' := fun j => by + apply Quiver.Hom.op_inj + dsimp + apply cone_morphism.w } } + unitIso := + NatIso.ofComponents + (fun c => + Cocones.ext (Iso.refl _) + (by + dsimp + simp)) + fun X Y f => by + ext + simp + counitIso := + NatIso.ofComponents + (fun c => by + induction c using Opposite.rec + dsimp + apply iso.op + exact + cones.ext (iso.refl _) + (by + dsimp + simp)) + fun X Y f => + Quiver.Hom.unop_inj + (ConeMorphism.ext _ _ + (by + dsimp + simp)) + functor_unitIso_comp' c := by + apply Quiver.Hom.unop_inj + ext + dsimp + apply comp_id +#align category_theory.limits.cocone_equivalence_op_cone_op CategoryTheory.Limits.coconeEquivalenceOpConeOp + +attribute [simps] cocone_equivalence_op_cone_op + +end + +section + +variable {F : J ⥤ Cᵒᵖ} + +-- Here and below we only automatically generate the `@[simp]` lemma for the `X` field, +-- as we can write a simpler `rfl` lemma for the components of the natural transformation by hand. +/-- Change a cocone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ +@[simps (config := + { rhsMd := semireducible + simpRhs := true })] +def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F + where + x := op c.x + π := NatTrans.removeLeftOp c.ι +#align category_theory.limits.cone_of_cocone_left_op CategoryTheory.Limits.coneOfCoconeLeftOp + +/-- Change a cone on `F : J ⥤ Cᵒᵖ` to a cocone on `F.left_op : Jᵒᵖ ⥤ C`. -/ +@[simps (config := + { rhsMd := semireducible + simpRhs := true })] +def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp + where + x := unop c.x + ι := NatTrans.leftOp c.π +#align category_theory.limits.cocone_left_op_of_cone CategoryTheory.Limits.coconeLeftOpOfCone + +/- When trying use `@[simps]` to generate the `ι_app` field of this definition, `@[simps]` tries to + reduce the RHS using `expr.dsimp` and `expr.simp`, but for some reason the expression is not + being simplified properly. -/ +/-- Change a cone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ +@[simps x] +def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F + where + x := op c.x + ι := NatTrans.removeLeftOp c.π +#align category_theory.limits.cocone_of_cone_left_op CategoryTheory.Limits.coconeOfConeLeftOp + +@[simp] +theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : + (coconeOfConeLeftOp c).ι.app j = (c.π.app (op j)).op := + by + dsimp only [cocone_of_cone_left_op] + simp +#align category_theory.limits.cocone_of_cone_left_op_ι_app CategoryTheory.Limits.coconeOfConeLeftOp_ι_app + +/-- Change a cocone on `F : J ⥤ Cᵒᵖ` to a cone on `F.left_op : Jᵒᵖ ⥤ C`. -/ +@[simps (config := + { rhsMd := semireducible + simpRhs := true })] +def coneLeftOpOfCocone (c : Cocone F) : Cone F.leftOp + where + x := unop c.x + π := NatTrans.leftOp c.ι +#align category_theory.limits.cone_left_op_of_cocone CategoryTheory.Limits.coneLeftOpOfCocone + +end + +section + +variable {F : Jᵒᵖ ⥤ C} + +/-- Change a cocone on `F.right_op : J ⥤ Cᵒᵖ` to a cone on `F : Jᵒᵖ ⥤ C`. -/ +@[simps] +def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F + where + x := unop c.x + π := NatTrans.removeRightOp c.ι +#align category_theory.limits.cone_of_cocone_right_op CategoryTheory.Limits.coneOfCoconeRightOp + +/-- Change a cone on `F : Jᵒᵖ ⥤ C` to a cocone on `F.right_op : Jᵒᵖ ⥤ C`. -/ +@[simps] +def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp + where + x := op c.x + ι := NatTrans.rightOp c.π +#align category_theory.limits.cocone_right_op_of_cone CategoryTheory.Limits.coconeRightOpOfCone + +/-- Change a cone on `F.right_op : J ⥤ Cᵒᵖ` to a cocone on `F : Jᵒᵖ ⥤ C`. -/ +@[simps] +def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F + where + x := unop c.x + ι := NatTrans.removeRightOp c.π +#align category_theory.limits.cocone_of_cone_right_op CategoryTheory.Limits.coconeOfConeRightOp + +/-- Change a cocone on `F : Jᵒᵖ ⥤ C` to a cone on `F.right_op : J ⥤ Cᵒᵖ`. -/ +@[simps] +def coneRightOpOfCocone (c : Cocone F) : Cone F.rightOp + where + x := op c.x + π := NatTrans.rightOp c.ι +#align category_theory.limits.cone_right_op_of_cocone CategoryTheory.Limits.coneRightOpOfCocone + +end + +section + +variable {F : Jᵒᵖ ⥤ Cᵒᵖ} + +/-- Change a cocone on `F.unop : J ⥤ C` into a cone on `F : Jᵒᵖ ⥤ Cᵒᵖ`. -/ +@[simps] +def coneOfCoconeUnop (c : Cocone F.unop) : Cone F + where + x := op c.x + π := NatTrans.removeUnop c.ι +#align category_theory.limits.cone_of_cocone_unop CategoryTheory.Limits.coneOfCoconeUnop + +/-- Change a cone on `F : Jᵒᵖ ⥤ Cᵒᵖ` into a cocone on `F.unop : J ⥤ C`. -/ +@[simps] +def coconeUnopOfCone (c : Cone F) : Cocone F.unop + where + x := unop c.x + ι := NatTrans.unop c.π +#align category_theory.limits.cocone_unop_of_cone CategoryTheory.Limits.coconeUnopOfCone + +/-- Change a cone on `F.unop : J ⥤ C` into a cocone on `F : Jᵒᵖ ⥤ Cᵒᵖ`. -/ +@[simps] +def coconeOfConeUnop (c : Cone F.unop) : Cocone F + where + x := op c.x + ι := NatTrans.removeUnop c.π +#align category_theory.limits.cocone_of_cone_unop CategoryTheory.Limits.coconeOfConeUnop + +/-- Change a cocone on `F : Jᵒᵖ ⥤ Cᵒᵖ` into a cone on `F.unop : J ⥤ C`. -/ +@[simps] +def coneUnopOfCocone (c : Cocone F) : Cone F.unop + where + x := unop c.x + π := NatTrans.unop c.ι +#align category_theory.limits.cone_unop_of_cocone CategoryTheory.Limits.coneUnopOfCocone + +end + +end CategoryTheory.Limits + +namespace CategoryTheory.Functor + +open CategoryTheory.Limits + +variable {F : J ⥤ C} + +section + +variable (G : C ⥤ D) + +/-- The opposite cocone of the image of a cone is the image of the opposite cocone. -/ +@[simps (config := { rhsMd := semireducible })] +def mapConeOp (t : Cone F) : (G.mapCone t).op ≅ G.op.mapCocone t.op := + Cocones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cone_op CategoryTheory.Functor.mapConeOp + +/-- The opposite cone of the image of a cocone is the image of the opposite cone. -/ +@[simps (config := { rhsMd := semireducible })] +def mapCoconeOp {t : Cocone F} : (G.mapCocone t).op ≅ G.op.mapCone t.op := + Cones.ext (Iso.refl _) (by tidy) +#align category_theory.functor.map_cocone_op CategoryTheory.Functor.mapCoconeOp + +end + +end CategoryTheory.Functor + From 1c31831334ebc9719047d6959ed54f8f27088d44 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 22:59:12 -0500 Subject: [PATCH 116/126] 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/Cones.lean | 23 +++++++++-------------- 2 files changed, 10 insertions(+), 14 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 1b810c4cb1e16..29582c2916584 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.Cones import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Opposites diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index c5c1bb53927c7..30b58f7208818 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -8,10 +8,10 @@ Authors: 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.Const -import Mathbin.CategoryTheory.DiscreteCategory -import Mathbin.CategoryTheory.Yoneda -import Mathbin.CategoryTheory.Functor.ReflectsIsomorphisms +import Mathlib.CategoryTheory.Functor.Const +import Mathlib.CategoryTheory.DiscreteCategory +import Mathlib.CategoryTheory.Yoneda +import Mathlib.CategoryTheory.Functor.ReflectsIsomorphisms /-! # Cones and cocones @@ -136,8 +136,7 @@ instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := @[simp, reassoc.1] theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : - c.π.app j ≫ F.map f = c.π.app j' := - by + c.π.app j ≫ F.map f = c.π.app j' := by rw [← c.π.naturality f] apply id_comp #align category_theory.limits.cone.w CategoryTheory.Limits.Cone.w @@ -160,8 +159,7 @@ instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := @[simp, reassoc.1] theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : - F.map f ≫ c.ι.app j' = c.ι.app j := - by + F.map f ≫ c.ι.app j' = c.ι.app j := by rw [c.ι.naturality f] apply comp_id #align category_theory.limits.cocone.w CategoryTheory.Limits.Cocone.w @@ -442,8 +440,7 @@ def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.Functor) : as well. -/ instance reflects_cone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : - ReflectsIsomorphisms (Cones.functoriality K F) := - by + ReflectsIsomorphisms (Cones.functoriality K F) := by constructor intros haveI : is_iso (F.map f.hom) := (cones.forget (K ⋙ F)).map_isIso ((cones.functoriality K F).map f) @@ -658,8 +655,7 @@ def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.Functo as well. -/ instance reflects_cocone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : - ReflectsIsomorphisms (Cocones.functoriality K F) := - by + ReflectsIsomorphisms (Cocones.functoriality K F) := by constructor intros haveI : is_iso (F.map f.hom) := @@ -973,8 +969,7 @@ def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F @[simp] theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : - (coconeOfConeLeftOp c).ι.app j = (c.π.app (op j)).op := - by + (coconeOfConeLeftOp c).ι.app j = (c.π.app (op j)).op := by dsimp only [cocone_of_cone_left_op] simp #align category_theory.limits.cocone_of_cone_left_op_ι_app CategoryTheory.Limits.coconeOfConeLeftOp_ι_app From 86f2d01767ca9a03c5513cde111ef02d6914d971 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Thu, 16 Feb 2023 23:47:10 -0500 Subject: [PATCH 117/126] fix some errors --- Mathlib/CategoryTheory/Limits/Cones.lean | 254 +++++++++++------------ 1 file changed, 125 insertions(+), 129 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 30b58f7208818..c8effc7e3c4f3 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -11,7 +11,7 @@ Authors: Stephen Morgan, Scott Morrison, Floris van Doorn import Mathlib.CategoryTheory.Functor.Const import Mathlib.CategoryTheory.DiscreteCategory import Mathlib.CategoryTheory.Yoneda -import Mathlib.CategoryTheory.Functor.ReflectsIsomorphisms +import Mathlib.CategoryTheory.Functor.ReflectsIso /-! # Cones and cocones @@ -64,13 +64,13 @@ namespace CategoryTheory namespace Functor -variable {J C} (F : J ⥤ C) +variable (F : J ⥤ C) /-- `F.cones` is the functor assigning to an object `X` the type of natural transformations from the constant functor with value `X` to `F`. An object representing this functor is a limit of `F`. -/ -@[simps] +@[simps!] def cones : Cᵒᵖ ⥤ Type max u₁ v₃ := (const J).op ⋙ yoneda.obj F #align category_theory.functor.cones CategoryTheory.Functor.cones @@ -79,7 +79,7 @@ def cones : Cᵒᵖ ⥤ Type max u₁ v₃ := natural transformations from `F` to the constant functor with value `X`. An object corepresenting this functor is a colimit of `F`. -/ -@[simps] +@[simps!] def cocones : C ⥤ Type max u₁ v₃ := const J ⋙ coyoneda.obj (op F) #align category_theory.functor.cocones CategoryTheory.Functor.cocones @@ -93,21 +93,19 @@ variable (J C) /-- Functorially associated to each functor `J ⥤ C`, we have the `C`-presheaf consisting of cones with a given cone point. -/ -@[simps] -def cones : (J ⥤ C) ⥤ Cᵒᵖ ⥤ Type max u₁ v₃ - where +@[simps!] +def cones : (J ⥤ C) ⥤ Cᵒᵖ ⥤ Type max u₁ v₃ where obj := Functor.cones - map F G f := whiskerLeft (const J).op (yoneda.map f) + map f := whiskerLeft (const J).op (yoneda.map f) #align category_theory.cones CategoryTheory.cones /-- Contravariantly associated to each functor `J ⥤ C`, we have the `C`-copresheaf consisting of cocones with a given cocone point. -/ -@[simps] -def cocones : (J ⥤ C)ᵒᵖ ⥤ C ⥤ Type max u₁ v₃ - where +@[simps!] +def cocones : (J ⥤ C)ᵒᵖ ⥤ C ⥤ Type max u₁ v₃ where obj F := Functor.cocones (unop F) - map F G f := whiskerLeft (const J) (coyoneda.map f) + map f := whiskerLeft (const J) (coyoneda.map f) #align category_theory.cocones CategoryTheory.cocones end @@ -116,7 +114,7 @@ namespace Limits section -attribute [local tidy] tactic.discrete_cases +-- attribute [local tidy] tactic.discrete_cases /-- A `c : cone F` is: * an object `c.X` and @@ -131,10 +129,12 @@ structure Cone (F : J ⥤ C) where instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := ⟨{ x := F.obj ⟨⟨⟩⟩ - π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ } }⟩ + π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ + naturality := aesop_cat } + }⟩ #align category_theory.limits.inhabited_cone CategoryTheory.Limits.inhabitedCone -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : c.π.app j ≫ F.map f = c.π.app j' := by rw [← c.π.naturality f] @@ -154,10 +154,11 @@ structure Cocone (F : J ⥤ C) where instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := ⟨{ x := F.obj ⟨⟨⟩⟩ - ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ } }⟩ + ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ + naturality := aesop_cat} }⟩ #align category_theory.limits.inhabited_cocone CategoryTheory.Limits.inhabitedCocone -@[simp, reassoc.1] +@[reassoc (attr := simp)] theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : F.map f ≫ c.ι.app j' = c.ι.app j := by rw [c.ι.naturality f] @@ -171,33 +172,34 @@ variable {F : J ⥤ C} namespace Cone /-- The isomorphism between a cone on `F` and an element of the functor `F.cones`. -/ -@[simps] +@[simps!] def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X where - Hom c := ⟨op c.x, c.π⟩ + hom c := ⟨op c.x, c.π⟩ inv c := { x := c.1.unop π := c.2 } - hom_inv_id' := by + hom_inv_id := by ext1 cases x rfl - inv_hom_id' := by + inv_hom_id := by ext1 cases x rfl #align category_theory.limits.cone.equiv CategoryTheory.Limits.Cone.equiv /-- A map to the vertex of a cone naturally induces a cone by composition. -/ -@[simps] -def extensions (c : Cone F) : yoneda.obj c.x ⋙ uliftFunctor.{u₁} ⟶ F.cones - where app X f := (const J).map f.down ≫ c.π +@[simps!] +def extensions (c : Cone F) : yoneda.obj c.x ⋙ uliftFunctor.{u₁} ⟶ F.cones where + app X f := (const J).map f.down ≫ c.π + naturality := sorry #align category_theory.limits.cone.extensions CategoryTheory.Limits.Cone.extensions /-- A map to the vertex of a cone induces a cone by composition. -/ @[simps] def extend (c : Cone F) {X : C} (f : X ⟶ c.x) : Cone F := - { x + { x := X π := c.extensions.app (op X) ⟨f⟩ } #align category_theory.limits.cone.extend CategoryTheory.Limits.Cone.extend @@ -216,15 +218,15 @@ namespace Cocone /-- The isomorphism between a cocone on `F` and an element of the functor `F.cocones`. -/ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X where - Hom c := ⟨c.x, c.ι⟩ + hom c := ⟨c.x, c.ι⟩ inv c := { x := c.1 ι := c.2 } - hom_inv_id' := by + hom_inv_id := by ext1 cases x rfl - inv_hom_id' := by + inv_hom_id := by ext1 cases x rfl @@ -232,14 +234,15 @@ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X /-- A map from the vertex of a cocone naturally induces a cocone by composition. -/ @[simps] -def extensions (c : Cocone F) : coyoneda.obj (op c.x) ⋙ uliftFunctor.{u₁} ⟶ F.cocones - where app X f := c.ι ≫ (const J).map f.down +def extensions (c : Cocone F) : coyoneda.obj (op c.x) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where + app X f := c.ι ≫ (const J).map f.down + naturality := sorry #align category_theory.limits.cocone.extensions CategoryTheory.Limits.Cocone.extensions /-- A map from the vertex of a cocone induces a cocone by composition. -/ @[simps] def extend (c : Cocone F) {X : C} (f : c.x ⟶ X) : Cocone F := - { x + { c with ι := c.extensions.app X ⟨f⟩ } #align category_theory.limits.cocone.extend CategoryTheory.Limits.Cocone.extend @@ -260,12 +263,11 @@ commutes with the cone legs. -/ @[ext] structure ConeMorphism (A B : Cone F) where Hom : A.x ⟶ B.x - w' : ∀ j : J, hom ≫ B.π.app j = A.π.app j := by obviously + w : ∀ j : J, Hom ≫ B.π.app j = A.π.app j := by aesop_cat #align category_theory.limits.cone_morphism CategoryTheory.Limits.ConeMorphism +#align category_theory.limits.cone_morphism.w' CategoryTheory.Limits.ConeMorphism.w -restate_axiom cone_morphism.w' - -attribute [simp, reassoc.1] cone_morphism.w +attribute [reassoc (attr := simp)] ConeMorphism.w instance inhabitedConeMorphism (A : Cone F) : Inhabited (ConeMorphism A A) := ⟨{ Hom := 𝟙 _ }⟩ @@ -276,7 +278,7 @@ instance inhabitedConeMorphism (A : Cone F) : Inhabited (ConeMorphism A A) := instance Cone.category : Category (Cone F) where Hom A B := ConeMorphism A B - comp X Y Z f g := { Hom := f.Hom ≫ g.Hom } + comp f g := { Hom := f.Hom ≫ g.Hom } id B := { Hom := 𝟙 B.x } #align category_theory.limits.cone.category CategoryTheory.Limits.Cone.category @@ -286,18 +288,18 @@ namespace Cones isomorphism between their vertices which commutes with the cone maps. -/ @[ext, simps] -def ext {c c' : Cone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.π.app j = φ.Hom ≫ c'.π.app j) : c ≅ c' +def ext {c c' : Cone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' where - Hom := { Hom := φ.Hom } + hom := { Hom := φ.hom } inv := { Hom := φ.inv - w' := fun j => φ.inv_comp_eq.mpr (w j) } + w := fun j => φ.inv_comp_eq.mpr (w j) } #align category_theory.limits.cones.ext CategoryTheory.Limits.Cones.ext /-- Eta rule for cones. -/ -@[simps] +@[simps!] def eta (c : Cone F) : c ≅ ⟨c.x, c.π⟩ := - Cones.ext (Iso.refl _) (by tidy) + Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cones.eta CategoryTheory.Limits.Cones.eta /-- Given a cone morphism whose object part is an isomorphism, produce an @@ -305,7 +307,7 @@ isomorphism of cones. -/ theorem cone_iso_of_hom_iso {K : J ⥤ C} {c d : Cone K} (f : c ⟶ d) [i : IsIso f.Hom] : IsIso f := ⟨⟨{ Hom := inv f.Hom - w' := fun j => (asIso f.Hom).inv_comp_eq.2 (f.w j).symm }, by tidy⟩⟩ + w := fun j => (asIso f.Hom).inv_comp_eq.2 (f.w j).symm }, by aesop_cat⟩⟩ #align category_theory.limits.cones.cone_iso_of_hom_iso CategoryTheory.Limits.Cones.cone_iso_of_hom_iso /-- @@ -317,21 +319,21 @@ def postcompose {G : J ⥤ C} (α : F ⟶ G) : Cone F ⥤ Cone G obj c := { x := c.x π := c.π ≫ α } - map c₁ c₂ f := { Hom := f.Hom } + map f := { Hom := f.Hom } #align category_theory.limits.cones.postcompose CategoryTheory.Limits.Cones.postcompose /-- Postcomposing a cone by the composite natural transformation `α ≫ β` is the same as postcomposing by `α` and then by `β`. -/ -@[simps] +@[simps!] def postcomposeComp {G H : J ⥤ C} (α : F ⟶ G) (β : G ⟶ H) : postcompose (α ≫ β) ≅ postcompose α ⋙ postcompose β := - NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cones.postcompose_comp CategoryTheory.Limits.Cones.postcomposeComp /-- Postcomposing by the identity does not change the cone up to isomorphism. -/ -@[simps] +@[simps!] def postcomposeId : postcompose (𝟙 F) ≅ 𝟭 (Cone F) := - NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cones.postcompose_id CategoryTheory.Limits.Cones.postcomposeId /-- If `F` and `G` are naturally isomorphic functors, then they have equivalent categories of @@ -340,10 +342,10 @@ cones. @[simps] def postcomposeEquivalence {G : J ⥤ C} (α : F ≅ G) : Cone F ≌ Cone G where - Functor := postcompose α.Hom + functor := postcompose α.hom inverse := postcompose α.inv - unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) - counitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) + counitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cones.postcompose_equivalence CategoryTheory.Limits.Cones.postcomposeEquivalence /-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cone F` to `cone (E ⋙ F)`. @@ -352,17 +354,17 @@ def postcomposeEquivalence {G : J ⥤ C} (α : F ≅ G) : Cone F ≌ Cone G def whiskering (E : K ⥤ J) : Cone F ⥤ Cone (E ⋙ F) where obj c := c.whisker E - map c c' f := { Hom := f.Hom } + map f := { Hom := f.Hom } #align category_theory.limits.cones.whiskering CategoryTheory.Limits.Cones.whiskering /-- Whiskering by an equivalence gives an equivalence between categories of cones. -/ @[simps] -def whiskeringEquivalence (e : K ≌ J) : Cone F ≌ Cone (e.Functor ⋙ F) +def whiskeringEquivalence (e : K ≌ J) : Cone F ≌ Cone (e.functor ⋙ F) where - Functor := whiskering e.Functor - inverse := whiskering e.inverse ⋙ postcompose (e.invFunIdAssoc F).Hom - unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by tidy)) (by tidy) + functor := whiskering e.functor + inverse := whiskering e.inverse ⋙ postcompose (e.invFunIdAssoc F).hom + unitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) counitIso := NatIso.ofComponents (fun s => @@ -371,15 +373,15 @@ def whiskeringEquivalence (e : K ≌ J) : Cone F ≌ Cone (e.Functor ⋙ F) intro k dsimp -- See library note [dsimp, simp] - simpa [e.counit_app_functor] using s.w (e.unit_inv.app k))) - (by tidy) + simpa [e.counit_app_functor] using s.w (e.unitInv.app k))) + (by aesop_cat) #align category_theory.limits.cones.whiskering_equivalence CategoryTheory.Limits.Cones.whiskeringEquivalence /-- The categories of cones over `F` and `G` are equivalent if `F` and `G` are naturally isomorphic (possibly after changing the indexing category by an equivalence). -/ -@[simps Functor inverse unitIso counitIso] -def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.Functor ⋙ F ≅ G) : Cone F ≌ Cone G := +@[simps! functor inverse unitIso counitIso] +def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.functor ⋙ F ≅ G) : Cone F ≌ Cone G := (whiskeringEquivalence e).trans (postcomposeEquivalence α) #align category_theory.limits.cones.equivalence_of_reindexing CategoryTheory.Limits.Cones.equivalenceOfReindexing @@ -391,33 +393,32 @@ variable (F) @[simps] def forget : Cone F ⥤ C where obj t := t.x - map s t f := f.Hom + map f := f.Hom #align category_theory.limits.cones.forget CategoryTheory.Limits.Cones.forget variable (G : C ⥤ D) /-- A functor `G : C ⥤ D` sends cones over `F` to cones over `F ⋙ G` functorially. -/ @[simps] -def functoriality : Cone F ⥤ Cone (F ⋙ G) - where +def functoriality : Cone F ⥤ Cone (F ⋙ G) where obj A := { x := G.obj A.x π := { app := fun j => G.map (A.π.app j) - naturality' := by intros <;> erw [← G.map_comp] <;> tidy } } + naturality := by intros <;> erw [← G.map_comp] <;> aesop_cat } } map X Y f := { Hom := G.map f.Hom w' := fun j => by simp [-cone_morphism.w, ← f.w j] } #align category_theory.limits.cones.functoriality CategoryTheory.Limits.Cones.functoriality -instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) - where preimage X Y t := +instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) where + preimage t := { Hom := G.preimage t.Hom - w' := fun j => G.map_injective (by simpa using t.w j) } + w := fun j => G.map_injective (by simpa using t.w j) } #align category_theory.limits.cones.functoriality_full CategoryTheory.Limits.Cones.functorialityFull -instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G) - where map_injective' X Y f g e := by +instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G) where + map_injective f g e := by ext1 injection e apply G.map_injective h_1 @@ -427,12 +428,12 @@ instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G equivalence between cones over `F` and cones over `F ⋙ e.functor`. -/ @[simps] -def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.Functor) := - let f : (F ⋙ e.Functor) ⋙ e.inverse ≅ F := +def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.functor) := + let f : (F ⋙ e.functor) ⋙ e.inverse ≅ F := Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ - { Functor := functoriality F e.Functor - inverse := functoriality (F ⋙ e.Functor) e.inverse ⋙ (postcomposeEquivalence f).Functor - unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by tidy)) (by tidy) + { functor := functoriality F e.functor + inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (postcomposeEquivalence f).functor + unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by aesop_cat)) (by tidy) counitIso := NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by tidy)) (by tidy) } #align category_theory.limits.cones.functoriality_equivalence CategoryTheory.Limits.Cones.functorialityEquivalence @@ -442,9 +443,9 @@ as well. instance reflects_cone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : ReflectsIsomorphisms (Cones.functoriality K F) := by constructor - intros - haveI : is_iso (F.map f.hom) := (cones.forget (K ⋙ F)).map_isIso ((cones.functoriality K F).map f) - haveI := reflects_isomorphisms.reflects F f.hom + intro A B f + haveI : IsIso (F.map f.Hom) := (Cones.forget (K ⋙ F)).map_isIso ((Cones.functoriality K F).map f) + haveI := ReflectsIsomorphisms.reflects F f.Hom apply cone_iso_of_hom_iso #align category_theory.limits.cones.reflects_cone_isomorphism CategoryTheory.Limits.Cones.reflects_cone_isomorphism @@ -457,22 +458,20 @@ which commutes with the cocone legs. -/ @[ext] structure CoconeMorphism (A B : Cocone F) where Hom : A.x ⟶ B.x - w' : ∀ j : J, A.ι.app j ≫ hom = B.ι.app j := by obviously + w : ∀ j : J, A.ι.app j ≫ Hom = B.ι.app j := by aesop_cat #align category_theory.limits.cocone_morphism CategoryTheory.Limits.CoconeMorphism +#align category_theory.limits.cocone_morphism.w' CategoryTheory.Limits.CoconeMorphism.w instance inhabitedCoconeMorphism (A : Cocone F) : Inhabited (CoconeMorphism A A) := ⟨{ Hom := 𝟙 _ }⟩ #align category_theory.limits.inhabited_cocone_morphism CategoryTheory.Limits.inhabitedCoconeMorphism -restate_axiom cocone_morphism.w' - -attribute [simp, reassoc.1] cocone_morphism.w +attribute [reassoc (attr := simp)] CoconeMorphism.w @[simps] -instance Cocone.category : Category (Cocone F) - where +instance Cocone.category : Category (Cocone F) where Hom A B := CoconeMorphism A B - comp _ _ _ f g := { Hom := f.Hom ≫ g.Hom } + comp f g := { Hom := f.Hom ≫ g.Hom } id B := { Hom := 𝟙 B.x } #align category_theory.limits.cocone.category CategoryTheory.Limits.Cocone.category @@ -482,18 +481,18 @@ namespace Cocones isomorphism between their vertices which commutes with the cocone maps. -/ @[ext, simps] -def ext {c c' : Cocone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.ι.app j ≫ φ.Hom = c'.ι.app j) : c ≅ c' +def ext {c c' : Cocone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) : c ≅ c' where - Hom := { Hom := φ.Hom } + hom := { Hom := φ.hom } inv := { Hom := φ.inv - w' := fun j => φ.comp_inv_eq.mpr (w j).symm } + w := fun j => φ.comp_inv_eq.mpr (w j).symm } #align category_theory.limits.cocones.ext CategoryTheory.Limits.Cocones.ext /-- Eta rule for cocones. -/ -@[simps] +@[simps!] def eta (c : Cocone F) : c ≅ ⟨c.x, c.ι⟩ := - Cocones.ext (Iso.refl _) (by tidy) + Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cocones.eta CategoryTheory.Limits.Cocones.eta /-- Given a cocone morphism whose object part is an isomorphism, produce an @@ -502,42 +501,40 @@ isomorphism of cocones. theorem cocone_iso_of_hom_iso {K : J ⥤ C} {c d : Cocone K} (f : c ⟶ d) [i : IsIso f.Hom] : IsIso f := ⟨⟨{ Hom := inv f.Hom - w' := fun j => (asIso f.Hom).comp_inv_eq.2 (f.w j).symm }, by tidy⟩⟩ + w := fun j => (asIso f.Hom).comp_inv_eq.2 (f.w j).symm }, by aesop_cat⟩⟩ #align category_theory.limits.cocones.cocone_iso_of_hom_iso CategoryTheory.Limits.Cocones.cocone_iso_of_hom_iso /-- Functorially precompose a cocone for `F` by a natural transformation `G ⟶ F` to give a cocone for `G`. -/ @[simps] -def precompose {G : J ⥤ C} (α : G ⟶ F) : Cocone F ⥤ Cocone G - where +def precompose {G : J ⥤ C} (α : G ⟶ F) : Cocone F ⥤ Cocone G where obj c := { x := c.x ι := α ≫ c.ι } - map c₁ c₂ f := { Hom := f.Hom } + map f := { Hom := f.Hom } #align category_theory.limits.cocones.precompose CategoryTheory.Limits.Cocones.precompose /-- Precomposing a cocone by the composite natural transformation `α ≫ β` is the same as precomposing by `β` and then by `α`. -/ def precomposeComp {G H : J ⥤ C} (α : F ⟶ G) (β : G ⟶ H) : precompose (α ≫ β) ≅ precompose β ⋙ precompose α := - NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cocones.precompose_comp CategoryTheory.Limits.Cocones.precomposeComp /-- Precomposing by the identity does not change the cocone up to isomorphism. -/ def precomposeId : precompose (𝟙 F) ≅ 𝟭 (Cocone F) := - NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cocones.precompose_id CategoryTheory.Limits.Cocones.precomposeId /-- If `F` and `G` are naturally isomorphic functors, then they have equivalent categories of cocones. -/ @[simps] -def precomposeEquivalence {G : J ⥤ C} (α : G ≅ F) : Cocone F ≌ Cocone G - where - Functor := precompose α.Hom +def precomposeEquivalence {G : J ⥤ C} (α : G ≅ F) : Cocone F ≌ Cocone G where + functor := precompose α.hom inverse := precompose α.inv - unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) - counitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) + counitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cocones.precompose_equivalence CategoryTheory.Limits.Cocones.precomposeEquivalence /-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cocone F` to `cocone (E ⋙ F)`. @@ -546,21 +543,21 @@ def precomposeEquivalence {G : J ⥤ C} (α : G ≅ F) : Cocone F ≌ Cocone G def whiskering (E : K ⥤ J) : Cocone F ⥤ Cocone (E ⋙ F) where obj c := c.whisker E - map c c' f := { Hom := f.Hom } + map f := { Hom := f.Hom } #align category_theory.limits.cocones.whiskering CategoryTheory.Limits.Cocones.whiskering /-- Whiskering by an equivalence gives an equivalence between categories of cones. -/ @[simps] -def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.Functor ⋙ F) +def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.functor ⋙ F) where - Functor := whiskering e.Functor + functor := whiskering e.functor inverse := whiskering e.inverse ⋙ precompose ((Functor.leftUnitor F).inv ≫ whiskerRight e.counitIso.inv F ≫ (Functor.associator _ _ _).inv) - unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by tidy)) (by tidy) + unitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) counitIso := NatIso.ofComponents (fun s => @@ -569,7 +566,7 @@ def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.Functor ⋙ F) intro k dsimp simpa [e.counit_inv_app_functor k] using s.w (e.unit.app k))) - (by tidy) + (by aesop_cat) #align category_theory.limits.cocones.whiskering_equivalence CategoryTheory.Limits.Cocones.whiskeringEquivalence /-- @@ -577,7 +574,7 @@ The categories of cocones over `F` and `G` are equivalent if `F` and `G` are nat (possibly after changing the indexing category by an equivalence). -/ @[simps functor_obj] -def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.Functor ⋙ F ≅ G) : Cocone F ≌ Cocone G := +def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.functor ⋙ F ≅ G) : Cocone F ≌ Cocone G := (whiskeringEquivalence e).trans (precomposeEquivalence α.symm) #align category_theory.limits.cocones.equivalence_of_reindexing CategoryTheory.Limits.Cocones.equivalenceOfReindexing @@ -589,33 +586,32 @@ variable (F) @[simps] def forget : Cocone F ⥤ C where obj t := t.x - map s t f := f.Hom + map f := f.Hom #align category_theory.limits.cocones.forget CategoryTheory.Limits.Cocones.forget variable (G : C ⥤ D) /-- A functor `G : C ⥤ D` sends cocones over `F` to cocones over `F ⋙ G` functorially. -/ @[simps] -def functoriality : Cocone F ⥤ Cocone (F ⋙ G) - where +def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where obj A := { x := G.obj A.x ι := { app := fun j => G.map (A.ι.app j) - naturality' := by intros <;> erw [← G.map_comp] <;> tidy } } - map _ _ f := + naturality := by intros <;> erw [← G.map_comp] <;> aesop_cat } } + map f := { Hom := G.map f.Hom - w' := by intros <;> rw [← functor.map_comp, cocone_morphism.w] } + w := by intros <;> rw [← Functor.map_comp, CoconeMorphism.w] } #align category_theory.limits.cocones.functoriality CategoryTheory.Limits.Cocones.functoriality -instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) - where preimage X Y t := +instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) where + preimage t := { Hom := G.preimage t.Hom - w' := fun j => G.map_injective (by simpa using t.w j) } + w := fun j => G.map_injective (by simpa using t.w j) } #align category_theory.limits.cocones.functoriality_full CategoryTheory.Limits.Cocones.functorialityFull -instance functoriality_faithful [Faithful G] : Faithful (functoriality F G) - where map_injective' X Y f g e := by +instance functoriality_faithful [Faithful G] : Faithful (functoriality F G) where + map_injective f g e := by ext1 injection e apply G.map_injective h_1 @@ -625,12 +621,12 @@ instance functoriality_faithful [Faithful G] : Faithful (functoriality F G) equivalence between cocones over `F` and cocones over `F ⋙ e.functor`. -/ @[simps] -def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.Functor) := - let f : (F ⋙ e.Functor) ⋙ e.inverse ≅ F := +def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.functor) := + let f : (F ⋙ e.functor) ⋙ e.inverse ≅ F := Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ - { Functor := functoriality F e.Functor - inverse := functoriality (F ⋙ e.Functor) e.inverse ⋙ (precomposeEquivalence f.symm).Functor - unitIso := NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by tidy)) (by tidy) + { functor := functoriality F e.functor + inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (precomposeEquivalence f.symm).Functor + unitIso := NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) counitIso := NatIso.ofComponents (fun c => @@ -640,8 +636,8 @@ def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.Functo -- In this configuration `simp` reaches a dead-end and needs help. intro j dsimp - simp only [← equivalence.counit_inv_app_functor, iso.inv_hom_id_app, map_comp, - equivalence.fun_inv_map, assoc, id_comp, iso.inv_hom_id_app_assoc] + simp only [← Equivalence.counit_inv_app_functor, Iso.inv_hom_id_app, map_comp, + Equivalence.fun_inv_map, assoc, id_comp, Iso.inv_hom_id_app_assoc] dsimp; simp))-- See note [dsimp, simp]. fun c c' f => by ext @@ -657,11 +653,11 @@ as well. instance reflects_cocone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : ReflectsIsomorphisms (Cocones.functoriality K F) := by constructor - intros - haveI : is_iso (F.map f.hom) := - (cocones.forget (K ⋙ F)).map_isIso ((cocones.functoriality K F).map f) - haveI := reflects_isomorphisms.reflects F f.hom - apply cocone_iso_of_hom_iso + intro A B f + haveI : IsIso (F.map f.Hom) := + (Cocones.forget (K ⋙ F)).map_isIso ((Cocones.functoriality K F).map f) + haveI := ReflectsIsomorphisms.reflects F f.Hom + apply Cocone_iso_of_hom_iso #align category_theory.limits.cocones.reflects_cocone_isomorphism CategoryTheory.Limits.Cocones.reflects_cocone_isomorphism end @@ -677,7 +673,7 @@ variable {F : J ⥤ C} {G : J ⥤ C} (H : C ⥤ D) open CategoryTheory.Limits /-- The image of a cone in C under a functor G : C ⥤ D is a cone in D. -/ -@[simps] +@[simps!] def mapCone (c : Cone F) : Cone (F ⋙ H) := (Cones.functoriality F H).obj c #align category_theory.functor.map_cone CategoryTheory.Functor.mapCone From ea4b2759f6e3615c49edffe607d361cddf8850b6 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 17 Feb 2023 09:09:06 -0500 Subject: [PATCH 118/126] fix some more --- Mathlib/CategoryTheory/Limits/Cones.lean | 139 ++++++++++++----------- 1 file changed, 71 insertions(+), 68 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index c8effc7e3c4f3..a8a161feb8414 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -123,14 +123,19 @@ section `cone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cones.obj X`. -/ structure Cone (F : J ⥤ C) where - x : C + X : C π : (const J).obj X ⟶ F #align category_theory.limits.cone CategoryTheory.Limits.Cone instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := - ⟨{ x := F.obj ⟨⟨⟩⟩ + ⟨{ X := F.obj ⟨⟨⟩⟩ π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ - naturality := aesop_cat } + naturality := by + dsimp [const] + intro _ _ f + rw [id_comp,id_comp] + sorry + } }⟩ #align category_theory.limits.inhabited_cone CategoryTheory.Limits.inhabitedCone @@ -148,14 +153,14 @@ theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : `cocone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cocones.obj X`. -/ structure Cocone (F : J ⥤ C) where - x : C + X : C ι : F ⟶ (const J).obj X #align category_theory.limits.cocone CategoryTheory.Limits.Cocone instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := - ⟨{ x := F.obj ⟨⟨⟩⟩ + ⟨{ X := F.obj ⟨⟨⟩⟩ ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ - naturality := aesop_cat} }⟩ + naturality := by sorry } }⟩ #align category_theory.limits.inhabited_cocone CategoryTheory.Limits.inhabitedCocone @[reassoc (attr := simp)] @@ -175,31 +180,31 @@ namespace Cone @[simps!] def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X where - hom c := ⟨op c.x, c.π⟩ + hom c := ⟨op c.X, c.π⟩ inv c := - { x := c.1.unop + { X := c.1.unop π := c.2 } hom_inv_id := by - ext1 - cases x + funext X + cases X rfl inv_hom_id := by - ext1 - cases x + funext X + cases X rfl #align category_theory.limits.cone.equiv CategoryTheory.Limits.Cone.equiv /-- A map to the vertex of a cone naturally induces a cone by composition. -/ @[simps!] -def extensions (c : Cone F) : yoneda.obj c.x ⋙ uliftFunctor.{u₁} ⟶ F.cones where +def extensions (c : Cone F) : yoneda.obj c.X ⋙ uliftFunctor.{u₁} ⟶ F.cones where app X f := (const J).map f.down ≫ c.π - naturality := sorry + naturality := by sorry #align category_theory.limits.cone.extensions CategoryTheory.Limits.Cone.extensions /-- A map to the vertex of a cone induces a cone by composition. -/ @[simps] -def extend (c : Cone F) {X : C} (f : X ⟶ c.x) : Cone F := - { x := X +def extend (c : Cone F) {X : C} (f : X ⟶ c.X) : Cone F := + { X := X π := c.extensions.app (op X) ⟨f⟩ } #align category_theory.limits.cone.extend CategoryTheory.Limits.Cone.extend @@ -207,7 +212,7 @@ def extend (c : Cone F) {X : C} (f : X ⟶ c.x) : Cone F := @[simps] def whisker (E : K ⥤ J) (c : Cone F) : Cone (E ⋙ F) where - x := c.x + X := c.X π := whiskerLeft E c.π #align category_theory.limits.cone.whisker CategoryTheory.Limits.Cone.whisker @@ -218,30 +223,30 @@ namespace Cocone /-- The isomorphism between a cocone on `F` and an element of the functor `F.cocones`. -/ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X where - hom c := ⟨c.x, c.ι⟩ + hom c := ⟨c.X, c.ι⟩ inv c := - { x := c.1 + { X := c.1 ι := c.2 } hom_inv_id := by - ext1 - cases x + funext X + cases X rfl inv_hom_id := by - ext1 - cases x + funext X + cases X rfl #align category_theory.limits.cocone.equiv CategoryTheory.Limits.Cocone.equiv /-- A map from the vertex of a cocone naturally induces a cocone by composition. -/ @[simps] -def extensions (c : Cocone F) : coyoneda.obj (op c.x) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where +def extensions (c : Cocone F) : coyoneda.obj (op c.X) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where app X f := c.ι ≫ (const J).map f.down naturality := sorry #align category_theory.limits.cocone.extensions CategoryTheory.Limits.Cocone.extensions /-- A map from the vertex of a cocone induces a cocone by composition. -/ @[simps] -def extend (c : Cocone F) {X : C} (f : c.x ⟶ X) : Cocone F := +def extend (c : Cocone F) {X : C} (f : c.X ⟶ X) : Cocone F := { c with ι := c.extensions.app X ⟨f⟩ } #align category_theory.limits.cocone.extend CategoryTheory.Limits.Cocone.extend @@ -252,7 +257,7 @@ version. @[simps] def whisker (E : K ⥤ J) (c : Cocone F) : Cocone (E ⋙ F) where - x := c.x + X := c.X ι := whiskerLeft E c.ι #align category_theory.limits.cocone.whisker CategoryTheory.Limits.Cocone.whisker @@ -262,7 +267,7 @@ end Cocone commutes with the cone legs. -/ @[ext] structure ConeMorphism (A B : Cone F) where - Hom : A.x ⟶ B.x + Hom : A.X ⟶ B.X w : ∀ j : J, Hom ≫ B.π.app j = A.π.app j := by aesop_cat #align category_theory.limits.cone_morphism CategoryTheory.Limits.ConeMorphism #align category_theory.limits.cone_morphism.w' CategoryTheory.Limits.ConeMorphism.w @@ -275,11 +280,10 @@ instance inhabitedConeMorphism (A : Cone F) : Inhabited (ConeMorphism A A) := /-- The category of cones on a given diagram. -/ @[simps] -instance Cone.category : Category (Cone F) - where +instance Cone.category : Category (Cone F) where Hom A B := ConeMorphism A B comp f g := { Hom := f.Hom ≫ g.Hom } - id B := { Hom := 𝟙 B.x } + id B := { Hom := 𝟙 B.X } #align category_theory.limits.cone.category CategoryTheory.Limits.Cone.category namespace Cones @@ -288,8 +292,7 @@ namespace Cones isomorphism between their vertices which commutes with the cone maps. -/ @[ext, simps] -def ext {c c' : Cone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' - where +def ext {c c' : Cone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' where hom := { Hom := φ.hom } inv := { Hom := φ.inv @@ -298,7 +301,7 @@ def ext {c c' : Cone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.π.app j = φ.hom ≫ /-- Eta rule for cones. -/ @[simps!] -def eta (c : Cone F) : c ≅ ⟨c.x, c.π⟩ := +def eta (c : Cone F) : c ≅ ⟨c.X, c.π⟩ := Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cones.eta CategoryTheory.Limits.Cones.eta @@ -317,7 +320,7 @@ Functorially postcompose a cone for `F` by a natural transformation `F ⟶ G` to def postcompose {G : J ⥤ C} (α : F ⟶ G) : Cone F ⥤ Cone G where obj c := - { x := c.x + { X := c.X π := c.π ≫ α } map f := { Hom := f.Hom } #align category_theory.limits.cones.postcompose CategoryTheory.Limits.Cones.postcompose @@ -392,7 +395,7 @@ variable (F) /-- Forget the cone structure and obtain just the cone point. -/ @[simps] def forget : Cone F ⥤ C where - obj t := t.x + obj t := t.X map f := f.Hom #align category_theory.limits.cones.forget CategoryTheory.Limits.Cones.forget @@ -402,13 +405,13 @@ variable (G : C ⥤ D) @[simps] def functoriality : Cone F ⥤ Cone (F ⋙ G) where obj A := - { x := G.obj A.x + { X := G.obj A.X π := { app := fun j => G.map (A.π.app j) - naturality := by intros <;> erw [← G.map_comp] <;> aesop_cat } } - map X Y f := + naturality := by intros; erw [← G.map_comp]; aesop_cat } } + map f := { Hom := G.map f.Hom - w' := fun j => by simp [-cone_morphism.w, ← f.w j] } + w := fun j => by simp [-ConeMorphism.w, ← f.w j] } #align category_theory.limits.cones.functoriality CategoryTheory.Limits.Cones.functoriality instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) where @@ -418,8 +421,8 @@ instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) wher #align category_theory.limits.cones.functoriality_full CategoryTheory.Limits.Cones.functorialityFull instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G) where - map_injective f g e := by - ext1 + map_injective _ f g e := by + funext injection e apply G.map_injective h_1 #align category_theory.limits.cones.functoriality_faithful CategoryTheory.Limits.Cones.functoriality_faithful @@ -457,7 +460,7 @@ end Cones which commutes with the cocone legs. -/ @[ext] structure CoconeMorphism (A B : Cocone F) where - Hom : A.x ⟶ B.x + Hom : A.X ⟶ B.X w : ∀ j : J, A.ι.app j ≫ Hom = B.ι.app j := by aesop_cat #align category_theory.limits.cocone_morphism CategoryTheory.Limits.CoconeMorphism #align category_theory.limits.cocone_morphism.w' CategoryTheory.Limits.CoconeMorphism.w @@ -472,7 +475,7 @@ attribute [reassoc (attr := simp)] CoconeMorphism.w instance Cocone.category : Category (Cocone F) where Hom A B := CoconeMorphism A B comp f g := { Hom := f.Hom ≫ g.Hom } - id B := { Hom := 𝟙 B.x } + id B := { Hom := 𝟙 B.X } #align category_theory.limits.cocone.category CategoryTheory.Limits.Cocone.category namespace Cocones @@ -481,8 +484,8 @@ namespace Cocones isomorphism between their vertices which commutes with the cocone maps. -/ @[ext, simps] -def ext {c c' : Cocone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) : c ≅ c' - where +def ext {c c' : Cocone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) + : c ≅ c' where hom := { Hom := φ.hom } inv := { Hom := φ.inv @@ -491,7 +494,7 @@ def ext {c c' : Cocone F} (φ : c.x ≅ c'.x) (w : ∀ j, c.ι.app j ≫ φ.hom /-- Eta rule for cocones. -/ @[simps!] -def eta (c : Cocone F) : c ≅ ⟨c.x, c.ι⟩ := +def eta (c : Cocone F) : c ≅ ⟨c.X, c.ι⟩ := Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cocones.eta CategoryTheory.Limits.Cocones.eta @@ -509,7 +512,7 @@ for `G`. -/ @[simps] def precompose {G : J ⥤ C} (α : G ⟶ F) : Cocone F ⥤ Cocone G where obj c := - { x := c.x + { X := c.X ι := α ≫ c.ι } map f := { Hom := f.Hom } #align category_theory.limits.cocones.precompose CategoryTheory.Limits.Cocones.precompose @@ -573,7 +576,7 @@ def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.functor ⋙ F) The categories of cocones over `F` and `G` are equivalent if `F` and `G` are naturally isomorphic (possibly after changing the indexing category by an equivalence). -/ -@[simps functor_obj] +@[simps! functor_obj] def equivalenceOfReindexing {G : K ⥤ C} (e : K ≌ J) (α : e.functor ⋙ F ≅ G) : Cocone F ≌ Cocone G := (whiskeringEquivalence e).trans (precomposeEquivalence α.symm) #align category_theory.limits.cocones.equivalence_of_reindexing CategoryTheory.Limits.Cocones.equivalenceOfReindexing @@ -585,7 +588,7 @@ variable (F) /-- Forget the cocone structure and obtain just the cocone point. -/ @[simps] def forget : Cocone F ⥤ C where - obj t := t.x + obj t := t.X map f := f.Hom #align category_theory.limits.cocones.forget CategoryTheory.Limits.Cocones.forget @@ -595,13 +598,13 @@ variable (G : C ⥤ D) @[simps] def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where obj A := - { x := G.obj A.x + { X := G.obj A.X ι := { app := fun j => G.map (A.ι.app j) - naturality := by intros <;> erw [← G.map_comp] <;> aesop_cat } } + naturality := by intros; erw [← G.map_comp]; aesop_cat } } map f := { Hom := G.map f.Hom - w := by intros <;> rw [← Functor.map_comp, CoconeMorphism.w] } + w := by intros; rw [← Functor.map_comp, CoconeMorphism.w] } #align category_theory.limits.cocones.functoriality CategoryTheory.Limits.Cocones.functoriality instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) where @@ -836,14 +839,14 @@ variable {F : J ⥤ C} /-- Change a `cocone F` into a `cone F.op`. -/ @[simps] def Cocone.op (c : Cocone F) : Cone F.op where - x := op c.x + X := op c.X π := NatTrans.op c.ι #align category_theory.limits.cocone.op CategoryTheory.Limits.Cocone.op /-- Change a `cone F` into a `cocone F.op`. -/ @[simps] def Cone.op (c : Cone F) : Cocone F.op where - x := op c.x + X := op c.X ι := NatTrans.op c.π #align category_theory.limits.cone.op CategoryTheory.Limits.Cone.op @@ -851,7 +854,7 @@ def Cone.op (c : Cone F) : Cocone F.op where @[simps] def Cocone.unop (c : Cocone F.op) : Cone F where - x := unop c.x + X := unop c.X π := NatTrans.removeOp c.ι #align category_theory.limits.cocone.unop CategoryTheory.Limits.Cocone.unop @@ -859,7 +862,7 @@ def Cocone.unop (c : Cocone F.op) : Cone F @[simps] def Cone.unop (c : Cone F.op) : Cocone F where - x := unop c.x + X := unop c.X ι := NatTrans.removeOp c.π #align category_theory.limits.cone.unop CategoryTheory.Limits.Cone.unop @@ -938,7 +941,7 @@ variable {F : J ⥤ Cᵒᵖ} simpRhs := true })] def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F where - x := op c.x + X := op c.X π := NatTrans.removeLeftOp c.ι #align category_theory.limits.cone_of_cocone_left_op CategoryTheory.Limits.coneOfCoconeLeftOp @@ -948,7 +951,7 @@ def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F simpRhs := true })] def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp where - x := unop c.x + X := unop c.X ι := NatTrans.leftOp c.π #align category_theory.limits.cocone_left_op_of_cone CategoryTheory.Limits.coconeLeftOpOfCone @@ -959,7 +962,7 @@ def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp @[simps x] def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F where - x := op c.x + X := op c.X ι := NatTrans.removeLeftOp c.π #align category_theory.limits.cocone_of_cone_left_op CategoryTheory.Limits.coconeOfConeLeftOp @@ -976,7 +979,7 @@ theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : simpRhs := true })] def coneLeftOpOfCocone (c : Cocone F) : Cone F.leftOp where - x := unop c.x + X := unop c.X π := NatTrans.leftOp c.ι #align category_theory.limits.cone_left_op_of_cocone CategoryTheory.Limits.coneLeftOpOfCocone @@ -990,7 +993,7 @@ variable {F : Jᵒᵖ ⥤ C} @[simps] def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F where - x := unop c.x + X := unop c.X π := NatTrans.removeRightOp c.ι #align category_theory.limits.cone_of_cocone_right_op CategoryTheory.Limits.coneOfCoconeRightOp @@ -998,7 +1001,7 @@ def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F @[simps] def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp where - x := op c.x + X := op c.X ι := NatTrans.rightOp c.π #align category_theory.limits.cocone_right_op_of_cone CategoryTheory.Limits.coconeRightOpOfCone @@ -1006,7 +1009,7 @@ def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp @[simps] def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F where - x := unop c.x + X := unop c.X ι := NatTrans.removeRightOp c.π #align category_theory.limits.cocone_of_cone_right_op CategoryTheory.Limits.coconeOfConeRightOp @@ -1014,7 +1017,7 @@ def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F @[simps] def coneRightOpOfCocone (c : Cocone F) : Cone F.rightOp where - x := op c.x + X := op c.X π := NatTrans.rightOp c.ι #align category_theory.limits.cone_right_op_of_cocone CategoryTheory.Limits.coneRightOpOfCocone @@ -1028,7 +1031,7 @@ variable {F : Jᵒᵖ ⥤ Cᵒᵖ} @[simps] def coneOfCoconeUnop (c : Cocone F.unop) : Cone F where - x := op c.x + X := op c.X π := NatTrans.removeUnop c.ι #align category_theory.limits.cone_of_cocone_unop CategoryTheory.Limits.coneOfCoconeUnop @@ -1036,7 +1039,7 @@ def coneOfCoconeUnop (c : Cocone F.unop) : Cone F @[simps] def coconeUnopOfCone (c : Cone F) : Cocone F.unop where - x := unop c.x + X := unop c.X ι := NatTrans.unop c.π #align category_theory.limits.cocone_unop_of_cone CategoryTheory.Limits.coconeUnopOfCone @@ -1044,7 +1047,7 @@ def coconeUnopOfCone (c : Cone F) : Cocone F.unop @[simps] def coconeOfConeUnop (c : Cone F.unop) : Cocone F where - x := op c.x + X := op c.X ι := NatTrans.removeUnop c.π #align category_theory.limits.cocone_of_cone_unop CategoryTheory.Limits.coconeOfConeUnop @@ -1052,7 +1055,7 @@ def coconeOfConeUnop (c : Cone F.unop) : Cocone F @[simps] def coneUnopOfCocone (c : Cocone F) : Cone F.unop where - x := unop c.x + X := unop c.X π := NatTrans.unop c.ι #align category_theory.limits.cone_unop_of_cocone CategoryTheory.Limits.coneUnopOfCocone From 8d19aa905c08e582695bf3de8de828ce99a42339 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 17 Feb 2023 16:29:22 -0500 Subject: [PATCH 119/126] fix errors --- Mathlib/CategoryTheory/Limits/Cones.lean | 233 ++++++++++++----------- 1 file changed, 125 insertions(+), 108 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index a8a161feb8414..08796ca03d20a 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -131,10 +131,10 @@ instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := ⟨{ X := F.obj ⟨⟨⟩⟩ π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ naturality := by - dsimp [const] - intro _ _ f - rw [id_comp,id_comp] - sorry + intro X Y f + match X, Y, f with + | .mk A, .mk B, .up g => + aesop_cat } }⟩ #align category_theory.limits.inhabited_cone CategoryTheory.Limits.inhabitedCone @@ -160,7 +160,13 @@ structure Cocone (F : J ⥤ C) where instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := ⟨{ X := F.obj ⟨⟨⟩⟩ ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ - naturality := by sorry } }⟩ + naturality := by + intro X Y f + match X, Y, f with + | .mk A, .mk B, .up g => + aesop_cat + } + }⟩ #align category_theory.limits.inhabited_cocone CategoryTheory.Limits.inhabitedCocone @[reassoc (attr := simp)] @@ -195,10 +201,14 @@ def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X #align category_theory.limits.cone.equiv CategoryTheory.Limits.Cone.equiv /-- A map to the vertex of a cone naturally induces a cone by composition. -/ -@[simps!] +@[simps] def extensions (c : Cone F) : yoneda.obj c.X ⋙ uliftFunctor.{u₁} ⟶ F.cones where app X f := (const J).map f.down ≫ c.π - naturality := by sorry + naturality := by + intros X Y f + dsimp [yoneda,const] + funext X + aesop_cat #align category_theory.limits.cone.extensions CategoryTheory.Limits.Cone.extensions /-- A map to the vertex of a cone induces a cone by composition. -/ @@ -241,14 +251,17 @@ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X @[simps] def extensions (c : Cocone F) : coyoneda.obj (op c.X) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where app X f := c.ι ≫ (const J).map f.down - naturality := sorry + naturality := fun {X} {Y} f => by + dsimp [coyoneda,const] + funext X + aesop_cat #align category_theory.limits.cocone.extensions CategoryTheory.Limits.Cocone.extensions /-- A map from the vertex of a cocone induces a cocone by composition. -/ @[simps] -def extend (c : Cocone F) {X : C} (f : c.X ⟶ X) : Cocone F := - { c with - ι := c.extensions.app X ⟨f⟩ } +def extend (c : Cocone F) {Y : C} (f : c.X ⟶ Y) : Cocone F where + X := Y + ι := c.extensions.app Y ⟨f⟩ #align category_theory.limits.cocone.extend CategoryTheory.Limits.Cocone.extend /-- Whisker a cocone by precomposition of a functor. See `whiskering` for a functorial @@ -291,7 +304,8 @@ namespace Cones /-- To give an isomorphism between cones, it suffices to give an isomorphism between their vertices which commutes with the cone maps. -/ -@[ext, simps] +-- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +@[aesop apply safe (rule_sets [CategoryTheory]), simps] def ext {c c' : Cone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' where hom := { Hom := φ.hom } inv := @@ -420,12 +434,12 @@ instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) wher w := fun j => G.map_injective (by simpa using t.w j) } #align category_theory.limits.cones.functoriality_full CategoryTheory.Limits.Cones.functorialityFull -instance functoriality_faithful [Faithful G] : Faithful (Cones.functoriality F G) where - map_injective _ f g e := by - funext - injection e - apply G.map_injective h_1 -#align category_theory.limits.cones.functoriality_faithful CategoryTheory.Limits.Cones.functoriality_faithful +instance functorialityFaithful [Faithful G] : Faithful (Cones.functoriality F G) where + map_injective {c} {c'} f g e := by + apply ConeMorphism.ext f g + let f := ConeMorphism.mk.inj e; dsimp [functoriality] + apply G.map_injective f +#align category_theory.limits.cones.functoriality_faithful CategoryTheory.Limits.Cones.functorialityFaithful /-- If `e : C ≌ D` is an equivalence of categories, then `functoriality F e.functor` induces an equivalence between cones over `F` and cones over `F ⋙ e.functor`. @@ -436,8 +450,8 @@ def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.functor) : Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ { functor := functoriality F e.functor inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (postcomposeEquivalence f).functor - unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by aesop_cat)) (by tidy) - counitIso := NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by tidy)) (by tidy) } + unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) + counitIso := NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by aesop_cat)) (by aesop_cat) } #align category_theory.limits.cones.functoriality_equivalence CategoryTheory.Limits.Cones.functorialityEquivalence /-- If `F` reflects isomorphisms, then `cones.functoriality F` reflects isomorphisms @@ -446,8 +460,9 @@ as well. instance reflects_cone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : ReflectsIsomorphisms (Cones.functoriality K F) := by constructor - intro A B f - haveI : IsIso (F.map f.Hom) := (Cones.forget (K ⋙ F)).map_isIso ((Cones.functoriality K F).map f) + intro A B f _ + haveI : IsIso (F.map f.Hom) := + (Cones.forget (K ⋙ F)).map_isIso ((Cones.functoriality K F).map f) haveI := ReflectsIsomorphisms.reflects F f.Hom apply cone_iso_of_hom_iso #align category_theory.limits.cones.reflects_cone_isomorphism CategoryTheory.Limits.Cones.reflects_cone_isomorphism @@ -483,7 +498,8 @@ namespace Cocones /-- To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps. -/ -@[ext, simps] +-- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule +@[aesop apply safe (rule_sets [CategoryTheory]), simps] def ext {c c' : Cocone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) : c ≅ c' where hom := { Hom := φ.hom } @@ -568,7 +584,7 @@ def whiskeringEquivalence (e : K ≌ J) : Cocone F ≌ Cocone (e.functor ⋙ F) (by intro k dsimp - simpa [e.counit_inv_app_functor k] using s.w (e.unit.app k))) + simpa [e.counitInv_app_functor k] using s.w (e.unit.app k))) (by aesop_cat) #align category_theory.limits.cocones.whiskering_equivalence CategoryTheory.Limits.Cocones.whiskeringEquivalence @@ -605,6 +621,8 @@ def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where map f := { Hom := G.map f.Hom w := by intros; rw [← Functor.map_comp, CoconeMorphism.w] } + map_id := by aesop_cat + map_comp := by aesop_cat #align category_theory.limits.cocones.functoriality CategoryTheory.Limits.Cocones.functoriality instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) where @@ -614,10 +632,10 @@ instance functorialityFull [Full G] [Faithful G] : Full (functoriality F G) wher #align category_theory.limits.cocones.functoriality_full CategoryTheory.Limits.Cocones.functorialityFull instance functoriality_faithful [Faithful G] : Faithful (functoriality F G) where - map_injective f g e := by - ext1 - injection e - apply G.map_injective h_1 + map_injective {X} {Y} f g e := by + apply CoconeMorphism.ext + let h := CoconeMorphism.mk.inj e + apply G.map_injective h #align category_theory.limits.cocones.functoriality_faithful CategoryTheory.Limits.Cocones.functoriality_faithful /-- If `e : C ≌ D` is an equivalence of categories, then `functoriality F e.functor` induces an @@ -628,7 +646,7 @@ def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.functo let f : (F ⋙ e.functor) ⋙ e.inverse ≅ F := Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ { functor := functoriality F e.functor - inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (precomposeEquivalence f.symm).Functor + inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (precomposeEquivalence f.symm).functor unitIso := NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) counitIso := NatIso.ofComponents @@ -639,15 +657,13 @@ def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.functo -- In this configuration `simp` reaches a dead-end and needs help. intro j dsimp - simp only [← Equivalence.counit_inv_app_functor, Iso.inv_hom_id_app, map_comp, + simp only [← Equivalence.counitInv_app_functor, Iso.inv_hom_id_app, map_comp, Equivalence.fun_inv_map, assoc, id_comp, Iso.inv_hom_id_app_assoc] dsimp; simp))-- See note [dsimp, simp]. - fun c c' f => by - ext - dsimp + fun {c} {c'} f => by + apply CoconeMorphism.ext simp - dsimp - simp } + } #align category_theory.limits.cocones.functoriality_equivalence CategoryTheory.Limits.Cocones.functorialityEquivalence /-- If `F` reflects isomorphisms, then `cocones.functoriality F` reflects isomorphisms @@ -656,11 +672,11 @@ as well. instance reflects_cocone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : ReflectsIsomorphisms (Cocones.functoriality K F) := by constructor - intro A B f + intro A B f _ haveI : IsIso (F.map f.Hom) := (Cocones.forget (K ⋙ F)).map_isIso ((Cocones.functoriality K F).map f) haveI := ReflectsIsomorphisms.reflects F f.Hom - apply Cocone_iso_of_hom_iso + apply cocone_iso_of_hom_iso #align category_theory.limits.cocones.reflects_cocone_isomorphism CategoryTheory.Limits.Cocones.reflects_cocone_isomorphism end @@ -675,6 +691,7 @@ variable {F : J ⥤ C} {G : J ⥤ C} (H : C ⥤ D) open CategoryTheory.Limits +/- Porting note: dot notation on the functor is broken for `mapCone` -/ /-- The image of a cone in C under a functor G : C ⥤ D is a cone in D. -/ @[simps!] def mapCone (c : Cone F) : Cone (F ⋙ H) := @@ -682,19 +699,19 @@ def mapCone (c : Cone F) : Cone (F ⋙ H) := #align category_theory.functor.map_cone CategoryTheory.Functor.mapCone /-- The image of a cocone in C under a functor G : C ⥤ D is a cocone in D. -/ -@[simps] +@[simps!] def mapCocone (c : Cocone F) : Cocone (F ⋙ H) := (Cocones.functoriality F H).obj c #align category_theory.functor.map_cocone CategoryTheory.Functor.mapCocone /-- Given a cone morphism `c ⟶ c'`, construct a cone morphism on the mapped cones functorially. -/ -def mapConeMorphism {c c' : Cone F} (f : c ⟶ c') : H.mapCone c ⟶ H.mapCone c' := +def mapConeMorphism {c c' : Cone F} (f : c ⟶ c') : mapCone H c ⟶ mapCone _ c' := (Cones.functoriality F H).map f #align category_theory.functor.map_cone_morphism CategoryTheory.Functor.mapConeMorphism /-- Given a cocone morphism `c ⟶ c'`, construct a cocone morphism on the mapped cocones functorially. -/ -def mapCoconeMorphism {c c' : Cocone F} (f : c ⟶ c') : H.mapCocone c ⟶ H.mapCocone c' := +def mapCoconeMorphism {c c' : Cocone F} (f : c ⟶ c') : mapCocone H c ⟶ mapCocone _ c' := (Cocones.functoriality F H).map f #align category_theory.functor.map_cocone_morphism CategoryTheory.Functor.mapCoconeMorphism @@ -735,19 +752,19 @@ def mapCoconeInvMapCocone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Coc #align category_theory.functor.map_cocone_inv_map_cocone CategoryTheory.Functor.mapCoconeInvMapCocone /-- `functoriality F _ ⋙ postcompose (whisker_left F _)` simplifies to `functoriality F _`. -/ -@[simps] +@[simps!] def functorialityCompPostcompose {H H' : C ⥤ D} (α : H ≅ H') : - Cones.functoriality F H ⋙ Cones.postcompose (whiskerLeft F α.Hom) ≅ Cones.functoriality F H' := - NatIso.ofComponents (fun c => Cones.ext (α.app _) (by tidy)) (by tidy) + Cones.functoriality F H ⋙ Cones.postcompose (whiskerLeft F α.hom) ≅ Cones.functoriality F H' := + NatIso.ofComponents (fun c => Cones.ext (α.app _) (by aesop_cat)) (by aesop_cat) #align category_theory.functor.functoriality_comp_postcompose CategoryTheory.Functor.functorialityCompPostcompose /-- For `F : J ⥤ C`, given a cone `c : cone F`, and a natural isomorphism `α : H ≅ H'` for functors `H H' : C ⥤ D`, the postcomposition of the cone `H.map_cone` using the isomorphism `α` is isomorphic to the cone `H'.map_cone`. -/ -@[simps] +@[simps!] def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) : - (Cones.postcompose (whiskerLeft F α.Hom : _)).obj (H.mapCone c) ≅ H'.mapCone c := + (Cones.postcompose (whiskerLeft F α.hom : _)).obj (mapCone H c) ≅ mapCone H' c := (functorialityCompPostcompose α).app c #align category_theory.functor.postcompose_whisker_left_map_cone CategoryTheory.Functor.postcomposeWhiskerLeftMapCone @@ -756,28 +773,28 @@ def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious ways of producing a cone over `G ⋙ H`, and they are both isomorphic. -/ -@[simps] +@[simps!] def mapConePostcompose {α : F ⟶ G} {c} : - H.mapCone ((Cones.postcompose α).obj c) ≅ - (Cones.postcompose (whiskerRight α H : _)).obj (H.mapCone c) := - Cones.ext (Iso.refl _) (by tidy) + mapCone H ((Cones.postcompose α).obj c) ≅ + (Cones.postcompose (whiskerRight α H : _)).obj (mapCone H c) := + Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_postcompose CategoryTheory.Functor.mapConePostcompose /-- `map_cone` commutes with `postcompose_equivalence` -/ -@[simps] +@[simps!] def mapConePostcomposeEquivalenceFunctor {α : F ≅ G} {c} : - H.mapCone ((Cones.postcomposeEquivalence α).Functor.obj c) ≅ - (Cones.postcomposeEquivalence (isoWhiskerRight α H : _)).Functor.obj (H.mapCone c) := - Cones.ext (Iso.refl _) (by tidy) + mapCone H ((Cones.postcomposeEquivalence α).functor.obj c) ≅ + (Cones.postcomposeEquivalence (isoWhiskerRight α H : _)).functor.obj (mapCone H c) := + Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_postcompose_equivalence_functor CategoryTheory.Functor.mapConePostcomposeEquivalenceFunctor /-- `functoriality F _ ⋙ precompose (whisker_left F _)` simplifies to `functoriality F _`. -/ -@[simps] +@[simps!] def functorialityCompPrecompose {H H' : C ⥤ D} (α : H ≅ H') : Cocones.functoriality F H ⋙ Cocones.precompose (whiskerLeft F α.inv) ≅ Cocones.functoriality F H' := - NatIso.ofComponents (fun c => Cocones.ext (α.app _) (by tidy)) (by tidy) + NatIso.ofComponents (fun c => Cocones.ext (α.app _) (by aesop_cat)) (by aesop_cat) #align category_theory.functor.functoriality_comp_precompose CategoryTheory.Functor.functorialityCompPrecompose /-- @@ -785,9 +802,9 @@ For `F : J ⥤ C`, given a cocone `c : cocone F`, and a natural isomorphism `α `H H' : C ⥤ D`, the precomposition of the cocone `H.map_cocone` using the isomorphism `α` is isomorphic to the cocone `H'.map_cocone`. -/ -@[simps] +@[simps!] def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone F) : - (Cocones.precompose (whiskerLeft F α.inv : _)).obj (H.mapCocone c) ≅ H'.mapCocone c := + (Cocones.precompose (whiskerLeft F α.inv : _)).obj (mapCocone H c) ≅ mapCocone H' c := (functorialityCompPrecompose α).app c #align category_theory.functor.precompose_whisker_left_map_cocone CategoryTheory.Functor.precomposeWhiskerLeftMapCocone @@ -795,35 +812,35 @@ def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone `c : cocone F`, a natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious ways of producing a cocone over `G ⋙ H`, and they are both isomorphic. -/ -@[simps] +@[simps!] def mapCoconePrecompose {α : F ⟶ G} {c} : - H.mapCocone ((Cocones.precompose α).obj c) ≅ - (Cocones.precompose (whiskerRight α H : _)).obj (H.mapCocone c) := - Cocones.ext (Iso.refl _) (by tidy) + mapCocone H ((Cocones.precompose α).obj c) ≅ + (Cocones.precompose (whiskerRight α H : _)).obj (mapCocone H c) := + Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_precompose CategoryTheory.Functor.mapCoconePrecompose /-- `map_cocone` commutes with `precompose_equivalence` -/ -@[simps] +@[simps!] def mapCoconePrecomposeEquivalenceFunctor {α : F ≅ G} {c} : - H.mapCocone ((Cocones.precomposeEquivalence α).Functor.obj c) ≅ - (Cocones.precomposeEquivalence (isoWhiskerRight α H : _)).Functor.obj (H.mapCocone c) := - Cocones.ext (Iso.refl _) (by tidy) + mapCocone H ((Cocones.precomposeEquivalence α).functor.obj c) ≅ + (Cocones.precomposeEquivalence (isoWhiskerRight α H : _)).functor.obj (mapCocone H c) := + Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_precompose_equivalence_functor CategoryTheory.Functor.mapCoconePrecomposeEquivalenceFunctor /-- `map_cone` commutes with `whisker` -/ -@[simps] -def mapConeWhisker {E : K ⥤ J} {c : Cone F} : H.mapCone (c.whisker E) ≅ (H.mapCone c).whisker E := - Cones.ext (Iso.refl _) (by tidy) +@[simps!] +def mapConeWhisker {E : K ⥤ J} {c : Cone F} : mapCone H (c.whisker E) ≅ (mapCone H c).whisker E := + Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_whisker CategoryTheory.Functor.mapConeWhisker /-- `map_cocone` commutes with `whisker` -/ -@[simps] +@[simps!] def mapCoconeWhisker {E : K ⥤ J} {c : Cocone F} : - H.mapCocone (c.whisker E) ≅ (H.mapCocone c).whisker E := - Cocones.ext (Iso.refl _) (by tidy) + mapCocone H (c.whisker E) ≅ (mapCocone H c).whisker E := + Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_whisker CategoryTheory.Functor.mapCoconeWhisker end Functor @@ -839,14 +856,14 @@ variable {F : J ⥤ C} /-- Change a `cocone F` into a `cone F.op`. -/ @[simps] def Cocone.op (c : Cocone F) : Cone F.op where - X := op c.X + X := Opposite.op c.X π := NatTrans.op c.ι #align category_theory.limits.cocone.op CategoryTheory.Limits.Cocone.op /-- Change a `cone F` into a `cocone F.op`. -/ @[simps] def Cone.op (c : Cone F) : Cocone F.op where - X := op c.X + X := Opposite.op c.X ι := NatTrans.op c.π #align category_theory.limits.cone.op CategoryTheory.Limits.Cone.op @@ -854,7 +871,7 @@ def Cone.op (c : Cone F) : Cocone F.op where @[simps] def Cocone.unop (c : Cocone F.op) : Cone F where - X := unop c.X + X := Opposite.unop c.X π := NatTrans.removeOp c.ι #align category_theory.limits.cocone.unop CategoryTheory.Limits.Cocone.unop @@ -862,7 +879,7 @@ def Cocone.unop (c : Cocone F.op) : Cone F @[simps] def Cone.unop (c : Cone F.op) : Cocone F where - X := unop c.X + X := Opposite.unop c.X ι := NatTrans.removeOp c.π #align category_theory.limits.cone.unop CategoryTheory.Limits.Cone.unop @@ -874,23 +891,23 @@ the category of cones on the opposite of `F`. -/ def coconeEquivalenceOpConeOp : Cocone F ≌ (Cone F.op)ᵒᵖ where - Functor := + functor := { obj := fun c => op (Cocone.op c) - map := fun X Y f => + map := fun {X} {Y} f => Quiver.Hom.op { Hom := f.Hom.op - w' := fun j => by + w := fun j => by apply Quiver.Hom.unop_inj dsimp - apply cocone_morphism.w } } + apply CoconeMorphism.w } } inverse := { obj := fun c => Cone.unop (unop c) - map := fun X Y f => + map := fun {X} {Y} f => { Hom := f.unop.Hom.unop - w' := fun j => by + w := fun j => by apply Quiver.Hom.op_inj dsimp - apply cone_morphism.w } } + apply ConeMorphism.w } } unitIso := NatIso.ofComponents (fun c => @@ -898,47 +915,49 @@ def coconeEquivalenceOpConeOp : Cocone F ≌ (Cone F.op)ᵒᵖ (by dsimp simp)) - fun X Y f => by - ext + fun {X} {Y} f => by + apply CoconeMorphism.ext simp counitIso := NatIso.ofComponents (fun c => by induction c using Opposite.rec dsimp - apply iso.op + apply Iso.op exact - cones.ext (iso.refl _) + Cones.ext (Iso.refl _) (by dsimp simp)) - fun X Y f => + fun {X} {Y} f => Quiver.Hom.unop_inj (ConeMorphism.ext _ _ (by dsimp simp)) - functor_unitIso_comp' c := by + functor_unitIso_comp c := by apply Quiver.Hom.unop_inj - ext + apply ConeMorphism.ext dsimp apply comp_id #align category_theory.limits.cocone_equivalence_op_cone_op CategoryTheory.Limits.coconeEquivalenceOpConeOp -attribute [simps] cocone_equivalence_op_cone_op +attribute [simps] coconeEquivalenceOpConeOp end section variable {F : J ⥤ Cᵒᵖ} - +/- Porting note: removed a few simps configs +`@[simps (config := + { rhsMd := semireducible + simpRhs := true })]` +and replace with `@[simps]`-/ -- Here and below we only automatically generate the `@[simp]` lemma for the `X` field, -- as we can write a simpler `rfl` lemma for the components of the natural transformation by hand. /-- Change a cocone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ -@[simps (config := - { rhsMd := semireducible - simpRhs := true })] +@[simps!] def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F where X := op c.X @@ -946,9 +965,7 @@ def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F #align category_theory.limits.cone_of_cocone_left_op CategoryTheory.Limits.coneOfCoconeLeftOp /-- Change a cone on `F : J ⥤ Cᵒᵖ` to a cocone on `F.left_op : Jᵒᵖ ⥤ C`. -/ -@[simps (config := - { rhsMd := semireducible - simpRhs := true })] +@[simps!] def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp where X := unop c.X @@ -959,7 +976,7 @@ def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp reduce the RHS using `expr.dsimp` and `expr.simp`, but for some reason the expression is not being simplified properly. -/ /-- Change a cone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ -@[simps x] +@[simps X] def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F where X := op c.X @@ -969,14 +986,12 @@ def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F @[simp] theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : (coconeOfConeLeftOp c).ι.app j = (c.π.app (op j)).op := by - dsimp only [cocone_of_cone_left_op] + dsimp only [coconeOfConeLeftOp] simp #align category_theory.limits.cocone_of_cone_left_op_ι_app CategoryTheory.Limits.coconeOfConeLeftOp_ι_app /-- Change a cocone on `F : J ⥤ Cᵒᵖ` to a cone on `F.left_op : Jᵒᵖ ⥤ C`. -/ -@[simps (config := - { rhsMd := semireducible - simpRhs := true })] +@[simps!] def coneLeftOpOfCocone (c : Cocone F) : Cone F.leftOp where X := unop c.X @@ -1074,15 +1089,17 @@ section variable (G : C ⥤ D) /-- The opposite cocone of the image of a cone is the image of the opposite cocone. -/ -@[simps (config := { rhsMd := semireducible })] -def mapConeOp (t : Cone F) : (G.mapCone t).op ≅ G.op.mapCocone t.op := - Cocones.ext (Iso.refl _) (by tidy) +-- @[simps (config := { rhsMd := semireducible })] +@[simps!] +def mapConeOp (t : Cone F) : (mapCone G t).op ≅ mapCocone G.op t.op := + Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_op CategoryTheory.Functor.mapConeOp /-- The opposite cone of the image of a cocone is the image of the opposite cone. -/ -@[simps (config := { rhsMd := semireducible })] -def mapCoconeOp {t : Cocone F} : (G.mapCocone t).op ≅ G.op.mapCone t.op := - Cones.ext (Iso.refl _) (by tidy) +-- @[simps (config := { rhsMd := semireducible })] +@[simps!] +def mapCoconeOp {t : Cocone F} : (mapCocone G t).op ≅ mapCone G.op t.op := + Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_op CategoryTheory.Functor.mapCoconeOp end From e1a7f9895bd2c4846fba5383c5e4f9d5463bb60c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 17 Feb 2023 19:10:12 -0500 Subject: [PATCH 120/126] lint --- Mathlib.lean | 2 +- Mathlib/CategoryTheory/Limits/Cones.lean | 110 ++++++++++++----------- 2 files changed, 61 insertions(+), 51 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index dcf15edc173f4..a540460a6b682 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -250,9 +250,9 @@ import Mathlib.CategoryTheory.Functor.InvIsos import Mathlib.CategoryTheory.Functor.ReflectsIso import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.Iso -import Mathlib.CategoryTheory.Limits.Cones import Mathlib.CategoryTheory.LiftingProperties.Adjunction import Mathlib.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Limits.Cones import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 08796ca03d20a..abf5a7c42648a 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -16,7 +16,7 @@ import Mathlib.CategoryTheory.Functor.ReflectsIso /-! # Cones and cocones -We define `cone F`, a cone over a functor `F`, +We define `Cone F`, a cone over a functor `F`, and `F.cones : Cᵒᵖ ⥤ Type`, the functor associating to `X` the cones over `F` with cone point `X`. A cone `c` is defined by specifying its cone point `c.X` and a natural transformation `c.π` @@ -27,11 +27,11 @@ as a wrapper for `c.π.naturality f` avoiding unneeded identity morphisms. We define `c.extend f`, where `c : cone F` and `f : Y ⟶ c.X` for some other `Y`, which replaces the cone point by `Y` and inserts `f` into each of the components of the cone. -Similarly we have `c.whisker F` producing a `cone (E ⋙ F)` +Similarly we have `c.whisker F` producing a `Cone (E ⋙ F)` We define morphisms of cones, and the category of cones. -We define `cone.postcompose α : cone F ⥤ cone G` for `α` a natural transformation `F ⟶ G`. +We define `Cone.postcompose α : cone F ⥤ cone G` for `α` a natural transformation `F ⟶ G`. And, of course, we dualise all this to cocones as well. @@ -39,7 +39,7 @@ For more results about the category of cones, see `cone_category.lean`. -/ --- morphism levels before object levels. See note [category_theory universes]. +-- morphism levels before object levels. See note [CategoryTheory universes]. universe v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ open CategoryTheory @@ -114,16 +114,16 @@ namespace Limits section --- attribute [local tidy] tactic.discrete_cases - -/-- A `c : cone F` is: +/-- A `c : Cone F` is: * an object `c.X` and * a natural transformation `c.π : c.X ⟶ F` from the constant `c.X` functor to `F`. -`cone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cones.obj X`. +`Cone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cones.obj X`. -/ structure Cone (F : J ⥤ C) where + /-- An object of `C` -/ X : C + /-- A natural transformation from the constant functor at `X` to `F` -/ π : (const J).obj X ⟶ F #align category_theory.limits.cone CategoryTheory.Limits.Cone @@ -146,14 +146,16 @@ theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : apply id_comp #align category_theory.limits.cone.w CategoryTheory.Limits.Cone.w -/-- A `c : cocone F` is +/-- A `c : Cocone F` is * an object `c.X` and * a natural transformation `c.ι : F ⟶ c.X` from `F` to the constant `c.X` functor. -`cocone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cocones.obj X`. +`Cocone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cocones.obj X`. -/ structure Cocone (F : J ⥤ C) where + /-- An object of `C` -/ X : C + /-- A natural transformation from `F` to the constant functor at `X` -/ ι : F ⟶ (const J).obj X #align category_theory.limits.cocone CategoryTheory.Limits.Cocone @@ -169,7 +171,7 @@ instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := }⟩ #align category_theory.limits.inhabited_cocone CategoryTheory.Limits.inhabitedCocone -@[reassoc (attr := simp)] +@[reassoc] theorem Cocone.w {F : J ⥤ C} (c : Cocone F) {j j' : J} (f : j ⟶ j') : F.map f ≫ c.ι.app j' = c.ι.app j := by rw [c.ι.naturality f] @@ -280,7 +282,9 @@ end Cocone commutes with the cone legs. -/ @[ext] structure ConeMorphism (A B : Cone F) where + /-- A morphism between the two vertex objects of the cones -/ Hom : A.X ⟶ B.X + /-- The triangle consisting of the two natural tranformations and `Hom` commutes -/ w : ∀ j : J, Hom ≫ B.π.app j = A.π.app j := by aesop_cat #align category_theory.limits.cone_morphism CategoryTheory.Limits.ConeMorphism #align category_theory.limits.cone_morphism.w' CategoryTheory.Limits.ConeMorphism.w @@ -365,7 +369,7 @@ def postcomposeEquivalence {G : J ⥤ C} (α : F ≅ G) : Cone F ≌ Cone G counitIso := NatIso.ofComponents (fun s => Cones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cones.postcompose_equivalence CategoryTheory.Limits.Cones.postcomposeEquivalence -/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cone F` to `cone (E ⋙ F)`. +/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `Cone F` to `Cone (E ⋙ F)`. -/ @[simps] def whiskering (E : K ⥤ J) : Cone F ⥤ Cone (E ⋙ F) @@ -450,11 +454,14 @@ def functorialityEquivalence (e : C ≌ D) : Cone F ≌ Cone (F ⋙ e.functor) : Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ { functor := functoriality F e.functor inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (postcomposeEquivalence f).functor - unitIso := NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) - counitIso := NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by aesop_cat)) (by aesop_cat) } + unitIso := + NatIso.ofComponents (fun c => Cones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) + counitIso := + NatIso.ofComponents (fun c => Cones.ext (e.counitIso.app _) (by aesop_cat)) (by aesop_cat) + } #align category_theory.limits.cones.functoriality_equivalence CategoryTheory.Limits.Cones.functorialityEquivalence -/-- If `F` reflects isomorphisms, then `cones.functoriality F` reflects isomorphisms +/-- If `F` reflects isomorphisms, then `Cones.functoriality F` reflects isomorphisms as well. -/ instance reflects_cone_isomorphism (F : C ⥤ D) [ReflectsIsomorphisms F] (K : J ⥤ C) : @@ -475,7 +482,9 @@ end Cones which commutes with the cocone legs. -/ @[ext] structure CoconeMorphism (A B : Cocone F) where + /-- A morphism between the (co)vertex objects in `C` -/ Hom : A.X ⟶ B.X + /-- The triangle made from the two natural transformations and `Hom` commutes -/ w : ∀ j : J, A.ι.app j ≫ Hom = B.ι.app j := by aesop_cat #align category_theory.limits.cocone_morphism CategoryTheory.Limits.CoconeMorphism #align category_theory.limits.cocone_morphism.w' CategoryTheory.Limits.CoconeMorphism.w @@ -556,7 +565,7 @@ def precomposeEquivalence {G : J ⥤ C} (α : G ≅ F) : Cocone F ≌ Cocone G w counitIso := NatIso.ofComponents (fun s => Cocones.ext (Iso.refl _) (by aesop_cat)) (by aesop_cat) #align category_theory.limits.cocones.precompose_equivalence CategoryTheory.Limits.Cocones.precomposeEquivalence -/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `cocone F` to `cocone (E ⋙ F)`. +/-- Whiskering on the left by `E : K ⥤ J` gives a functor from `Cocone F` to `Cocone (E ⋙ F)`. -/ @[simps] def whiskering (E : K ⥤ J) : Cocone F ⥤ Cocone (E ⋙ F) @@ -647,7 +656,8 @@ def functorialityEquivalence (e : C ≌ D) : Cocone F ≌ Cocone (F ⋙ e.functo Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ e.unitIso.symm ≪≫ Functor.rightUnitor _ { functor := functoriality F e.functor inverse := functoriality (F ⋙ e.functor) e.inverse ⋙ (precomposeEquivalence f.symm).functor - unitIso := NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) + unitIso := + NatIso.ofComponents (fun c => Cocones.ext (e.unitIso.app _) (by aesop_cat)) (by aesop_cat) counitIso := NatIso.ofComponents (fun c => @@ -715,37 +725,37 @@ def mapCoconeMorphism {c c' : Cocone F} (f : c ⟶ c') : mapCocone H c ⟶ mapCo (Cocones.functoriality F H).map f #align category_theory.functor.map_cocone_morphism CategoryTheory.Functor.mapCoconeMorphism -/-- If `H` is an equivalence, we invert `H.map_cone` and get a cone for `F` from a cone +/-- If `H` is an equivalence, we invert `H.mapCone` and get a cone for `F` from a cone for `F ⋙ H`.-/ def mapConeInv [IsEquivalence H] (c : Cone (F ⋙ H)) : Cone F := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).inverse.obj c #align category_theory.functor.map_cone_inv CategoryTheory.Functor.mapConeInv -/-- `map_cone` is the left inverse to `map_cone_inv`. -/ +/-- `mapCone` is the left inverse to `mapConeInv`. -/ def mapConeMapConeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone (F ⋙ H)) : mapCone H (mapConeInv H c) ≅ c := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).counitIso.app c #align category_theory.functor.map_cone_map_cone_inv CategoryTheory.Functor.mapConeMapConeInv -/-- `map_cone` is the right inverse to `map_cone_inv`. -/ +/-- `MapCone` is the right inverse to `mapConeInv`. -/ def mapConeInvMapCone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone F) : mapConeInv H (mapCone H c) ≅ c := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c #align category_theory.functor.map_cone_inv_map_cone CategoryTheory.Functor.mapConeInvMapCone -/-- If `H` is an equivalence, we invert `H.map_cone` and get a cone for `F` from a cone +/-- If `H` is an equivalence, we invert `H.mapCone` and get a cone for `F` from a cone for `F ⋙ H`.-/ def mapCoconeInv [IsEquivalence H] (c : Cocone (F ⋙ H)) : Cocone F := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).inverse.obj c #align category_theory.functor.map_cocone_inv CategoryTheory.Functor.mapCoconeInv -/-- `map_cocone` is the left inverse to `map_cocone_inv`. -/ +/-- `mapCocone` is the left inverse to `mapCoconeInv`. -/ def mapCoconeMapCoconeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone (F ⋙ H)) : mapCocone H (mapCoconeInv H c) ≅ c := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).counitIso.app c #align category_theory.functor.map_cocone_map_cocone_inv CategoryTheory.Functor.mapCoconeMapCoconeInv -/-- `map_cocone` is the right inverse to `map_cocone_inv`. -/ +/-- `mapCocone` is the right inverse to `mapCoconeInv`. -/ def mapCoconeInvMapCocone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone F) : mapCoconeInv H (mapCocone H c) ≅ c := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c @@ -758,9 +768,9 @@ def functorialityCompPostcompose {H H' : C ⥤ D} (α : H ≅ H') : NatIso.ofComponents (fun c => Cones.ext (α.app _) (by aesop_cat)) (by aesop_cat) #align category_theory.functor.functoriality_comp_postcompose CategoryTheory.Functor.functorialityCompPostcompose -/-- For `F : J ⥤ C`, given a cone `c : cone F`, and a natural isomorphism `α : H ≅ H'` for functors -`H H' : C ⥤ D`, the postcomposition of the cone `H.map_cone` using the isomorphism `α` is -isomorphic to the cone `H'.map_cone`. +/-- For `F : J ⥤ C`, given a cone `c : Cone F`, and a natural isomorphism `α : H ≅ H'` for functors +`H H' : C ⥤ D`, the postcomposition of the cone `H.mapCone` using the isomorphism `α` is +isomorphic to the cone `H'.mapCone`. -/ @[simps!] def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) : @@ -769,7 +779,7 @@ def postcomposeWhiskerLeftMapCone {H H' : C ⥤ D} (α : H ≅ H') (c : Cone F) #align category_theory.functor.postcompose_whisker_left_map_cone CategoryTheory.Functor.postcomposeWhiskerLeftMapCone /-- -`map_cone` commutes with `postcompose`. In particular, for `F : J ⥤ C`, given a cone `c : cone F`, a +`mapCone` commutes with `postcompose`. In particular, for `F : J ⥤ C`, given a cone `c : Cone F`, a natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious ways of producing a cone over `G ⋙ H`, and they are both isomorphic. -/ @@ -780,7 +790,7 @@ def mapConePostcompose {α : F ⟶ G} {c} : Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_postcompose CategoryTheory.Functor.mapConePostcompose -/-- `map_cone` commutes with `postcompose_equivalence` +/-- `mapCone` commutes with `postcomposeEquivalence` -/ @[simps!] def mapConePostcomposeEquivalenceFunctor {α : F ≅ G} {c} : @@ -789,7 +799,7 @@ def mapConePostcomposeEquivalenceFunctor {α : F ≅ G} {c} : Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_postcompose_equivalence_functor CategoryTheory.Functor.mapConePostcomposeEquivalenceFunctor -/-- `functoriality F _ ⋙ precompose (whisker_left F _)` simplifies to `functoriality F _`. -/ +/-- `functoriality F _ ⋙ precompose (whiskerLeft F _)` simplifies to `functoriality F _`. -/ @[simps!] def functorialityCompPrecompose {H H' : C ⥤ D} (α : H ≅ H') : Cocones.functoriality F H ⋙ Cocones.precompose (whiskerLeft F α.inv) ≅ @@ -798,9 +808,9 @@ def functorialityCompPrecompose {H H' : C ⥤ D} (α : H ≅ H') : #align category_theory.functor.functoriality_comp_precompose CategoryTheory.Functor.functorialityCompPrecompose /-- -For `F : J ⥤ C`, given a cocone `c : cocone F`, and a natural isomorphism `α : H ≅ H'` for functors -`H H' : C ⥤ D`, the precomposition of the cocone `H.map_cocone` using the isomorphism `α` is -isomorphic to the cocone `H'.map_cocone`. +For `F : J ⥤ C`, given a cocone `c : Cocone F`, and a natural isomorphism `α : H ≅ H'` for functors +`H H' : C ⥤ D`, the precomposition of the cocone `H.mapCocone` using the isomorphism `α` is +isomorphic to the cocone `H'.mapCocone`. -/ @[simps!] def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone F) : @@ -809,7 +819,7 @@ def precomposeWhiskerLeftMapCocone {H H' : C ⥤ D} (α : H ≅ H') (c : Cocone #align category_theory.functor.precompose_whisker_left_map_cocone CategoryTheory.Functor.precomposeWhiskerLeftMapCocone /-- `map_cocone` commutes with `precompose`. In particular, for `F : J ⥤ C`, given a cocone -`c : cocone F`, a natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious +`c : Cocone F`, a natural transformation `α : F ⟶ G` and a functor `H : C ⥤ D`, we have two obvious ways of producing a cocone over `G ⋙ H`, and they are both isomorphic. -/ @[simps!] @@ -819,7 +829,7 @@ def mapCoconePrecompose {α : F ⟶ G} {c} : Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_precompose CategoryTheory.Functor.mapCoconePrecompose -/-- `map_cocone` commutes with `precompose_equivalence` +/-- `mapCocone` commutes with `precomposeEquivalence` -/ @[simps!] def mapCoconePrecomposeEquivalenceFunctor {α : F ≅ G} {c} : @@ -828,14 +838,14 @@ def mapCoconePrecomposeEquivalenceFunctor {α : F ≅ G} {c} : Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cocone_precompose_equivalence_functor CategoryTheory.Functor.mapCoconePrecomposeEquivalenceFunctor -/-- `map_cone` commutes with `whisker` +/-- `mapCone` commutes with `whisker` -/ @[simps!] def mapConeWhisker {E : K ⥤ J} {c : Cone F} : mapCone H (c.whisker E) ≅ (mapCone H c).whisker E := Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_whisker CategoryTheory.Functor.mapConeWhisker -/-- `map_cocone` commutes with `whisker` +/-- `mapCocone` commutes with `whisker` -/ @[simps!] def mapCoconeWhisker {E : K ⥤ J} {c : Cocone F} : @@ -853,21 +863,21 @@ section variable {F : J ⥤ C} -/-- Change a `cocone F` into a `cone F.op`. -/ +/-- Change a `Cocone F` into a `Cone F.op`. -/ @[simps] def Cocone.op (c : Cocone F) : Cone F.op where X := Opposite.op c.X π := NatTrans.op c.ι #align category_theory.limits.cocone.op CategoryTheory.Limits.Cocone.op -/-- Change a `cone F` into a `cocone F.op`. -/ +/-- Change a `Cone F` into a `Cocone F.op`. -/ @[simps] def Cone.op (c : Cone F) : Cocone F.op where X := Opposite.op c.X ι := NatTrans.op c.π #align category_theory.limits.cone.op CategoryTheory.Limits.Cone.op -/-- Change a `cocone F.op` into a `cone F`. -/ +/-- Change a `Cocone F.op` into a `Cone F`. -/ @[simps] def Cocone.unop (c : Cocone F.op) : Cone F where @@ -875,7 +885,7 @@ def Cocone.unop (c : Cocone F.op) : Cone F π := NatTrans.removeOp c.ι #align category_theory.limits.cocone.unop CategoryTheory.Limits.Cocone.unop -/-- Change a `cone F.op` into a `cocone F`. -/ +/-- Change a `Cone F.op` into a `Cocone F`. -/ @[simps] def Cone.unop (c : Cone F.op) : Cocone F where @@ -956,7 +966,7 @@ variable {F : J ⥤ Cᵒᵖ} and replace with `@[simps]`-/ -- Here and below we only automatically generate the `@[simp]` lemma for the `X` field, -- as we can write a simpler `rfl` lemma for the components of the natural transformation by hand. -/-- Change a cocone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ +/-- Change a cocone on `F.leftOp : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ @[simps!] def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F where @@ -964,7 +974,7 @@ def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F π := NatTrans.removeLeftOp c.ι #align category_theory.limits.cone_of_cocone_left_op CategoryTheory.Limits.coneOfCoconeLeftOp -/-- Change a cone on `F : J ⥤ Cᵒᵖ` to a cocone on `F.left_op : Jᵒᵖ ⥤ C`. -/ +/-- Change a cone on `F : J ⥤ Cᵒᵖ` to a cocone on `F.leftOp : Jᵒᵖ ⥤ C`. -/ @[simps!] def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp where @@ -975,7 +985,7 @@ def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp /- When trying use `@[simps]` to generate the `ι_app` field of this definition, `@[simps]` tries to reduce the RHS using `expr.dsimp` and `expr.simp`, but for some reason the expression is not being simplified properly. -/ -/-- Change a cone on `F.left_op : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ +/-- Change a cone on `F.leftOp : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ @[simps X] def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F where @@ -990,7 +1000,7 @@ theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : simp #align category_theory.limits.cocone_of_cone_left_op_ι_app CategoryTheory.Limits.coconeOfConeLeftOp_ι_app -/-- Change a cocone on `F : J ⥤ Cᵒᵖ` to a cone on `F.left_op : Jᵒᵖ ⥤ C`. -/ +/-- Change a cocone on `F : J ⥤ Cᵒᵖ` to a cone on `F.leftOp : Jᵒᵖ ⥤ C`. -/ @[simps!] def coneLeftOpOfCocone (c : Cocone F) : Cone F.leftOp where @@ -1004,7 +1014,7 @@ section variable {F : Jᵒᵖ ⥤ C} -/-- Change a cocone on `F.right_op : J ⥤ Cᵒᵖ` to a cone on `F : Jᵒᵖ ⥤ C`. -/ +/-- Change a cocone on `F.rightOp : J ⥤ Cᵒᵖ` to a cone on `F : Jᵒᵖ ⥤ C`. -/ @[simps] def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F where @@ -1012,7 +1022,7 @@ def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F π := NatTrans.removeRightOp c.ι #align category_theory.limits.cone_of_cocone_right_op CategoryTheory.Limits.coneOfCoconeRightOp -/-- Change a cone on `F : Jᵒᵖ ⥤ C` to a cocone on `F.right_op : Jᵒᵖ ⥤ C`. -/ +/-- Change a cone on `F : Jᵒᵖ ⥤ C` to a cocone on `F.rightOp : Jᵒᵖ ⥤ C`. -/ @[simps] def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp where @@ -1020,7 +1030,7 @@ def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp ι := NatTrans.rightOp c.π #align category_theory.limits.cocone_right_op_of_cone CategoryTheory.Limits.coconeRightOpOfCone -/-- Change a cone on `F.right_op : J ⥤ Cᵒᵖ` to a cocone on `F : Jᵒᵖ ⥤ C`. -/ +/-- Change a cone on `F.rightOp : J ⥤ Cᵒᵖ` to a cocone on `F : Jᵒᵖ ⥤ C`. -/ @[simps] def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F where @@ -1028,7 +1038,7 @@ def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F ι := NatTrans.removeRightOp c.π #align category_theory.limits.cocone_of_cone_right_op CategoryTheory.Limits.coconeOfConeRightOp -/-- Change a cocone on `F : Jᵒᵖ ⥤ C` to a cone on `F.right_op : J ⥤ Cᵒᵖ`. -/ +/-- Change a cocone on `F : Jᵒᵖ ⥤ C` to a cone on `F.rightOp : J ⥤ Cᵒᵖ`. -/ @[simps] def coneRightOpOfCocone (c : Cocone F) : Cone F.rightOp where @@ -1089,14 +1099,14 @@ section variable (G : C ⥤ D) /-- The opposite cocone of the image of a cone is the image of the opposite cocone. -/ --- @[simps (config := { rhsMd := semireducible })] +-- Porting note: removed @[simps (config := { rhsMd := semireducible })] and replaced with @[simps!] def mapConeOp (t : Cone F) : (mapCone G t).op ≅ mapCocone G.op t.op := Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.functor.map_cone_op CategoryTheory.Functor.mapConeOp /-- The opposite cone of the image of a cocone is the image of the opposite cone. -/ --- @[simps (config := { rhsMd := semireducible })] +-- Porting note: removed @[simps (config := { rhsMd := semireducible })] and replaced with @[simps!] def mapCoconeOp {t : Cocone F} : (mapCocone G t).op ≅ mapCone G.op t.op := Cones.ext (Iso.refl _) (by aesop_cat) From bd0f2cfb394806ac7fbe118e298dd39e80f4af25 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 11:09:19 -0500 Subject: [PATCH 121/126] add updated files --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 10 +-- Mathlib/CategoryTheory/CommSq.lean | 19 +++-- Mathlib/CategoryTheory/DiscreteCategory.lean | 78 ++++++------------- .../Limits/Shapes/StrongEpi.lean | 47 ++++++----- Mathlib/CategoryTheory/Yoneda.lean | 33 -------- 5 files changed, 60 insertions(+), 127 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index a5b65f05aaa50..610789b9be365 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -177,14 +177,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 = 𝟙 _ := 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 = 𝟙 _ := by ext; dsimp erw [← adj.homEquiv_unit, ← Equiv.eq_symm_apply, adj.homEquiv_counit] simp @@ -387,10 +387,10 @@ 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 +/- 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 +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 diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index 12ffba8bc2641..9a7f103ff036c 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -121,7 +121,7 @@ structure LiftStruct (sq : CommSq f i p g) where namespace LiftStruct -/-- A `lift_struct` for a commutative square gives a `lift_struct` for the +/-- A `LiftStruct` for a commutative square gives a `LiftStruct` for the corresponding square in the opposite category. -/ @[simps] def op {sq : CommSq f i p g} (l : LiftStruct sq) : LiftStruct sq.op @@ -131,8 +131,8 @@ def op {sq : CommSq f i p g} (l : LiftStruct sq) : LiftStruct sq.op 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. -/ +/-- A `LiftStruct` for a commutative square in the opposite category +gives a `LiftStruct` 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 @@ -142,7 +142,7 @@ def unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B 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 +/-- Equivalences of `LiftStruct` 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 @@ -153,7 +153,7 @@ def opEquiv (sq : CommSq f i p g) : LiftStruct sq ≃ LiftStruct sq.op 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 +/-- Equivalences of `LiftStruct` 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 @@ -171,8 +171,7 @@ instance subsingleton_liftStruct_of_epi (sq : CommSq f i p g) [Epi i] : ⟨fun l₁ l₂ => by ext rw [← cancel_epi i] - simp only [LiftStruct.fac_left] - ⟩ + 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] : @@ -186,9 +185,9 @@ 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`. -/ +/-- The assertion that a square has a `LiftStruct`. -/ class HasLift : Prop where - /-- Square has a `lift_struct`. -/ + /-- Square has a `LiftStruct`. -/ exists_lift : Nonempty sq.LiftStruct #align category_theory.comm_sq.has_lift CategoryTheory.CommSq.HasLift @@ -220,7 +219,7 @@ theorem iff_unop {A B X Y : Cᵒᵖ} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} { end HasLift -/-- A choice of a diagonal morphism that is part of a `lift_struct` when +/-- A choice of a diagonal morphism that is part of a `LiftStruct` when the square has a lift. -/ noncomputable def lift [hsq : HasLift sq] : B ⟶ X := hsq.exists_lift.some.l diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index cac60eb13ddee..9518b44d11cdb 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -44,7 +44,7 @@ universe v₁ v₂ v₃ u₁ u₁' u₂ u₃ -- This is intentionally a structure rather than a type synonym -- to enforce using `DiscreteEquiv` (or `Discrete.mk` and `Discrete.as`) to move between --- `discrete α` and `α`. Otherwise there is too much API leakage. +-- `Discrete α` and `α`. Otherwise there is too much API leakage. /-- A wrapper for promoting any type to a category, with the only morphisms being equalities. -/ @@ -64,8 +64,7 @@ theorem Discrete.mk_as {α : Type u₁} (X : Discrete α) : Discrete.mk X.as = X /-- `Discrete α` is equivalent to the original type `α`.-/ @[simps] -def discreteEquiv {α : Type u₁} : Discrete α ≃ α - where +def discreteEquiv {α : Type u₁} : Discrete α ≃ α where toFun := Discrete.as invFun := Discrete.mk left_inv := by aesop_cat @@ -82,8 +81,7 @@ somewhat annoyingly we have to define `X ⟶ Y` as `ULift (PLift (X = Y))`. See -/ -instance discreteCategory (α : Type u₁) : SmallCategory (Discrete α) - where +instance discreteCategory (α : Type u₁) : SmallCategory (Discrete α) where Hom X Y := ULift (PLift (X.as = Y.as)) id X := ULift.up (PLift.up rfl) comp {X Y Z} g f := by @@ -107,22 +105,15 @@ instance [Subsingleton α] : Subsingleton (Discrete α) := ext apply Subsingleton.elim⟩ -/- ./././Mathport/Syntax/Translate/Expr.lean:334:4: warning: unsupported (TODO): `[tacs] -/ -/-- A simple tactic to run `cases` on any `discrete α` hypotheses. -/ ---unsafe def _root_.tactic.discrete_cases : tactic Unit := --- sorry ---#align tactic.discrete_cases tactic.discrete_cases - - ---run_cmd --- add_interactive [`` tactic.discrete_cases] +/- +Porting note: It seems that `aesop` currently has no way to add lemmas locally. ---attribute [local tidy] tactic.discrete_cases ---`[cases_matching* [discrete _, (_ : discrete _) ⟶ (_ : discrete _), plift _]] +attribute [local tidy] tactic.discrete_cases +`[cases_matching* [discrete _, (_ : discrete _) ⟶ (_ : discrete _), PLift _]] +-/ ---macro (name := discrete_cases) "discrete_cases" : tactic => --- `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) /- Porting note: rewrote `discrete_cases` tactic -/ +/-- A simple tactic to run `cases` on any `discrete α` hypotheses. -/ macro "discrete_cases": tactic => `(tactic|casesm* Discrete _, (_ : Discrete _) ⟶ (_ : Discrete _), PLift _) @@ -152,7 +143,7 @@ protected abbrev eqToIso {X Y : Discrete α} (h : X.as = Y.as) : X ≅ Y := exact h) #align category_theory.discrete.eq_to_iso CategoryTheory.Discrete.eqToIso -/-- A variant of `eq_to_hom` that lifts terms to the discrete category. -/ +/-- A variant of `eqToHom` that lifts terms to the discrete category. -/ abbrev eqToHom' {a b : α} (h : a = b) : Discrete.mk a ⟶ Discrete.mk b := Discrete.eqToHom h #align category_theory.discrete.eq_to_hom' CategoryTheory.Discrete.eqToHom' @@ -172,12 +163,8 @@ variable {C : Type u₂} [Category.{v₂} C] instance {I : Type u₁} {i j : Discrete I} (f : i ⟶ j) : IsIso f := ⟨⟨Discrete.eqToHom (eq_of_hom f).symm, by aesop_cat⟩⟩ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ -/-- Any function `I → C` gives a functor `Discrete I ⥤ C`. --/ -def functor {I : Type u₁} (F : I → C) : Discrete I ⥤ C - where +/-- Any function `I → C` gives a functor `Discrete I ⥤ C`.-/ +def functor {I : Type u₁} (F : I → C) : Discrete I ⥤ C where obj := F ∘ Discrete.as map {X Y} f := by dsimp @@ -208,8 +195,6 @@ def functorComp {I : Type u₁} {J : Type u₁'} (f : J → C) (g : I → J) : NatIso.ofComponents (fun X => Iso.refl _) (by aesop_cat) #align category_theory.discrete.functor_comp CategoryTheory.Discrete.functorComp -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ /-- For functors out of a discrete category, a natural transformation is just a collection of maps, as the naturality squares are trivial. @@ -218,24 +203,20 @@ as the naturality squares are trivial. def natTrans {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ⟶ G.obj i) : F ⟶ G where app := f - naturality {X Y} g := by - rcases g with ⟨⟨g⟩⟩ + naturality := fun {X Y} ⟨⟨g⟩⟩ => by discrete_cases rcases g change F.map (𝟙 _) ≫ _ = _ ≫ G.map (𝟙 _) simp #align category_theory.discrete.nat_trans CategoryTheory.Discrete.natTrans -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ /-- For functors out of a discrete category, a natural isomorphism is just a collection of isomorphisms, as the naturality squares are trivial. -/ @[simps!] def natIso {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) : F ≅ G := - NatIso.ofComponents f fun g => by - rcases g with ⟨⟨g⟩⟩ + NatIso.ofComponents f fun ⟨⟨g⟩⟩ => by discrete_cases rcases g change F.map (𝟙 _) ≫ _ = _ ≫ G.map (𝟙 _) @@ -247,8 +228,6 @@ theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discret (i : Discrete I) : (Discrete.natIso f).app i = f i := by aesop_cat #align category_theory.discrete.nat_iso_app CategoryTheory.Discrete.natIso_app -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ /-- Every functor `F` from a discrete category is naturally isomorphic (actually, equal) to `discrete.functor (F.obj)`. -/ @[simp] @@ -263,16 +242,11 @@ def compNatIsoDiscrete {I : Type u₁} {D : Type u₃} [Category.{v₃} D] (F : natIso fun _ => Iso.refl _ #align category_theory.discrete.comp_nat_iso_discrete CategoryTheory.Discrete.compNatIsoDiscrete -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ -/-- We can promote a type-level `equiv` to +/-- We can promote a type-level `Equiv` to an equivalence between the corresponding `discrete` categories. -/ @[simps] -def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ Discrete J - where +def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ Discrete J where functor := Discrete.functor (Discrete.mk ∘ (e : I → J)) inverse := Discrete.functor (Discrete.mk ∘ (e.symm : J → I)) unitIso := @@ -289,10 +263,9 @@ def equivalence {I : Type u₁} {J : Type u₂} (e : I ≃ J) : Discrete I ≌ D simp) #align category_theory.discrete.equivalence CategoryTheory.Discrete.equivalence -/-- We can convert an equivalence of `discrete` categories to a type-level `equiv`. -/ +/-- We can convert an equivalence of `discrete` categories to a type-level `Equiv`. -/ @[simps] -def equivOfEquivalence {α : Type u₁} {β : Type u₂} (h : Discrete α ≌ Discrete β) : α ≃ β - where +def equivOfEquivalence {α : Type u₁} {β : Type u₂} (h : Discrete α ≌ Discrete β) : α ≃ β where toFun := Discrete.as ∘ h.functor.obj ∘ Discrete.mk invFun := Discrete.as ∘ h.inverse.obj ∘ Discrete.mk left_inv a := by simpa using eq_of_hom (h.unitIso.app (Discrete.mk a)).2 @@ -307,27 +280,23 @@ variable {J : Type v₁} open Opposite -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:72:18: unsupported -\ non-interactive tactic tactic.op_induction' -/ -/- ./././Mathport/Syntax/Translate/Tactic/Builtin.lean:76:14: unsupported -\ tactic `discrete_cases #[] -/ /-- A discrete category is equivalent to its opposite category. -/ @[simps! functor_obj_as inverse_obj] protected def opposite (α : Type u₁) : (Discrete α)ᵒᵖ ≌ Discrete α := let F : Discrete α ⥤ (Discrete α)ᵒᵖ := Discrete.functor fun x => op (Discrete.mk x) Equivalence.mk F.leftOp F - (NatIso.ofComponents (fun ⟨X⟩ => Iso.refl _) - <| fun {X Y} f => by + (NatIso.ofComponents (fun ⟨X⟩ => Iso.refl _) <| fun {X Y} ⟨⟨f⟩⟩ => by induction X using Opposite.rec induction Y using Opposite.rec - rcases f with ⟨⟨f⟩⟩ discrete_cases rcases f aesop_cat) (Discrete.natIso <| fun ⟨X⟩ => Iso.refl _) + /- + Porting note: + The following is what was generated by mathport: + refine' Equivalence.mk (F.leftOp) F _ (Discrete.natIso fun X => @@ -361,4 +330,3 @@ theorem functor_map_id (F : Discrete J ⥤ C) {j : Discrete J} (f : j ⟶ j) : F end Discrete end CategoryTheory - diff --git a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean index a71d94949664b..ef32f913c8954 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/StrongEpi.lean @@ -26,7 +26,7 @@ Besides the definition, we show that * 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 +We also define classes `StrongMonoCategory` and `StrongEpiCategory` for categories in which every monomorphism or epimorphism is strong, and deduce that these categories are balanced. ## TODO @@ -59,11 +59,11 @@ class StrongEpi (f : P ⟶ Q) : Prop where theorem StrongEpi.mk' {f : P ⟶ Q} [Epi f] - (hf : ∀ (X Y : C) (z : X ⟶ Y) + (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⟩ } + 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 @@ -76,10 +76,10 @@ class StrongMono (f : P ⟶ Q) : Prop where #align category_theory.strong_mono CategoryTheory.StrongMono theorem StrongMono.mk' {f : P ⟶ Q} [Mono f] - (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 + (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⟩ + 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] StrongEpi.llp @@ -117,7 +117,7 @@ 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 := fun {X} {Y} z _ => by + 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] @@ -130,27 +130,25 @@ theorem strongEpi_of_strongEpi [StrongEpi (f ≫ g)] : StrongEpi g := /-- 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 := fun {X} {Y} z => by + rlp := fun {X Y} z => by intros constructor intro u v sq - have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by + 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 +instance (priority := 100) strongEpi_of_isIso [IsIso f] : StrongEpi f where epi := by infer_instance - llp {X} {Y} z := HasLiftingProperty.of_left_iso _ _ + 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 +instance (priority := 100) strongMono_of_isIso [IsIso f] : StrongMono f where mono := by infer_instance - rlp {X} {Y} z := HasLiftingProperty.of_right_iso _ _ + 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'} @@ -159,7 +157,7 @@ theorem StrongEpi.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} 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 HasLiftingProperty.of_arrow_iso_left e z } #align category_theory.strong_epi.of_arrow_iso CategoryTheory.StrongEpi.of_arrow_iso @@ -170,7 +168,7 @@ theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} 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 HasLiftingProperty.of_arrow_iso_right z e } #align category_theory.strong_mono.of_arrow_iso CategoryTheory.StrongMono.of_arrow_iso @@ -178,13 +176,13 @@ theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} 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[StrongEpi.of_arrow_iso e, StrongEpi.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[StrongMono.of_arrow_iso e, StrongMono.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 @@ -208,12 +206,14 @@ 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 +#align category_theory.strong_epi_category.strong_epi_of_epi CategoryTheory.StrongEpiCategory.strongEpi_of_epi /-- 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 +#align category_theory.strong_mono_category.strong_mono_of_mono CategoryTheory.StrongMonoCategory.strongMono_of_mono end @@ -229,8 +229,8 @@ section attribute [local instance] strongEpi_of_epi -instance (priority := 100) balanced_of_strongEpiCategory [StrongEpiCategory C] : Balanced C - where isIso_of_mono_of_epi _ _ _ := isIso_of_mono_of_strongEpi _ +instance (priority := 100) balanced_of_strongEpiCategory [StrongEpiCategory C] : Balanced C where + isIso_of_mono_of_epi _ _ _ := isIso_of_mono_of_strongEpi _ #align category_theory.balanced_of_strong_epi_category CategoryTheory.balanced_of_strongEpiCategory end @@ -239,11 +239,10 @@ section attribute [local instance] strongMono_of_mono -instance (priority := 100) balanced_of_strongMonoCategory [StrongMonoCategory C] : Balanced C - where isIso_of_mono_of_epi _ _ _ := isIso_of_epi_of_strongMono _ +instance (priority := 100) balanced_of_strongMonoCategory [StrongMonoCategory C] : Balanced C 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 diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 7773489921444..6d0fbd315294c 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -217,24 +217,6 @@ noncomputable def reprF : yoneda.obj F.reprX ⟶ F := Representable.has_representation.choose_spec.choose #align category_theory.functor.repr_f CategoryTheory.Functor.reprF -/- warning: category_theory.functor.repr_x clashes with -\ category_theory.functor.repr_X -> CategoryTheory.Functor.reprX -warning: category_theory.functor.repr_x -> CategoryTheory.Functor.reprX is a -dubious translation: lean 3 declaration is forall {C : Type.{u2}} [_inst_1 : -CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, -succ u1} (Opposite.{succ u2} C) (CategoryTheory.Category.opposite.{u1, u2} C -_inst_1) Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : -CategoryTheory.Functor.Representable.{u1, u2} C _inst_1 F], -CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) -(CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} -CategoryTheory.types.{u1} F (Opposite.op.{succ u2} C -(CategoryTheory.Functor.reprX.{u1, u2} C _inst_1 F _inst_2)) but is expected to -have type forall {C : Type.{u2}} [_inst_1 : CategoryTheory.Category.{u1, u2} C] -(F : CategoryTheory.Functor.{u1, u1, u2, succ u1} (Opposite.{succ u2} C) -(CategoryTheory.Category.opposite.{u1, u2} C _inst_1) Type.{u1} -CategoryTheory.types.{u1}) [_inst_2 : CategoryTheory.Functor.Representable.{u1, -u2} C _inst_1 F], C Case conversion may be inaccurate. Consider using '#align -category_theory.functor.repr_x CategoryTheory.Functor.reprXₓ'. -/ /-- The representing element for the representable functor `F`, sometimes called the universal element of the functor. -/ @@ -284,21 +266,6 @@ noncomputable def coreprF : coyoneda.obj (op F.coreprX) ⟶ F := Corepresentable.has_corepresentation.choose_spec.choose #align category_theory.functor.corepr_f CategoryTheory.Functor.coreprF -/- warning: category_theory.functor.corepr_x clashes with -\ category_theory.functor.corepr_X -> CategoryTheory.Functor.coreprX -warning: category_theory.functor.corepr_x -> CategoryTheory.Functor.coreprX is -a dubious translation: lean 3 declaration is forall {C : Type.{u2}} [_inst_1 : -CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, -succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : -CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], -CategoryTheory.Functor.obj.{u1, u1, u2, succ u1} C _inst_1 Type.{u1} -CategoryTheory.types.{u1} F (CategoryTheory.Functor.coreprX.{u1, u2} C _inst_1 -F _inst_2) but is expected to have type forall {C : Type.{u2}} [_inst_1 : -CategoryTheory.Category.{u1, u2} C] (F : CategoryTheory.Functor.{u1, u1, u2, -succ u1} C _inst_1 Type.{u1} CategoryTheory.types.{u1}) [_inst_2 : -CategoryTheory.Functor.Corepresentable.{u1, u2} C _inst_1 F], C Case conversion -may be inaccurate. Consider using '#align category_theory.functor.corepr_x -CategoryTheory.Functor.coreprXₓ'. -/ /-- The representing element for the corepresentable functor `F`, sometimes called the universal element of the functor. -/ From 64d878d946a81ccd014ee4cbda51c3d84dd9447a Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 11:14:32 -0500 Subject: [PATCH 122/126] replace X with pt for cone field --- Mathlib/CategoryTheory/Limits/Cones.lean | 110 +++++++++++------------ 1 file changed, 55 insertions(+), 55 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index abf5a7c42648a..c24b968f77ed5 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -19,13 +19,13 @@ import Mathlib.CategoryTheory.Functor.ReflectsIso We define `Cone F`, a cone over a functor `F`, and `F.cones : Cᵒᵖ ⥤ Type`, the functor associating to `X` the cones over `F` with cone point `X`. -A cone `c` is defined by specifying its cone point `c.X` and a natural transformation `c.π` -from the constant `c.X` valued functor to `F`. +A cone `c` is defined by specifying its cone point `c.pt` and a natural transformation `c.π` +from the constant `c.pt` valued functor to `F`. We provide `c.w f : c.π.app j ≫ F.map f = c.π.app j'` for any `f : j ⟶ j'` as a wrapper for `c.π.naturality f` avoiding unneeded identity morphisms. -We define `c.extend f`, where `c : cone F` and `f : Y ⟶ c.X` for some other `Y`, +We define `c.extend f`, where `c : cone F` and `f : Y ⟶ c.pt` for some other `Y`, which replaces the cone point by `Y` and inserts `f` into each of the components of the cone. Similarly we have `c.whisker F` producing a `Cone (E ⋙ F)` @@ -115,20 +115,20 @@ namespace Limits section /-- A `c : Cone F` is: -* an object `c.X` and -* a natural transformation `c.π : c.X ⟶ F` from the constant `c.X` functor to `F`. +* an object `c.pt` and +* a natural transformation `c.π : c.pt ⟶ F` from the constant `c.pt` functor to `F`. `Cone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cones.obj X`. -/ structure Cone (F : J ⥤ C) where /-- An object of `C` -/ - X : C + pt : C /-- A natural transformation from the constant functor at `X` to `F` -/ - π : (const J).obj X ⟶ F + π : (const J).obj pt ⟶ F #align category_theory.limits.cone CategoryTheory.Limits.Cone instance inhabitedCone (F : Discrete PUnit ⥤ C) : Inhabited (Cone F) := - ⟨{ X := F.obj ⟨⟨⟩⟩ + ⟨{ pt := F.obj ⟨⟨⟩⟩ π := { app := fun ⟨⟨⟩⟩ => 𝟙 _ naturality := by intro X Y f @@ -147,20 +147,20 @@ theorem Cone.w {F : J ⥤ C} (c : Cone F) {j j' : J} (f : j ⟶ j') : #align category_theory.limits.cone.w CategoryTheory.Limits.Cone.w /-- A `c : Cocone F` is -* an object `c.X` and -* a natural transformation `c.ι : F ⟶ c.X` from `F` to the constant `c.X` functor. +* an object `c.pt` and +* a natural transformation `c.ι : F ⟶ c.pt` from `F` to the constant `c.pt` functor. `Cocone F` is equivalent, via `cone.equiv` below, to `Σ X, F.cocones.obj X`. -/ structure Cocone (F : J ⥤ C) where /-- An object of `C` -/ - X : C + pt : C /-- A natural transformation from `F` to the constant functor at `X` -/ - ι : F ⟶ (const J).obj X + ι : F ⟶ (const J).obj pt #align category_theory.limits.cocone CategoryTheory.Limits.Cocone instance inhabitedCocone (F : Discrete PUnit ⥤ C) : Inhabited (Cocone F) := - ⟨{ X := F.obj ⟨⟨⟩⟩ + ⟨{ pt := F.obj ⟨⟨⟩⟩ ι := { app := fun ⟨⟨⟩⟩ => 𝟙 _ naturality := by intro X Y f @@ -188,9 +188,9 @@ namespace Cone @[simps!] def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X where - hom c := ⟨op c.X, c.π⟩ + hom c := ⟨op c.pt, c.π⟩ inv c := - { X := c.1.unop + { pt := c.1.unop π := c.2 } hom_inv_id := by funext X @@ -204,7 +204,7 @@ def equiv (F : J ⥤ C) : Cone F ≅ ΣX, F.cones.obj X /-- A map to the vertex of a cone naturally induces a cone by composition. -/ @[simps] -def extensions (c : Cone F) : yoneda.obj c.X ⋙ uliftFunctor.{u₁} ⟶ F.cones where +def extensions (c : Cone F) : yoneda.obj c.pt ⋙ uliftFunctor.{u₁} ⟶ F.cones where app X f := (const J).map f.down ≫ c.π naturality := by intros X Y f @@ -215,8 +215,8 @@ def extensions (c : Cone F) : yoneda.obj c.X ⋙ uliftFunctor.{u₁} ⟶ F.cones /-- A map to the vertex of a cone induces a cone by composition. -/ @[simps] -def extend (c : Cone F) {X : C} (f : X ⟶ c.X) : Cone F := - { X := X +def extend (c : Cone F) {X : C} (f : X ⟶ c.pt) : Cone F := + { pt := X π := c.extensions.app (op X) ⟨f⟩ } #align category_theory.limits.cone.extend CategoryTheory.Limits.Cone.extend @@ -224,7 +224,7 @@ def extend (c : Cone F) {X : C} (f : X ⟶ c.X) : Cone F := @[simps] def whisker (E : K ⥤ J) (c : Cone F) : Cone (E ⋙ F) where - X := c.X + pt := c.pt π := whiskerLeft E c.π #align category_theory.limits.cone.whisker CategoryTheory.Limits.Cone.whisker @@ -235,9 +235,9 @@ namespace Cocone /-- The isomorphism between a cocone on `F` and an element of the functor `F.cocones`. -/ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X where - hom c := ⟨c.X, c.ι⟩ + hom c := ⟨c.pt, c.ι⟩ inv c := - { X := c.1 + { pt := c.1 ι := c.2 } hom_inv_id := by funext X @@ -251,7 +251,7 @@ def equiv (F : J ⥤ C) : Cocone F ≅ ΣX, F.cocones.obj X /-- A map from the vertex of a cocone naturally induces a cocone by composition. -/ @[simps] -def extensions (c : Cocone F) : coyoneda.obj (op c.X) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where +def extensions (c : Cocone F) : coyoneda.obj (op c.pt) ⋙ uliftFunctor.{u₁} ⟶ F.cocones where app X f := c.ι ≫ (const J).map f.down naturality := fun {X} {Y} f => by dsimp [coyoneda,const] @@ -261,8 +261,8 @@ def extensions (c : Cocone F) : coyoneda.obj (op c.X) ⋙ uliftFunctor.{u₁} /-- A map from the vertex of a cocone induces a cocone by composition. -/ @[simps] -def extend (c : Cocone F) {Y : C} (f : c.X ⟶ Y) : Cocone F where - X := Y +def extend (c : Cocone F) {Y : C} (f : c.pt ⟶ Y) : Cocone F where + pt := Y ι := c.extensions.app Y ⟨f⟩ #align category_theory.limits.cocone.extend CategoryTheory.Limits.Cocone.extend @@ -272,7 +272,7 @@ version. @[simps] def whisker (E : K ⥤ J) (c : Cocone F) : Cocone (E ⋙ F) where - X := c.X + pt := c.pt ι := whiskerLeft E c.ι #align category_theory.limits.cocone.whisker CategoryTheory.Limits.Cocone.whisker @@ -283,7 +283,7 @@ commutes with the cone legs. -/ @[ext] structure ConeMorphism (A B : Cone F) where /-- A morphism between the two vertex objects of the cones -/ - Hom : A.X ⟶ B.X + Hom : A.pt ⟶ B.pt /-- The triangle consisting of the two natural tranformations and `Hom` commutes -/ w : ∀ j : J, Hom ≫ B.π.app j = A.π.app j := by aesop_cat #align category_theory.limits.cone_morphism CategoryTheory.Limits.ConeMorphism @@ -300,7 +300,7 @@ instance inhabitedConeMorphism (A : Cone F) : Inhabited (ConeMorphism A A) := instance Cone.category : Category (Cone F) where Hom A B := ConeMorphism A B comp f g := { Hom := f.Hom ≫ g.Hom } - id B := { Hom := 𝟙 B.X } + id B := { Hom := 𝟙 B.pt } #align category_theory.limits.cone.category CategoryTheory.Limits.Cone.category namespace Cones @@ -310,7 +310,7 @@ namespace Cones maps. -/ -- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets [CategoryTheory]), simps] -def ext {c c' : Cone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' where +def ext {c c' : Cone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.π.app j = φ.hom ≫ c'.π.app j) : c ≅ c' where hom := { Hom := φ.hom } inv := { Hom := φ.inv @@ -319,7 +319,7 @@ def ext {c c' : Cone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.π.app j = φ.hom ≫ /-- Eta rule for cones. -/ @[simps!] -def eta (c : Cone F) : c ≅ ⟨c.X, c.π⟩ := +def eta (c : Cone F) : c ≅ ⟨c.pt, c.π⟩ := Cones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cones.eta CategoryTheory.Limits.Cones.eta @@ -338,7 +338,7 @@ Functorially postcompose a cone for `F` by a natural transformation `F ⟶ G` to def postcompose {G : J ⥤ C} (α : F ⟶ G) : Cone F ⥤ Cone G where obj c := - { X := c.X + { pt := c.pt π := c.π ≫ α } map f := { Hom := f.Hom } #align category_theory.limits.cones.postcompose CategoryTheory.Limits.Cones.postcompose @@ -413,7 +413,7 @@ variable (F) /-- Forget the cone structure and obtain just the cone point. -/ @[simps] def forget : Cone F ⥤ C where - obj t := t.X + obj t := t.pt map f := f.Hom #align category_theory.limits.cones.forget CategoryTheory.Limits.Cones.forget @@ -423,7 +423,7 @@ variable (G : C ⥤ D) @[simps] def functoriality : Cone F ⥤ Cone (F ⋙ G) where obj A := - { X := G.obj A.X + { pt := G.obj A.pt π := { app := fun j => G.map (A.π.app j) naturality := by intros; erw [← G.map_comp]; aesop_cat } } @@ -483,7 +483,7 @@ which commutes with the cocone legs. -/ @[ext] structure CoconeMorphism (A B : Cocone F) where /-- A morphism between the (co)vertex objects in `C` -/ - Hom : A.X ⟶ B.X + Hom : A.pt ⟶ B.pt /-- The triangle made from the two natural transformations and `Hom` commutes -/ w : ∀ j : J, A.ι.app j ≫ Hom = B.ι.app j := by aesop_cat #align category_theory.limits.cocone_morphism CategoryTheory.Limits.CoconeMorphism @@ -499,7 +499,7 @@ attribute [reassoc (attr := simp)] CoconeMorphism.w instance Cocone.category : Category (Cocone F) where Hom A B := CoconeMorphism A B comp f g := { Hom := f.Hom ≫ g.Hom } - id B := { Hom := 𝟙 B.X } + id B := { Hom := 𝟙 B.pt } #align category_theory.limits.cocone.category CategoryTheory.Limits.Cocone.category namespace Cocones @@ -509,7 +509,7 @@ namespace Cocones maps. -/ -- Porting note: `@[ext]` used to accept lemmas like this. Now we add an aesop rule @[aesop apply safe (rule_sets [CategoryTheory]), simps] -def ext {c c' : Cocone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) +def ext {c c' : Cocone F} (φ : c.pt ≅ c'.pt) (w : ∀ j, c.ι.app j ≫ φ.hom = c'.ι.app j) : c ≅ c' where hom := { Hom := φ.hom } inv := @@ -519,7 +519,7 @@ def ext {c c' : Cocone F} (φ : c.X ≅ c'.X) (w : ∀ j, c.ι.app j ≫ φ.hom /-- Eta rule for cocones. -/ @[simps!] -def eta (c : Cocone F) : c ≅ ⟨c.X, c.ι⟩ := +def eta (c : Cocone F) : c ≅ ⟨c.pt, c.ι⟩ := Cocones.ext (Iso.refl _) (by aesop_cat) #align category_theory.limits.cocones.eta CategoryTheory.Limits.Cocones.eta @@ -537,7 +537,7 @@ for `G`. -/ @[simps] def precompose {G : J ⥤ C} (α : G ⟶ F) : Cocone F ⥤ Cocone G where obj c := - { X := c.X + { pt := c.pt ι := α ≫ c.ι } map f := { Hom := f.Hom } #align category_theory.limits.cocones.precompose CategoryTheory.Limits.Cocones.precompose @@ -613,7 +613,7 @@ variable (F) /-- Forget the cocone structure and obtain just the cocone point. -/ @[simps] def forget : Cocone F ⥤ C where - obj t := t.X + obj t := t.pt map f := f.Hom #align category_theory.limits.cocones.forget CategoryTheory.Limits.Cocones.forget @@ -623,7 +623,7 @@ variable (G : C ⥤ D) @[simps] def functoriality : Cocone F ⥤ Cocone (F ⋙ G) where obj A := - { X := G.obj A.X + { pt := G.obj A.pt ι := { app := fun j => G.map (A.ι.app j) naturality := by intros; erw [← G.map_comp]; aesop_cat } } @@ -866,14 +866,14 @@ variable {F : J ⥤ C} /-- Change a `Cocone F` into a `Cone F.op`. -/ @[simps] def Cocone.op (c : Cocone F) : Cone F.op where - X := Opposite.op c.X + pt := Opposite.op c.pt π := NatTrans.op c.ι #align category_theory.limits.cocone.op CategoryTheory.Limits.Cocone.op /-- Change a `Cone F` into a `Cocone F.op`. -/ @[simps] def Cone.op (c : Cone F) : Cocone F.op where - X := Opposite.op c.X + pt := Opposite.op c.pt ι := NatTrans.op c.π #align category_theory.limits.cone.op CategoryTheory.Limits.Cone.op @@ -881,7 +881,7 @@ def Cone.op (c : Cone F) : Cocone F.op where @[simps] def Cocone.unop (c : Cocone F.op) : Cone F where - X := Opposite.unop c.X + pt := Opposite.unop c.pt π := NatTrans.removeOp c.ι #align category_theory.limits.cocone.unop CategoryTheory.Limits.Cocone.unop @@ -889,7 +889,7 @@ def Cocone.unop (c : Cocone F.op) : Cone F @[simps] def Cone.unop (c : Cone F.op) : Cocone F where - X := Opposite.unop c.X + pt := Opposite.unop c.pt ι := NatTrans.removeOp c.π #align category_theory.limits.cone.unop CategoryTheory.Limits.Cone.unop @@ -970,7 +970,7 @@ and replace with `@[simps]`-/ @[simps!] def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F where - X := op c.X + pt := op c.pt π := NatTrans.removeLeftOp c.ι #align category_theory.limits.cone_of_cocone_left_op CategoryTheory.Limits.coneOfCoconeLeftOp @@ -978,7 +978,7 @@ def coneOfCoconeLeftOp (c : Cocone F.leftOp) : Cone F @[simps!] def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp where - X := unop c.X + pt := unop c.pt ι := NatTrans.leftOp c.π #align category_theory.limits.cocone_left_op_of_cone CategoryTheory.Limits.coconeLeftOpOfCone @@ -989,7 +989,7 @@ def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp @[simps X] def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F where - X := op c.X + pt := op c.pt ι := NatTrans.removeLeftOp c.π #align category_theory.limits.cocone_of_cone_left_op CategoryTheory.Limits.coconeOfConeLeftOp @@ -1004,7 +1004,7 @@ theorem coconeOfConeLeftOp_ι_app (c : Cone F.leftOp) (j) : @[simps!] def coneLeftOpOfCocone (c : Cocone F) : Cone F.leftOp where - X := unop c.X + pt := unop c.pt π := NatTrans.leftOp c.ι #align category_theory.limits.cone_left_op_of_cocone CategoryTheory.Limits.coneLeftOpOfCocone @@ -1018,7 +1018,7 @@ variable {F : Jᵒᵖ ⥤ C} @[simps] def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F where - X := unop c.X + pt := unop c.pt π := NatTrans.removeRightOp c.ι #align category_theory.limits.cone_of_cocone_right_op CategoryTheory.Limits.coneOfCoconeRightOp @@ -1026,7 +1026,7 @@ def coneOfCoconeRightOp (c : Cocone F.rightOp) : Cone F @[simps] def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp where - X := op c.X + pt := op c.pt ι := NatTrans.rightOp c.π #align category_theory.limits.cocone_right_op_of_cone CategoryTheory.Limits.coconeRightOpOfCone @@ -1034,7 +1034,7 @@ def coconeRightOpOfCone (c : Cone F) : Cocone F.rightOp @[simps] def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F where - X := unop c.X + pt := unop c.pt ι := NatTrans.removeRightOp c.π #align category_theory.limits.cocone_of_cone_right_op CategoryTheory.Limits.coconeOfConeRightOp @@ -1042,7 +1042,7 @@ def coconeOfConeRightOp (c : Cone F.rightOp) : Cocone F @[simps] def coneRightOpOfCocone (c : Cocone F) : Cone F.rightOp where - X := op c.X + pt := op c.pt π := NatTrans.rightOp c.ι #align category_theory.limits.cone_right_op_of_cocone CategoryTheory.Limits.coneRightOpOfCocone @@ -1056,7 +1056,7 @@ variable {F : Jᵒᵖ ⥤ Cᵒᵖ} @[simps] def coneOfCoconeUnop (c : Cocone F.unop) : Cone F where - X := op c.X + pt := op c.pt π := NatTrans.removeUnop c.ι #align category_theory.limits.cone_of_cocone_unop CategoryTheory.Limits.coneOfCoconeUnop @@ -1064,7 +1064,7 @@ def coneOfCoconeUnop (c : Cocone F.unop) : Cone F @[simps] def coconeUnopOfCone (c : Cone F) : Cocone F.unop where - X := unop c.X + pt := unop c.pt ι := NatTrans.unop c.π #align category_theory.limits.cocone_unop_of_cone CategoryTheory.Limits.coconeUnopOfCone @@ -1072,7 +1072,7 @@ def coconeUnopOfCone (c : Cone F) : Cocone F.unop @[simps] def coconeOfConeUnop (c : Cone F.unop) : Cocone F where - X := op c.X + pt := op c.pt ι := NatTrans.removeUnop c.π #align category_theory.limits.cocone_of_cone_unop CategoryTheory.Limits.coconeOfConeUnop @@ -1080,7 +1080,7 @@ def coconeOfConeUnop (c : Cone F.unop) : Cocone F @[simps] def coneUnopOfCocone (c : Cocone F) : Cone F.unop where - X := unop c.X + pt := unop c.pt π := NatTrans.unop c.ι #align category_theory.limits.cone_unop_of_cocone CategoryTheory.Limits.coneUnopOfCocone From 0ae8a12093995bb64a01d9a27d07dc5421731932 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 11:17:15 -0500 Subject: [PATCH 123/126] fix import file --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 09ff72af6b009..b29c05af979ba 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -269,6 +269,7 @@ import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.IsomorphismClasses import Mathlib.CategoryTheory.LiftingProperties.Adjunction import Mathlib.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Limits.Cones import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi import Mathlib.CategoryTheory.Monoidal.Category import Mathlib.CategoryTheory.Monoidal.Functor From 7abe319c5caabec1753d2b8bcc61833c3efacd63 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 11:17:29 -0500 Subject: [PATCH 124/126] remove script --- 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 cf0ddc2cf2a060ff4b763118a33742f0f5fd5539 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 11:20:59 -0500 Subject: [PATCH 125/126] fix one error --- Mathlib/CategoryTheory/Limits/Cones.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index c24b968f77ed5..0bc31340aed2d 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -986,7 +986,7 @@ def coconeLeftOpOfCone (c : Cone F) : Cocone F.leftOp reduce the RHS using `expr.dsimp` and `expr.simp`, but for some reason the expression is not being simplified properly. -/ /-- Change a cone on `F.leftOp : Jᵒᵖ ⥤ C` to a cocone on `F : J ⥤ Cᵒᵖ`. -/ -@[simps X] +@[simps pt] def coconeOfConeLeftOp (c : Cone F.leftOp) : Cocone F where pt := op c.pt From a46ad8be64c1e7e39ff20c5c1c9a1a3f4116389a Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Fri, 24 Feb 2023 14:53:02 -0500 Subject: [PATCH 126/126] move comment --- Mathlib/CategoryTheory/Limits/Cones.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index 0bc31340aed2d..158e8eb08eb8a 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -701,7 +701,6 @@ variable {F : J ⥤ C} {G : J ⥤ C} (H : C ⥤ D) open CategoryTheory.Limits -/- Porting note: dot notation on the functor is broken for `mapCone` -/ /-- The image of a cone in C under a functor G : C ⥤ D is a cone in D. -/ @[simps!] def mapCone (c : Cone F) : Cone (F ⋙ H) := @@ -714,6 +713,7 @@ def mapCocone (c : Cocone F) : Cocone (F ⋙ H) := (Cocones.functoriality F H).obj c #align category_theory.functor.map_cocone CategoryTheory.Functor.mapCocone +/- Porting note: dot notation on the functor is broken for `mapCone` -/ /-- Given a cone morphism `c ⟶ c'`, construct a cone morphism on the mapped cones functorially. -/ def mapConeMorphism {c c' : Cone F} (f : c ⟶ c') : mapCone H c ⟶ mapCone _ c' := (Cones.functoriality F H).map f