Skip to content

Commit

Permalink
refactor(category_theory/functor): a folder for concepts directly rel…
Browse files Browse the repository at this point in the history
…ated to functors (#12346)
  • Loading branch information
TwoFX committed Mar 1, 2022
1 parent 456898c commit 73dd4b5
Show file tree
Hide file tree
Showing 40 changed files with 35 additions and 36 deletions.
2 changes: 1 addition & 1 deletion docs/tutorial/category_theory/intro.lean
@@ -1,4 +1,4 @@
import category_theory.functor_category -- this transitively imports
import category_theory.functor.category -- this transitively imports
-- category_theory.category
-- category_theory.functor
-- category_theory.natural_transformation
Expand Down
2 changes: 1 addition & 1 deletion src/algebra/category/Mon/basic.lean
Expand Up @@ -6,7 +6,7 @@ Authors: Scott Morrison
import tactic.elementwise
import category_theory.concrete_category.bundled_hom
import algebra.punit_instances
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms

/-!
# Category instances for monoid, add_monoid, comm_monoid, and add_comm_monoid.
Expand Down
2 changes: 1 addition & 1 deletion src/algebra/category/Semigroup/basic.lean
Expand Up @@ -5,7 +5,7 @@ Authors: Julian Kuelshammer
-/
import algebra.pempty_instances
import category_theory.concrete_category.bundled_hom
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms

/-!
# Category instances for has_mul, has_add, semigroup and add_semigroup
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/abelian/ext.lean
Expand Up @@ -5,7 +5,7 @@ Authors: Scott Morrison, Adam Topaz
-/
import algebra.category.Group.basic
import algebra.category.Module.abelian
import category_theory.derived
import category_theory.functor.derived
import category_theory.linear.yoneda
import category_theory.abelian.opposite
import category_theory.abelian.projective
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/adjunction/reflective.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta
-/
import category_theory.adjunction.fully_faithful
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms
import category_theory.epi_mono

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/comma.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison, Johan Commelin, Bhavik Mehta
-/
import category_theory.isomorphism
import category_theory.functor_category
import category_theory.functor.category
import category_theory.eq_to_hom

/-!
Expand Down
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.concrete_category.basic
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms

/-!
A `forget₂ C D` forgetful functor between concrete categories `C` and `D`
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/equivalence.lean
Expand Up @@ -3,7 +3,7 @@ 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
-/
import category_theory.fully_faithful
import category_theory.functor.fully_faithful
import category_theory.full_subcategory
import category_theory.whiskering
import category_theory.essential_image
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/full_subcategory.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2017 Scott Morrison. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison, Reid Barton
-/
import category_theory.fully_faithful
import category_theory.functor.fully_faithful

/-!
# Induced categories and full subcategories
Expand Down
File renamed without changes.
File renamed without changes.
File renamed without changes.
File renamed without changes.
1 change: 1 addition & 0 deletions src/category_theory/functor/default.lean
@@ -0,0 +1 @@
import category_theory.functor.basic
File renamed without changes.
File renamed without changes.
File renamed without changes.
File renamed without changes.
File renamed without changes.
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2020 Bhavik Mehta. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta
-/
import category_theory.fully_faithful
import category_theory.functor.fully_faithful

/-!
# Functors which reflect isomorphisms
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/limits/colimit_limit.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.limits.types
import category_theory.currying
import category_theory.functor.currying
import category_theory.limits.functor_category

/-!
Expand Down
4 changes: 2 additions & 2 deletions src/category_theory/limits/cones.lean
Expand Up @@ -3,10 +3,10 @@ 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
-/
import category_theory.const
import category_theory.functor.const
import category_theory.discrete_category
import category_theory.yoneda
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms

-- morphism levels before object levels. See note [category_theory universes].
universes v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/limits/fubini.lean
Expand Up @@ -5,7 +5,7 @@ Authors: Scott Morrison
-/
import category_theory.limits.has_limits
import category_theory.products.basic
import category_theory.currying
import category_theory.functor.currying

/-!
# A Fubini theorem for categorical limits
Expand Down
1 change: 0 additions & 1 deletion src/category_theory/limits/is_limit.lean
Expand Up @@ -5,7 +5,6 @@ Authors: Reid Barton, Mario Carneiro, Scott Morrison, Floris van Doorn
-/
import category_theory.adjunction.basic
import category_theory.limits.cones
import category_theory.reflects_isomorphisms

/-!
# Limits and colimits
Expand Down
1 change: 0 additions & 1 deletion src/category_theory/monad/algebra.lean
Expand Up @@ -5,7 +5,6 @@ Authors: Scott Morrison, Bhavik Mehta
-/
import category_theory.monad.basic
import category_theory.adjunction.basic
import category_theory.reflects_isomorphisms

/-!
# Eilenberg-Moore (co)algebras for a (co)monad
Expand Down
6 changes: 3 additions & 3 deletions src/category_theory/monad/basic.lean
Expand Up @@ -3,9 +3,9 @@ Copyright (c) 2019 Scott Morrison. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison, Bhavik Mehta, Adam Topaz
-/
import category_theory.functor_category
import category_theory.fully_faithful
import category_theory.reflects_isomorphisms
import category_theory.functor.category
import category_theory.functor.fully_faithful
import category_theory.functor.reflects_isomorphisms

namespace category_theory
open category
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/monoidal/center.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.monoidal.braided
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms

/-!
# Half braidings and the Drinfeld center of a monoidal category
Expand Down
4 changes: 2 additions & 2 deletions src/category_theory/monoidal/functor_category.lean
Expand Up @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.monoidal.braided
import category_theory.functor_category
import category_theory.const
import category_theory.functor.category
import category_theory.functor.const

/-!
# Monoidal structure on `C ⥤ D` when `D` is monoidal.
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/monoidal/functorial.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.monoidal.functor
import category_theory.functorial
import category_theory.functor.functorial

/-!
# Unbundled lax monoidal functors
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/monoidal/tor.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.derived
import category_theory.functor.derived
import category_theory.monoidal.preadditive

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/natural_isomorphism.lean
Expand Up @@ -3,7 +3,7 @@ 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
-/
import category_theory.functor_category
import category_theory.functor.category
import category_theory.isomorphism

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/over.lean
Expand Up @@ -5,7 +5,7 @@ Authors: Johan Commelin, Bhavik Mehta
-/
import category_theory.structured_arrow
import category_theory.punit
import category_theory.reflects_isomorphisms
import category_theory.functor.reflects_isomorphisms
import category_theory.epi_mono

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/punit.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2018 Scott Morrison. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison, Bhavik Mehta
-/
import category_theory.const
import category_theory.functor.const
import category_theory.discrete_category

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/sigma/basic.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta
-/
import category_theory.whiskering
import category_theory.fully_faithful
import category_theory.functor.fully_faithful
import category_theory.natural_isomorphism
import data.sigma.basic

Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/sites/cover_preserving.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Andrew Yang
-/
import category_theory.sites.limits
import category_theory.flat_functors
import category_theory.functor.flat
import category_theory.limits.preserves.filtered
import category_theory.sites.left_exact

Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/subobject/mono_over.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2020 Bhavik Mehta. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta, Scott Morrison
-/
import category_theory.currying
import category_theory.functor.currying
import category_theory.limits.over
import category_theory.limits.shapes.images
import category_theory.adjunction.reflective
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/thin.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison, Bhavik Mehta. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison, Bhavik Mehta
-/
import category_theory.functor_category
import category_theory.functor.category
import category_theory.isomorphism

/-!
Expand Down
2 changes: 1 addition & 1 deletion src/category_theory/types.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Stephen Morgan, Scott Morrison, Johannes Hölzl
-/
import category_theory.epi_mono
import category_theory.fully_faithful
import category_theory.functor.fully_faithful
import data.equiv.basic

/-!
Expand Down
4 changes: 2 additions & 2 deletions src/category_theory/whiskering.lean
Expand Up @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.isomorphism
import category_theory.functor_category
import category_theory.fully_faithful
import category_theory.functor.category
import category_theory.functor.fully_faithful

/-!
# Whiskering
Expand Down
4 changes: 2 additions & 2 deletions src/category_theory/yoneda.lean
Expand Up @@ -3,8 +3,8 @@ Copyright (c) 2017 Scott Morrison. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import category_theory.hom_functor
import category_theory.currying
import category_theory.functor.hom
import category_theory.functor.currying
import category_theory.products.basic

/-!
Expand Down

0 comments on commit 73dd4b5

Please sign in to comment.