1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79 80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 96 97 98 99 100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119 120 121 122 123 124 125 126 127 128 129 130 131 132 133 134 135 136 137 138 139 140 141 142 143 144 145 146 147 148 149 150 151 152 153 154 155 156 157 158 159 160 161 162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179 180 181 182 183 184 185 186 187 188 189 190 191 192 193 194 195 196 197 198 199 200 201 202 203 204 205 206 207 208 209 210 211 212 213 214 215 216 217 218 219 220 221 222 223 224 225 226 227 228 229 230 231 232 233 234 235 236 237 238 239 240 241 242
|
Require Import TestSuite.admit.
Set Implicit Arguments.
Generalizable All Variables.
Set Asymmetric Patterns.
Set Universe Polymorphism.
Delimit Scope object_scope with object.
Delimit Scope morphism_scope with morphism.
Delimit Scope category_scope with category.
Delimit Scope functor_scope with functor.
Local Open Scope category_scope.
Record SpecializedCategory (obj : Type) :=
{
Object :> _ := obj;
Morphism : obj -> obj -> Type;
Compose : forall s d d', Morphism d d' -> Morphism s d -> Morphism s d'
}.
Bind Scope category_scope with SpecializedCategory.
Bind Scope object_scope with Object.
Bind Scope morphism_scope with Morphism.
Arguments Object {obj%_type} C%_category / : rename.
Arguments Morphism {obj%_type} !C%_category s d : rename. (* , simpl nomatch. *)
Arguments Compose {obj%_type} [!C%_category s%_object d%_object d'%_object] m1%_morphism m2%_morphism : rename.
Record Category := {
CObject : Type;
UnderlyingCategory :> @SpecializedCategory CObject
}.
Definition GeneralizeCategory `(C : @SpecializedCategory obj) : Category.
refine {| CObject := C.(Object) |}; auto. (* Changing this [auto] to [assumption] removes the universe inconsistency. *)
Defined.
Coercion GeneralizeCategory : SpecializedCategory >-> Category.
Record SpecializedFunctor
`(C : @SpecializedCategory objC)
`(D : @SpecializedCategory objD)
:= {
ObjectOf :> objC -> objD;
MorphismOf : forall s d, C.(Morphism) s d -> D.(Morphism) (ObjectOf s) (ObjectOf d)
}.
Section Functor.
Context (C D : Category).
Definition Functor := SpecializedFunctor C D.
End Functor.
Bind Scope functor_scope with SpecializedFunctor.
Bind Scope functor_scope with Functor.
Arguments SpecializedFunctor {objC} C {objD} D.
Arguments Functor C D.
Arguments ObjectOf {objC%_type C%_category objD%_type D%_category} F%_functor c%_object : rename, simpl nomatch.
Arguments MorphismOf {objC%_type} [C%_category] {objD%_type} [D%_category] F%_functor [s%_object d%_object] m%_morphism : rename, simpl nomatch.
Section FunctorComposition.
Context `(B : @SpecializedCategory objB).
Context `(C : @SpecializedCategory objC).
Context `(D : @SpecializedCategory objD).
Context `(E : @SpecializedCategory objE).
Definition ComposeFunctors (G : SpecializedFunctor D E) (F : SpecializedFunctor C D) : SpecializedFunctor C E
:= Build_SpecializedFunctor C E
(fun c => G (F c))
(fun _ _ m => G.(MorphismOf) (F.(MorphismOf) m)).
End FunctorComposition.
Record SpecializedNaturalTransformation
`{C : @SpecializedCategory objC}
`{D : @SpecializedCategory objD}
(F G : SpecializedFunctor C D)
:= {
ComponentsOf :> forall c, D.(Morphism) (F c) (G c)
}.
Definition ProductCategory
`(C : @SpecializedCategory objC)
`(D : @SpecializedCategory objD)
: @SpecializedCategory (objC * objD)%type
:= @Build_SpecializedCategory _
(fun s d => (C.(Morphism) (fst s) (fst d) * D.(Morphism) (snd s) (snd d))%type)
(fun s d d' m2 m1 => (Compose (fst m2) (fst m1), Compose (snd m2) (snd m1))).
Infix "*" := ProductCategory : category_scope.
Section ProductCategoryFunctors.
Context `{C : @SpecializedCategory objC}.
Context `{D : @SpecializedCategory objD}.
Definition fst_Functor : SpecializedFunctor (C * D) C
:= Build_SpecializedFunctor (C * D) C
(@fst _ _)
(fun _ _ => @fst _ _).
Definition snd_Functor : SpecializedFunctor (C * D) D
:= Build_SpecializedFunctor (C * D) D
(@snd _ _)
(fun _ _ => @snd _ _).
End ProductCategoryFunctors.
Definition OppositeCategory `(C : @SpecializedCategory objC) : @SpecializedCategory objC :=
@Build_SpecializedCategory objC
(fun s d => Morphism C d s)
(fun _ _ _ m1 m2 => Compose m2 m1).
Section OppositeFunctor.
Context `(C : @SpecializedCategory objC).
Context `(D : @SpecializedCategory objD).
Variable F : SpecializedFunctor C D.
Let COp := OppositeCategory C.
Let DOp := OppositeCategory D.
Definition OppositeFunctor : SpecializedFunctor COp DOp
:= Build_SpecializedFunctor COp DOp
(fun c : COp => F c : DOp)
(fun (s d : COp) (m : C.(Morphism) d s) => MorphismOf F (s := d) (d := s) m).
End OppositeFunctor.
Section FunctorProduct.
Context `(C : @SpecializedCategory objC).
Context `(D : @SpecializedCategory objD).
Context `(D' : @SpecializedCategory objD').
Definition FunctorProduct (F : Functor C D) (F' : Functor C D') : SpecializedFunctor C (D * D').
match goal with
| [ |- SpecializedFunctor ?C0 ?D0 ] =>
refine (Build_SpecializedFunctor
C0 D0
(fun c => (F c, F' c))
(fun s d m => (MorphismOf F m, MorphismOf F' m)))
end.
Defined.
End FunctorProduct.
Section FunctorProduct'.
Context `(C : @SpecializedCategory objC).
Context `(D : @SpecializedCategory objD).
Context `(C' : @SpecializedCategory objC').
Context `(D' : @SpecializedCategory objD').
Variable F : Functor C D.
Variable F' : Functor C' D'.
Definition FunctorProduct' : SpecializedFunctor (C * C') (D * D')
:= FunctorProduct (ComposeFunctors F fst_Functor) (ComposeFunctors F' snd_Functor).
End FunctorProduct'.
(** XXX TODO(jgross): Change this to [FunctorProduct]. *)
Infix "*" := FunctorProduct' : functor_scope.
Definition TypeCat : @SpecializedCategory Type :=
(@Build_SpecializedCategory Type
(fun s d => s -> d)
(fun _ _ _ f g => (fun x => f (g x)))).
Section HomFunctor.
Context `(C : @SpecializedCategory objC).
Let COp := OppositeCategory C.
Definition CovariantHomFunctor (A : COp) : SpecializedFunctor C TypeCat
:= Build_SpecializedFunctor C TypeCat
(fun X : C => C.(Morphism) A X : TypeCat)
(fun X Y f => (fun g : C.(Morphism) A X => Compose f g)).
Definition hom_functor_object_of (c'c : COp * C) := C.(Morphism) (fst c'c) (snd c'c) : TypeCat.
Definition hom_functor_morphism_of (s's : (COp * C)%type) (d'd : (COp * C)%type) (hf : (COp * C).(Morphism) s's d'd) :
TypeCat.(Morphism) (hom_functor_object_of s's) (hom_functor_object_of d'd).
unfold hom_functor_object_of in *.
destruct s's as [ s' s ], d'd as [ d' d ].
destruct hf as [ h f ].
intro g.
exact (Compose f (Compose g h)).
Defined.
Definition HomFunctor : SpecializedFunctor (COp * C) TypeCat
:= Build_SpecializedFunctor (COp * C) TypeCat
(fun c'c : COp * C => C.(Morphism) (fst c'c) (snd c'c) : TypeCat)
(fun X Y (hf : (COp * C).(Morphism) X Y) => hom_functor_morphism_of hf).
End HomFunctor.
Section FullFaithful.
Context `(C : @SpecializedCategory objC).
Context `(D : @SpecializedCategory objD).
Variable F : SpecializedFunctor C D.
Let COp := OppositeCategory C.
Let DOp := OppositeCategory D.
Let FOp := OppositeFunctor F.
Definition InducedHomNaturalTransformation :
SpecializedNaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F))
:= (Build_SpecializedNaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F))
(fun sd : (COp * C) =>
MorphismOf F (s := _) (d := _))).
End FullFaithful.
Definition FunctorCategory
`(C : @SpecializedCategory objC)
`(D : @SpecializedCategory objD)
: @SpecializedCategory (SpecializedFunctor C D).
refine (@Build_SpecializedCategory _
(SpecializedNaturalTransformation (C := C) (D := D))
_);
admit.
Defined.
Notation "C ^ D" := (FunctorCategory D C) : category_scope.
Section Yoneda.
Context `(C : @SpecializedCategory objC).
Let COp := OppositeCategory C.
Section Yoneda.
Let Yoneda_NT s d (f : COp.(Morphism) s d) : SpecializedNaturalTransformation (CovariantHomFunctor C s) (CovariantHomFunctor C d)
:= Build_SpecializedNaturalTransformation
(CovariantHomFunctor C s)
(CovariantHomFunctor C d)
(fun c : C => (fun m : C.(Morphism) _ _ => Compose m f)).
Definition Yoneda : SpecializedFunctor COp (TypeCat ^ C)
:= Build_SpecializedFunctor COp (TypeCat ^ C)
(fun c : COp => CovariantHomFunctor C c : TypeCat ^ C)
(fun s d (f : Morphism COp s d) => Yoneda_NT s d f).
End Yoneda.
End Yoneda.
Section FullyFaithful.
Context `(C : @SpecializedCategory objC).
Set Printing Universes.
Check InducedHomNaturalTransformation (Yoneda C).
(* Error: Universe inconsistency (cannot enforce Top.865 = Top.851 because
Top.851 < Top.869 <= Top.864 <= Top.865). *)
End FullyFaithful.
|