Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit cf260ca

Browse files
kim-emjohoelzl
authored andcommitted
feat(category_theory/*): some lemmas about universes
1 parent 04c4abf commit cf260ca

File tree

3 files changed

+38
-3
lines changed

3 files changed

+38
-3
lines changed

category_theory/category.lean

Lines changed: 8 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -118,12 +118,17 @@ end
118118

119119
section
120120
variable (C : Type u)
121-
variable [small_category C]
121+
variable [category.{u v} C]
122122

123-
instance : large_category (ulift.{(u+1)} C) :=
123+
universe u'
124+
125+
instance ulift_category : category.{(max u u') v} (ulift.{u'} C) :=
124126
{ hom := λ X Y, (X.down ⟶ Y.down),
125127
id := λ X, 𝟙 X.down,
126-
comp := λ _ _ _ f g, f ≫ g }
128+
comp := λ _ _ _ f g, f ≫ g }
129+
130+
-- We verify that this previous instance can lift small categories to large categories.
131+
example (D : Type u) [small_category D] : large_category (ulift.{u+1} D) := by apply_instance
127132
end
128133

129134
variables (α : Type u)

category_theory/functor.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -103,6 +103,19 @@ infixr ` ⋙ `:80 := comp
103103
(F ⋙ G).map f = G.map (F.map f) := rfl
104104
end
105105

106+
section
107+
variables (C : Type u₁) [𝒞 : category.{u₁ v₁} C]
108+
include 𝒞
109+
110+
@[simp] def ulift_down : (ulift.{u₂} C) ⥤ C :=
111+
{ obj := λ X, X.down,
112+
map' := λ X Y f, f }
113+
114+
@[simp] def ulift_up : C ⥤ (ulift.{u₂} C) :=
115+
{ obj := λ X, ⟨ X ⟩,
116+
map' := λ X Y f, f }
117+
end
118+
106119
end functor
107120

108121
def bundled.map {c : Type u → Type v} {d : Type u → Type v} (f : Π{a}, c a → d a) (s : bundled c) : bundled d :=

category_theory/natural_isomorphism.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -87,6 +87,7 @@ namespace category_theory.functor
8787

8888
universes u₁ u₂ v₁ v₂
8989

90+
section
9091
variables {C : Type u₁} [𝒞 : category.{u₁ v₁} C]
9192
{D : Type u₂} [𝒟 : category.{u₂ v₂} D]
9293
include 𝒞 𝒟
@@ -121,4 +122,20 @@ variables (F : A ⥤ B) (G : B ⥤ C) (H : C ⥤ D)
121122
-- When it's time to define monoidal categories and 2-categories,
122123
-- we'll need to add lemmas relating these natural isomorphisms,
123124
-- in particular the pentagon for the associator.
125+
end
126+
127+
section
128+
variables {C : Type u₁} [𝒞 : category.{u₁ v₁} C]
129+
include 𝒞
130+
131+
def ulift_down_up : ulift_down.{u₁ v₁ u₂} C ⋙ ulift_up C ≅ functor.id (ulift.{u₂} C) :=
132+
{ hom := { app := λ X, @category.id (ulift.{u₂} C) _ X },
133+
inv := { app := λ X, @category.id (ulift.{u₂} C) _ X } }
134+
135+
def ulift_up_down : ulift_up.{u₁ v₁ u₂} C ⋙ ulift_down C ≅ functor.id C :=
136+
{ hom := { app := λ X, 𝟙 X },
137+
inv := { app := λ X, 𝟙 X } }
138+
139+
end
140+
124141
end category_theory.functor

0 commit comments

Comments
 (0)