diff -r 4eb6407c6ca3 Adjoint.v
--- a/Adjoint.v	Fri May 24 20:09:37 2013 -0400
+++ b/Adjoint.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export SpecializedCategory Category Functor NaturalTransformation NaturalEquivalence AdjointUnit.
+Require Export Category Category Functor NaturalTransformation NaturalEquivalence AdjointUnit.
 Require Import Common Hom ProductCategory FunctorProduct Duals.
 
 Set Implicit Arguments.
@@ -11,17 +11,17 @@
 Set Universe Polymorphism.
 
 Section Adjunction.
-  Context `{C : @SpecializedCategory objC}.
-  Context `{D : @SpecializedCategory objD}.
-  Variable F : SpecializedFunctor C D.
-  Variable G : SpecializedFunctor D C.
+  Context {C : Category}.
+  Context {D : Category}.
+  Variable F : Functor C D.
+  Variable G : Functor D C.
 
   Let COp := OppositeCategory C.
   Let DOp := OppositeCategory D.
   Let FOp := OppositeFunctor F.
 
-  Let HomCPreFunctor : SpecializedFunctor (COp * D) (COp * C) := ((IdentityFunctor _) * G)%functor.
-  Let HomDPreFunctor : SpecializedFunctor (COp * D) (DOp * D) := (FOp * (IdentityFunctor _))%functor.
+  Let HomCPreFunctor : Functor (COp * D) (COp * C) := ((IdentityFunctor _) * G)%functor.
+  Let HomDPreFunctor : Functor (COp * D) (DOp * D) := (FOp * (IdentityFunctor _))%functor.
 
   Record Adjunction := {
     AMateOf :> NaturalIsomorphism
@@ -120,20 +120,20 @@
 Arguments AIsomorphism {objC C objD D} [F G] T A A' : simpl nomatch.
 
 Section AdjunctionEquivalences.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
-  Variable G : SpecializedFunctor D C.
+  Variable C : Category.
+  Variable D : Category.
+  Variable F : Functor C D.
+  Variable G : Functor D C.
 
   Definition HomAdjunction2Adjunction_AMateOf (A : HomAdjunction F G) :
-    SpecializedNaturalTransformation
+    NaturalTransformation
     (ComposeFunctors (HomFunctor D)
       (OppositeFunctor F * IdentityFunctor D))
     (ComposeFunctors (HomFunctor C)
       (IdentityFunctor (OppositeCategory C) * G)).
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun cd : objC * objD => A.(AComponentsOf) (fst cd) (snd cd))
           _
         )
@@ -209,7 +209,7 @@
      [η c = ϕ(1_{F c})]
      *)
   Definition UnitOf (A : HomAdjunction F G) : AdjunctionUnit F G.
-    eexists (Build_SpecializedNaturalTransformation (IdentityFunctor C) (ComposeFunctors G F)
+    eexists (Build_NaturalTransformation (IdentityFunctor C) (ComposeFunctors G F)
                                                     (fun c => A.(AComponentsOf) c (F c) (Identity _))
                                                     _
             ).
@@ -239,7 +239,7 @@
 
 
   Definition CounitOf (A : HomAdjunction F G) : AdjunctionCounit F G.
-    eexists (Build_SpecializedNaturalTransformation (ComposeFunctors F G) (IdentityFunctor D)
+    eexists (Build_NaturalTransformation (ComposeFunctors F G) (IdentityFunctor D)
       (fun d => proj1_sig (A.(AIsomorphism) (G d) d) (Identity _))
       _
     ).
diff -r 4eb6407c6ca3 AdjointComposition.v
--- a/AdjointComposition.v	Fri May 24 20:09:37 2013 -0400
+++ b/AdjointComposition.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,14 +10,14 @@
 Set Universe Polymorphism.
 
 Section compose.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Context `(E : SpecializedCategory objE).
+  Variable C : Category.
+  Variable D : Category.
+  Variable E : Category.
 
-  Variable F : SpecializedFunctor C D.
-  Variable F' : SpecializedFunctor D E.
-  Variable G : SpecializedFunctor D C.
-  Variable G' : SpecializedFunctor E D.
+  Variable F : Functor C D.
+  Variable F' : Functor D E.
+  Variable G : Functor D C.
+  Variable G' : Functor E D.
 
   Variable A' : Adjunction F' G'.
   Variable A : Adjunction F G.
@@ -33,8 +33,8 @@
                                                η)));
       nt_solve_associator.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
                                                        (fun _ => Identity _)
                                                        _)
     end.
diff -r 4eb6407c6ca3 AdjointPointwise.v
--- a/AdjointPointwise.v	Fri May 24 20:09:37 2013 -0400
+++ b/AdjointPointwise.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,23 +10,23 @@
 Set Universe Polymorphism.
 
 Section AdjointPointwise.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
-  Context `(C' : @SpecializedCategory objC').
-  Context `(D' : @SpecializedCategory objD').
+  Variable C : Category.
+  Variable D : Category.
+  Variable E : Category.
+  Context `(C' : @Category objC').
+  Context `(D' : @Category objD').
 
-  Variable F : SpecializedFunctor C D.
-  Variable G : SpecializedFunctor D C.
+  Variable F : Functor C D.
+  Variable G : Functor D C.
 
   Variable A : Adjunction F G.
   Let A' : AdjunctionUnitCounit F G := A.
 
-  Variables F' G' : SpecializedFunctor C' D'.
+  Variables F' G' : Functor C' D'.
 
-(*  Variable T' : SpecializedNaturalTransformation F' G'.*)
+(*  Variable T' : NaturalTransformation F' G'.*)
 
-  Definition AdjointPointwise_NT_Unit : SpecializedNaturalTransformation (IdentityFunctor (C ^ E))
+  Definition AdjointPointwise_NT_Unit : NaturalTransformation (IdentityFunctor (C ^ E))
                                                                     (ComposeFunctors (G ^ IdentityFunctor E) (F ^ IdentityFunctor E)).
     clearbody A'; clear A.
     pose proof (A' : AdjunctionUnit _ _) as A''.
@@ -37,7 +37,7 @@
     exact (LeftIdentityFunctorNaturalTransformation2 _).
   Defined.
 
-  Definition AdjointPointwise_NT_Counit : SpecializedNaturalTransformation (ComposeFunctors (F ^ IdentityFunctor E) (G ^ IdentityFunctor E))
+  Definition AdjointPointwise_NT_Counit : NaturalTransformation (ComposeFunctors (F ^ IdentityFunctor E) (G ^ IdentityFunctor E))
                                                                            (IdentityFunctor (D ^ E)).
     clearbody A'; clear A.
     pose proof (A' : AdjunctionCounit _ _) as A''.
diff -r 4eb6407c6ca3 AdjointUniversalMorphisms.v
--- a/AdjointUniversalMorphisms.v	Fri May 24 20:09:37 2013 -0400
+++ b/AdjointUniversalMorphisms.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,10 +10,10 @@
 Set Universe Polymorphism.
 
 Section AdjunctionUniversal.
-  Context `{C : @SpecializedCategory objC}.
-  Context `{D : @SpecializedCategory objD}.
-  Variable F : SpecializedFunctor C D.
-  Variable G : SpecializedFunctor D C.
+  Context {C : Category}.
+  Context {D : Category}.
+  Variable F : Functor C D.
+  Variable G : Functor D C.
 
   Definition InitialMorphismOfAdjunction (A : Adjunction F G) Y : InitialMorphism Y G.
     pose (projT1 (A : AdjunctionCounit F G)) as ε.
@@ -80,8 +80,8 @@
 End AdjunctionUniversal.
 
 Section AdjunctionFromUniversal.
-  Context `{C : @SpecializedCategory objC}.
-  Context `{D : @SpecializedCategory objD}.
+  Context {C : Category}.
+  Context {D : Category}.
 
   Local Ltac diagonal_transitivity_pre_then tac :=
     repeat rewrite AssociativityNoEvar by noEvar;
@@ -120,11 +120,11 @@
     diagonal_transitivity_then_solve_rewrite.
 
   Section initial.
-    Variable G : SpecializedFunctor D C.
+    Variable G : Functor D C.
     Variable M : forall Y, InitialMorphism Y G.
 
-    Definition AdjointFunctorOfInitialMorphism : SpecializedFunctor C D.
-      refine (Build_SpecializedFunctor C D
+    Definition AdjointFunctorOfInitialMorphism : Functor C D.
+      refine (Build_Functor C D
                                        (fun Y => let ηY := InitialMorphism_Morphism (M Y) in
                                                  let F0Y := InitialMorphism_Object (M Y) in
                                                  F0Y)
@@ -138,7 +138,7 @@
 
     Definition AdjunctionOfInitialMorphism : Adjunction AdjointFunctorOfInitialMorphism G.
       refine (_ : AdjunctionUnit AdjointFunctorOfInitialMorphism G).
-      exists (Build_SpecializedNaturalTransformation (IdentityFunctor C)
+      exists (Build_NaturalTransformation (IdentityFunctor C)
                                                      (ComposeFunctors G AdjointFunctorOfInitialMorphism)
                                                      (fun x => InitialMorphism_Morphism (M x))
                                                      (fun s d m => eq_sym (proj1 (InitialProperty (M s) _ _)))).
@@ -148,11 +148,11 @@
   End initial.
 
   Section terminal.
-    Variable F : SpecializedFunctor C D.
+    Variable F : Functor C D.
     Variable M : forall X, TerminalMorphism F X.
 
-    Definition AdjointFunctorOfTerminalMorphism : SpecializedFunctor D C.
-      refine (Build_SpecializedFunctor D C
+    Definition AdjointFunctorOfTerminalMorphism : Functor D C.
+      refine (Build_Functor D C
                                        (fun X => let εX := TerminalMorphism_Morphism (M X) in
                                                  let G0X := TerminalMorphism_Object (M X) in
                                                  G0X)
@@ -168,7 +168,7 @@
     Definition AdjunctionOfTerminalMorphism : Adjunction F AdjointFunctorOfTerminalMorphism.
       refine (_ : AdjunctionCounit F AdjointFunctorOfTerminalMorphism).
       hnf.
-      exists (Build_SpecializedNaturalTransformation (ComposeFunctors F AdjointFunctorOfTerminalMorphism)
+      exists (Build_NaturalTransformation (ComposeFunctors F AdjointFunctorOfTerminalMorphism)
                                                      (IdentityFunctor D)
                                                      (fun x => TerminalMorphism_Morphism (M x))
                                                      (fun s d m => proj1 (TerminalProperty (M d) _ _))).
diff -r 4eb6407c6ca3 BoolCategory.v
--- a/BoolCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/BoolCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import Bool.
 
 Set Implicit Arguments.
@@ -29,8 +29,8 @@
   Global Arguments BoolCat_Compose [s d d'] m1 m2 : simpl never.
   Global Arguments BoolCat_Identity x : simpl never.
 
-  Definition BoolCat : @SpecializedCategory bool.
-    refine (@Build_SpecializedCategory _
+  Definition BoolCat : @Category bool.
+    refine (@Build_Category _
                                        mor
                                        BoolCat_Identity
                                        BoolCat_Compose
diff -r 4eb6407c6ca3 CanonicalStructureSimplification.v
--- a/CanonicalStructureSimplification.v	Fri May 24 20:09:37 2013 -0400
+++ b/CanonicalStructureSimplification.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Export LtacReifiedSimplification.
-Require Import SpecializedCategory Functor NaturalTransformation.
+Require Import Category Functor NaturalTransformation.
 
 Set Implicit Arguments.
 
@@ -9,9 +9,11 @@
 
 Set Universe Polymorphism.
 
+Local Open Scope morphism_scope.
+
 Section SimplifiedMorphism.
   Section single_category.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     (* structure for packaging a morphism and its reification *)
 
@@ -52,13 +54,13 @@
                  reflexivity.
 
   Section more_single_category.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     Lemma reifyIdentity x : Identity x = ReifiedMorphismDenote (ReifiedIdentityMorphism C x). reflexivity. Qed.
     Canonical Structure reify_identity_morphism x := ReifyMorphism (identity_tag _) _ (reifyIdentity x).
 
     Lemma reifyCompose s d d'
-          `(m1' : @SimplifiedMorphism objC C d d') `(m2' : @SimplifiedMorphism objC C s d)
+          `(m1' : @SimplifiedMorphism C d d') `(m2' : @SimplifiedMorphism C s d)
     : Compose (untag (morphism_of m1')) (untag (morphism_of m2'))
       = ReifiedMorphismDenote (ReifiedComposedMorphism (reified_morphism_of m1') (reified_morphism_of m2')).
       t.
@@ -67,11 +69,11 @@
   End more_single_category.
 
   Section functor.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
-    Variable F : SpecializedFunctor C D.
+    Context {C : Category}.
+    Context {D : Category}.
+    Variable F : Functor C D.
 
-    Lemma reifyFunctor `(m' : @SimplifiedMorphism objC C s d)
+    Lemma reifyFunctor `(m' : @SimplifiedMorphism C s d)
     : MorphismOf F (untag (morphism_of m')) = ReifiedMorphismDenote (ReifiedFunctorMorphism F (reified_morphism_of m')).
       t.
     Qed.
@@ -79,16 +81,15 @@
   End functor.
 
   Section natural_transformation.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
-    Variables F G : SpecializedFunctor C D.
-    Variable T : SpecializedNaturalTransformation F G.
+    Context {C D : Category}.
+    Variables F G : Functor C D.
+    Variable T : NaturalTransformation F G.
 
     Lemma reifyNT (x : C) : T x = ReifiedMorphismDenote (ReifiedNaturalTransformationMorphism T x). reflexivity. Qed.
     Canonical Structure reify_nt_morphism x := ReifyMorphism (nt_tag _) _ (@reifyNT x).
   End natural_transformation.
   Section generic.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     Lemma reifyGeneric s d (m : Morphism C s d) : m = ReifiedMorphismDenote (ReifiedGenericMorphism C s d m). reflexivity. Qed.
     Canonical Structure reify_generic_morphism s d m := ReifyMorphism (generic_tag m) _ (@reifyGeneric s d m).
@@ -113,24 +114,23 @@
 (*******************************************************************************)
 Section good_examples.
   Section id.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objC).
-    Variables F G : SpecializedFunctor C D.
-    Variable T : SpecializedNaturalTransformation F G.
+    Variables C D : Category.
+    Variables F G : Functor C D.
+    Variable T : NaturalTransformation F G.
 
-    Lemma good_example_00001 (x : C) :Compose (Identity x) (Identity x) = Identity (C := C) x.
+    Lemma good_example_00001 (x : C) : Identity x ∘ Identity x = Identity (C := C) x.
       rsimplify_morphisms.
       reflexivity.
     Qed.
 
     Lemma good_example_00002 s d (m : Morphism C s d)
-    : MorphismOf F (Compose m (Identity s)) = MorphismOf F m.
+    : MorphismOf F (m ∘ Identity s) = MorphismOf F m.
       rsimplify_morphisms.
       reflexivity.
     Qed.
 
     Lemma good_example_00003 s d (m : Morphism C s d)
-    : MorphismOf F (Compose (Identity d) m) = MorphismOf F m.
+    : MorphismOf F (Identity d ∘ m) = MorphismOf F m.
       rsimplify_morphisms.
       reflexivity.
     Qed.
@@ -143,16 +143,15 @@
 Section bad_examples.
   Require Import SumCategory.
   Section bad_example_0001.
-    Context `(C0 : SpecializedCategory objC0).
-    Context `(C1 : SpecializedCategory objC1).
-    Context `(D : SpecializedCategory objD).
+    Variables C0 C1 : Category.
+    Variable D : Category.
 
     Variables s d d' : C0.
     Variable m1 : Morphism C0 s d.
     Variable m2 : Morphism C0 d d'.
-    Variable F : SpecializedFunctor (C0 + C1) D.
+    Variable F : Functor (C0 + C1) D.
 
-    Goal MorphismOf F (s := inl _) (d := inl _) (Compose m2 m1) = Compose (MorphismOf F (s := inl _) (d := inl _) m2) (MorphismOf F (s := inl _) (d := inl _) m1).
+    Goal MorphismOf F (s := inl _) (d := inl _) (m2 ∘ m1) = (MorphismOf F (s := inl _) (d := inl _) m2) ∘ (MorphismOf F (s := inl _) (d := inl _) m1).
     simpl in *.
     etransitivity; [ refine (rsimplify_morphisms _) | ].
     Fail reflexivity.
diff -r 4eb6407c6ca3 Cat.v
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Cat.v	Fri Jun 21 15:07:08 2013 -0400
@@ -0,0 +1,46 @@
+Require Import FunctionalExtensionality JMeq ProofIrrelevance.
+Require Export Category CategoryIsomorphisms InitialTerminalCategory Functor ComputableCategory.
+Require Import Common FEqualDep.
+
+Set Implicit Arguments.
+
+Set Asymmetric Patterns.
+
+Set Universe Polymorphism.
+
+Section Cat.
+  Definition Cat := @ComputableCategory _ (@id Category).
+End Cat.
+
+Local Ltac destruct_simple_types :=
+  repeat match goal with
+           | [ |- context[?T] ] => let T' := type of T in
+                                   let T'' := fresh in
+                                   match eval hnf in T' with
+                                     | unit => set (T'' := T); destruct T''
+                                     | _ = _ => set (T'' := T); destruct T''
+                                   end
+         end.
+
+Section Objects.
+  Hint Unfold Morphism Object.
+
+  Hint Extern 1 => destruct_simple_types.
+  Hint Extern 3 => destruct_to_empty_set.
+
+  Lemma TerminalCategory_Terminal : IsTerminalObject (C := Cat) TerminalCategory.
+    repeat intro;
+    exists (FunctorToTerminal _).
+    abstract (
+        repeat intro; functor_eq; eauto
+      ).
+  Defined.
+
+  Lemma InitialCategory_Initial : IsInitialObject (C := Cat) InitialCategory.
+    repeat intro;
+    exists (FunctorFromInitial _).
+    abstract (
+        repeat intro; functor_eq; eauto
+      ).
+  Defined.
+End Objects.
diff -r 4eb6407c6ca3 Category.v
--- a/Category.v	Fri May 24 20:09:37 2013 -0400
+++ b/Category.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,6 @@
-Require Export SpecializedCategory.
-Require Import Common.
+Require Import JMeq ProofIrrelevance.
+Require Export Notations.
+Require Import Common StructureEquality FEqualDep.
 
 Set Implicit Arguments.
 
@@ -9,28 +10,308 @@
 
 Set Universe Polymorphism.
 
-Record > Category := {
-  CObject : Type;
+Local Infix "==" := JMeq.
 
-  UnderlyingCategory :> @SpecializedCategory CObject
-}.
+Record Category :=
+  Build_Category' {
+      Object :> Type;
+      Morphism : Object -> Object -> Type;
 
-Record > SmallCategory := {
-  SObject : Set;
+      Identity : forall x, Morphism x x;
+      Compose : forall s d d', Morphism d d' -> Morphism s d -> Morphism s d' where "f ∘ g" := (@Compose _ _ _ f g);
 
-  SUnderlyingCategory :> @SmallSpecializedCategory SObject
-}.
-
-Record > LocallySmallCategory := {
-  LSObject : Type;
-
-  LSUnderlyingCategory :> @LocallySmallSpecializedCategory LSObject
-}.
+      Associativity : forall o1 o2 o3 o4
+                             (m1 : Morphism o1 o2)
+                             (m2 : Morphism o2 o3)
+                             (m3 : Morphism o3 o4),
+                        (m3 ∘ m2) ∘ m1 = m3 ∘ (m2 ∘ m1);
+      (* ask for [eq_sym (Associativity ...)], so that C^{op}^{op} is convertible with C *)
+      Associativity_sym : forall o1 o2 o3 o4
+                                 (m1 : Morphism o1 o2)
+                                 (m2 : Morphism o2 o3)
+                                 (m3 : Morphism o3 o4),
+                            m3 ∘ (m2 ∘ m1) = (m3 ∘ m2) ∘ m1;
+      LeftIdentity : forall a b (f : Morphism a b), Identity b ∘ f = f;
+      RightIdentity : forall a b (f : Morphism a b), f ∘ Identity a = f
+    }.
 
 Bind Scope category_scope with Category.
-Bind Scope category_scope with SmallCategory.
-Bind Scope category_scope with LocallySmallCategory.
+Bind Scope object_scope with Object.
+Bind Scope morphism_scope with Morphism.
 
+Arguments Object C%category : rename.
+Arguments Morphism !C%category s d : rename. (* , simpl nomatch. *)
+Arguments Identity [!C%category] x%object : rename.
+Arguments Compose [!C%category s%object d%object d'%object] m1%morphism m2%morphism : rename.
+
+Infix "∘" := Compose : morphism_scope.
+
+Section CategoryInterface.
+  Definition Build_Category (Object' : Type)
+             (Morphism' : Object' -> Object' -> Type)
+             (Identity' : forall o : Object', Morphism' o o)
+             (Compose' : forall s d d' : Object', Morphism' d d' -> Morphism' s d -> Morphism' s d')
+             (Associativity' : forall (o1 o2 o3 o4 : Object') (m1 : Morphism' o1 o2) (m2 : Morphism' o2 o3) (m3 : Morphism' o3 o4),
+                                 Compose' o1 o2 o4 (Compose' o2 o3 o4 m3 m2) m1 = Compose' o1 o3 o4 m3 (Compose' o1 o2 o3 m2 m1))
+             (LeftIdentity' : forall (a b : Object') (f : Morphism' a b), Compose' a b b (Identity' b) f = f)
+             (RightIdentity' : forall (a b : Object') (f : Morphism' a b), Compose' a a b f (Identity' a) = f)
+  : Category
+    := @Build_Category' Object'
+                        Morphism'
+                        Identity'
+                        Compose'
+                        Associativity'
+                        (fun _ _ _ _ _ _ _ => eq_sym (Associativity' _ _ _ _ _ _ _))
+                        LeftIdentity'
+                        RightIdentity'.
+End CategoryInterface.
+
+(* create a hint db for all category theory things *)
+Create HintDb category discriminated.
+(* create a hint db for morphisms in categories *)
+Create HintDb morphism discriminated.
+
+Hint Extern 1 => symmetry : category morphism. (* TODO(jgross): Why do I need this? *)
+
+Ltac category_hideProofs :=
+  repeat match goal with
+             | [ |- context[{|
+                               Associativity := ?pf0;
+                               Associativity_sym := ?pf1;
+                               LeftIdentity := ?pf2;
+                               RightIdentity := ?pf3
+                             |}] ] =>
+               hideProofs pf0 pf1 pf2 pf3
+         end.
+
+Hint Resolve @LeftIdentity @RightIdentity @Associativity : category morphism.
+Hint Rewrite @LeftIdentity @RightIdentity : category.
+Hint Rewrite @LeftIdentity @RightIdentity : morphism.
+
+Section Categories_Equal.
+(*  Lemma Category_contr_eq' (C D : Category)
+        (C_morphism_proof_irrelevance
+         : forall s d (m1 m2 : Morphism C s d) (pf1 pf2 : m1 = m2),
+             pf1 = pf2)
+  : forall (HO : Object C = Object D)
+           (HM : match HO in _ = y return y -> y -> Type with
+                   | eq_refl => @Morphism C
+                 end = @Morphism D),
+      match HO in (_ = c) return _ with
+        | eq_refl => match HM in (_ = y) return (forall x : D, y x x) with
+                       | eq_refl => Identity (C := C)
+                     end
+      end = Identity (C := D)
+      -> match
+        HM in (_ = y) return (forall s d d' : C, y d d' -> y s d -> y s d')
+      with
+        | eq_refl => Compose (C := C)
+      end = Compose (C := D)
+      -> C = D.
+    intros.
+    destruct C, D;
+      subst_body;
+      intros;
+      simpl in *.
+    subst.
+    repeat f_equal;
+      repeat (apply functional_extensionality_dep; intro);
+      trivial.
+  Qed.
+
+  Lemma Category_contr_eq `(C : @Category objC) `(D : @Category objC)
+        (C_morphism_proof_irrelevance
+         : forall s d (m1 m2 : Morphism C s d) (pf1 pf2 : m1 = m2),
+             pf1 = pf2)
+        (C_morphism_type_contr
+         : forall s d (pf1 pf2 : Morphism C s d = Morphism C s d),
+             pf1 = pf2)
+  : forall (HM : forall s d, @Morphism _ C s d = @Morphism _ D s d),
+      (forall x,
+         match HM x x in (_ = y) return y with
+           | eq_refl => Identity (C := C) x
+         end = Identity (C := D) x)
+      -> (forall s d d' (m : Morphism D d d') (m' : Morphism D s d),
+            match HM s d' in (_ = y) return y with
+              | eq_refl =>
+                match HM s d in (_ = y) return (y -> Morphism C s d') with
+                  | eq_refl =>
+                    match
+                      HM d d' in (_ = y) return (y -> Morphism C s d -> Morphism C s d')
+                    with
+                      | eq_refl => Compose (d':=d')
+                    end m
+                end m'
+            end = Compose m m')
+      -> C = D.
+    intros HM HI HC.
+    assert (HM' : @Morphism _ C = @Morphism _ D)
+      by (repeat (apply functional_extensionality_dep; intro); trivial).
+    apply (Category_contr_eq' _ _ C_morphism_proof_irrelevance HM');
+      revert HI HC C_morphism_proof_irrelevance C_morphism_type_contr;
+      destruct C, D; simpl in *; clear;
+      intros;
+      subst_body;
+      simpl in *;
+      repeat (subst || intro || apply functional_extensionality_dep);
+      rewrite_rev_hyp;
+      generalize_eq_match;
+      subst_eq_refl_dec;
+      trivial.
+  Qed.
+*)
+  Lemma Category_eq (C D : Category)
+  : Object C = Object D
+    -> @Morphism C == @Morphism D
+    -> @Identity C == @Identity D
+    -> @Compose C == @Compose D
+    -> C = D.
+  Proof.
+    intros.
+    destruct_head @Category;
+      simpl in *; repeat subst;
+      f_equal; apply proof_irrelevance.
+  Qed.
+End Categories_Equal.
+
+Ltac cat_eq_step_with tac :=
+  structures_eq_step_with_tac ltac:(apply Category_eq) tac.
+
+Ltac cat_eq_with tac := repeat cat_eq_step_with tac.
+
+Ltac cat_eq_step := cat_eq_step_with idtac.
+Ltac cat_eq := category_hideProofs; cat_eq_with idtac.
+
+Ltac solve_for_identity :=
+  match goal with
+    | [ |- @Compose ?C ?s ?s ?d ?a ?b = ?b ]
+      => cut (a = @Identity C s);
+        [ try solve [ let H := fresh in intro H; rewrite H; apply LeftIdentity ] | ]
+    | [ |- @Compose ?C ?s ?d ?d ?a ?b = ?a ]
+      => cut (b = @Identity C d );
+        [ try solve [ let H := fresh in intro H; rewrite H; apply RightIdentity ] | ]
+  end.
+
+Local Open Scope morphism_scope.
+
+(** * Version of [Associativity] that avoids going off into the weeds in the presence of unification variables *)
+
+Definition NoEvar T (_ : T) := True.
+
+Lemma AssociativityNoEvar (C : Category) : forall (o1 o2 o3 o4 : C) (m1 : C.(Morphism) o1 o2)
+  (m2 : C.(Morphism) o2 o3) (m3 : C.(Morphism) o3 o4),
+  NoEvar (m1, m2) \/ NoEvar (m2, m3) \/ NoEvar (m1, m3)
+  -> (m3 ∘ m2) ∘ m1 = m3 ∘ (m2 ∘ m1).
+  intros; apply Associativity.
+Qed.
+
+Ltac noEvar := match goal with
+                 | [ |- context[NoEvar ?X] ] => (has_evar X; fail 1)
+                                                  || cut (NoEvar X); [ intro; tauto | constructor ]
+               end.
+
+Hint Rewrite @AssociativityNoEvar using noEvar : category.
+Hint Rewrite @AssociativityNoEvar using noEvar : morphism.
+
+Ltac try_associativity_quick tac := try_rewrite Associativity tac.
+Ltac try_associativity tac := try_rewrite_by AssociativityNoEvar ltac:(idtac; noEvar) tac.
+
+Ltac find_composition_to_identity :=
+  match goal with
+    | [ H : @Compose _ _ _ _ _ ?a ?b = @Identity _ _ _ |- context[@Compose ?A ?B ?C ?D ?E ?c ?d] ]
+      => let H' := fresh in
+        assert (H' : b = d /\ a = c) by (split; reflexivity); clear H';
+          assert (H' : @Compose A B C D E c d = @Identity _ _ _) by (
+            exact H ||
+              (simpl in H |- *; exact H || (rewrite H; reflexivity))
+          );
+          first [
+            rewrite H'
+            | simpl in H' |- *; rewrite H'
+            | let H'T := type of H' in fail 2 "error in rewriting a found identity" H "[" H'T "]"
+          ]; clear H'
+  end.
+
+(** * Back to the main content.... *)
+
+Section Category.
+  Variable C : Category.
+
+  (* Quoting Wikipedia,
+    In category theory, an epimorphism (also called an epic
+    morphism or, colloquially, an epi) is a morphism [f : X → Y]
+    which is right-cancellative in the sense that, for all
+    morphisms [g, g' : Y → Z],
+    [g ∘ f = g' ∘ f -> g = g']
+
+    Epimorphisms are analogues of surjective functions, but they
+    are not exactly the same. The dual of an epimorphism is a
+    monomorphism (i.e. an epimorphism in a category [C] is a
+    monomorphism in the dual category [OppositeCategory C]).
+    *)
+  Definition IsEpimorphism x y (m : C.(Morphism) x y) : Prop :=
+    forall z (m1 m2 : C.(Morphism) y z), m1 ∘ m = m2 ∘ m ->
+      m1 = m2.
+  Definition IsMonomorphism x y (m : C.(Morphism) x y) : Prop :=
+    forall z (m1 m2 : C.(Morphism) z x), m ∘ m1 = m ∘ m2 ->
+      m1 = m2.
+
+  Section properties.
+    Lemma IdentityIsEpimorphism x : IsEpimorphism _ _ (Identity x).
+      repeat intro; autorewrite with category in *; trivial.
+    Qed.
+
+    Lemma IdentityIsMonomorphism x : IsMonomorphism _ _ (Identity x).
+    Proof.
+      repeat intro; autorewrite with category in *; trivial.
+    Qed.
+
+    Lemma EpimorphismComposition s d d' m0 m1
+    : IsEpimorphism _ _ m0
+      -> IsEpimorphism _ _ m1
+      -> IsEpimorphism _ _ (Compose (C := C) (s := s) (d := d) (d' := d') m0 m1).
+    Proof.
+      repeat intro.
+      repeat match goal with | [ H : _ |- _ ] => rewrite <- Associativity in H end.
+      intuition.
+    Qed.
+
+    Lemma MonomorphismComposition s d d' m0 m1
+    : IsMonomorphism _ _ m0
+      -> IsMonomorphism _ _ m1
+      -> IsMonomorphism _ _ (Compose (C := C) (s := s) (d := d) (d' := d') m0 m1).
+    Proof.
+      repeat intro.
+      repeat match goal with | [ H : _ |- _ ] => rewrite Associativity in H end.
+      intuition.
+    Qed.
+  End properties.
+End Category.
+
+Hint Immediate @IdentityIsEpimorphism @IdentityIsMonomorphism @MonomorphismComposition @EpimorphismComposition : category morphism.
+
+Arguments IsEpimorphism [C x y] m.
+Arguments IsMonomorphism [C x y] m.
+
+Section AssociativityComposition.
+  Variable C : Category.
+  Variables o0 o1 o2 o3 o4 : C.
+
+  Lemma compose4associativity_helper
+    (a : Morphism C o3 o4) (b : Morphism C o2 o3)
+    (c : Morphism C o1 o2) (d : Morphism C o0 o1)
+  : (a ∘ b) ∘ (c ∘ d) = (a ∘ ((b ∘ c) ∘ d)).
+  Proof.
+    repeat rewrite Associativity; reflexivity.
+  Qed.
+End AssociativityComposition.
+
+Ltac compose4associativity' a b c d := transitivity (Compose a (Compose (Compose b c) d)); try solve [ apply compose4associativity_helper ].
+Ltac compose4associativity :=
+  match goal with
+    | [ |- Compose (Compose ?a ?b) (Compose ?c ?d) = _ ] => compose4associativity' a b c d
+    | [ |- _ = Compose (Compose ?a ?b) (Compose ?c ?d) ] => compose4associativity' a b c d
+  end.
 
 
 (** * The saturation prover: up to some bound on number of steps, consider all ways to extend equivalences with pre- or post-composition. *)
diff -r 4eb6407c6ca3 CategoryIsomorphisms.v
--- a/CategoryIsomorphisms.v	Fri May 24 20:09:37 2013 -0400
+++ b/CategoryIsomorphisms.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import ProofIrrelevance Setoid.
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import Common StructureEquality.
 
 Set Implicit Arguments.
@@ -10,22 +10,25 @@
 
 Set Universe Polymorphism.
 
+Local Open Scope morphism_scope.
+
 Section Category.
-  Context `{C : @SpecializedCategory obj}.
+  Context {C : Category}.
 
   (* [m'] is the inverse of [m] if both compositions are
      equivalent to the relevant identity morphisms. *)
   (* [Definitions] don't get sort-polymorphism :-(  *)
   Definition IsInverseOf1 (s d : C) (m : C.(Morphism) s d) (m' : C.(Morphism) d s) : Prop :=
-    Compose m' m = Identity s.
+    m' ∘ m = Identity s.
   Definition IsInverseOf2 (s d : C) (m : C.(Morphism) s d) (m' : C.(Morphism) d s) : Prop :=
-    Compose m m' = Identity d.
+    m ∘ m' = Identity d.
 
   Global Arguments IsInverseOf1 / _ _ _ _.
   Global Arguments IsInverseOf2 / _ _ _ _.
 
-  Definition IsInverseOf {s d : C} (m : C.(Morphism) s d) (m' : C.(Morphism) d s) : Prop := Eval simpl in
-    @IsInverseOf1 s d m m' /\ @IsInverseOf2 s d m m'.
+  Definition IsInverseOf {s d : C} (m : C.(Morphism) s d) (m' : C.(Morphism) d s) : Prop :=
+    Eval simpl in
+      @IsInverseOf1 s d m m' /\ @IsInverseOf2 s d m m'.
 
   Lemma IsInverseOf_sym s d m m' : @IsInverseOf s d m m' -> @IsInverseOf d s m' m.
     firstorder.
@@ -48,30 +51,31 @@
   (* [Record]s are [Inductive] and get sort-polymorphism *)
   Section IsomorphismOf.
     (* A morphism is an isomorphism if it has an inverse *)
-    Record IsomorphismOf {s d : C} (m : C.(Morphism) s d) := {
-      IsomorphismOf_Morphism :> C.(Morphism) s d := m;
-      Inverse : C.(Morphism) d s;
-      LeftInverse : Compose Inverse m = Identity s;
-      RightInverse : Compose m Inverse = Identity d
-    }.
+    Record IsomorphismOf {s d : C} (m : C.(Morphism) s d) :=
+      {
+        IsomorphismOf_Morphism :> C.(Morphism) s d := m;
+        Inverse : C.(Morphism) d s;
+        LeftInverse : Inverse ∘ m = Identity s;
+        RightInverse : m ∘ Inverse = Identity d
+      }.
 
     Hint Resolve RightInverse LeftInverse : category.
     Hint Resolve RightInverse LeftInverse : morphism.
 
     Definition IsomorphismOf_sig2 {s d : C} (m : C.(Morphism) s d) (i : @IsomorphismOf s d m) :
-      { m' | Compose m' m = Identity s & Compose m m' = Identity d }.
+      { m' | m' ∘ m = Identity s & m ∘ m' = Identity d }.
       exists (Inverse i);
-        [ apply LeftInverse | apply RightInverse ].
+      [ apply LeftInverse | apply RightInverse ].
     Defined.
 
-    Definition IsomorphismOf_sig {s d : C} (m : C.(Morphism) s d) := { m' | Compose m' m = Identity s & Compose m m' = Identity d }.
+    Definition IsomorphismOf_sig {s d : C} (m : C.(Morphism) s d) := { m' | m' ∘ m = Identity s & m ∘ m' = Identity d }.
 
     Global Identity Coercion Isomorphism_sig : IsomorphismOf_sig >-> sig2.
 
     Definition sig2_IsomorphismOf {s d : C} (m : C.(Morphism) s d) (i : @IsomorphismOf_sig s d m) :
       @IsomorphismOf s d m.
       exists (proj1_sig i);
-        [ apply (proj2_sig i) | apply (proj3_sig i) ].
+      [ apply (proj2_sig i) | apply (proj3_sig i) ].
     Defined.
 
     Global Coercion IsomorphismOf_sig2 : IsomorphismOf >-> sig2.
@@ -85,9 +89,9 @@
       exists (i : Morphism C _ _); auto with morphism.
     Defined.
 
-    Definition ComposeIsomorphismOf {s d d' : C} {m1 : C.(Morphism) d d'} {m2 : C.(Morphism) s d} (i1 : IsomorphismOf m1) (i2 : IsomorphismOf m2) :
-      IsomorphismOf (Compose m1 m2).
-      exists (Compose (Inverse i2) (Inverse i1));
+    Definition ComposeIsomorphismOf {s d d' : C} {m1 : C.(Morphism) d d'} {m2 : C.(Morphism) s d} (i1 : IsomorphismOf m1) (i2 : IsomorphismOf m2)
+    : IsomorphismOf (m1 ∘ m2).
+      exists (Inverse i2 ∘ Inverse i1);
       abstract (
           simpl; compose4associativity;
           destruct i1, i2; simpl;
@@ -98,10 +102,11 @@
   End IsomorphismOf.
 
   Section Isomorphism.
-    Record Isomorphism (s d : C) := {
-      Isomorphism_Morphism : C.(Morphism) s d;
-      Isomorphism_Of :> IsomorphismOf Isomorphism_Morphism
-    }.
+    Record Isomorphism (s d : C) :=
+      {
+        Isomorphism_Morphism : C.(Morphism) s d;
+        Isomorphism_Of :> IsomorphismOf Isomorphism_Morphism
+      }.
 
     Global Coercion Build_Isomorphism : IsomorphismOf >-> Isomorphism.
   End Isomorphism.
@@ -142,16 +147,16 @@
     Qed.
 
     Local Ltac t_iso' := intros;
-      repeat match goal with
-               | [ i : Isomorphic _ _ |- _ ] => destruct (Isomorphic_Isomorphism i) as [ ? [] ] ; clear i
-               | [ |- Isomorphic _ _ ] => apply Isomrphism_Isomorphic
-             end.
+                        repeat match goal with
+                                 | [ i : Isomorphic _ _ |- _ ] => destruct (Isomorphic_Isomorphism i) as [ ? [] ] ; clear i
+                                 | [ |- Isomorphic _ _ ] => apply Isomrphism_Isomorphic
+                               end.
 
     Local Ltac t_iso:= t_iso';
-      repeat match goal with
-               | [ i : Isomorphism _ _ |- _ ] => unique_pose (Isomorphism_Of i); try clear i
-               | [ |- Isomorphism _ _ ] => eapply Build_Isomorphism
-             end.
+                      repeat match goal with
+                               | [ i : Isomorphism _ _ |- _ ] => unique_pose (Isomorphism_Of i); try clear i
+                               | [ |- Isomorphism _ _ ] => eapply Build_Isomorphism
+                             end.
 
     Hint Resolve @IsomorphismOf_Identity @InverseOf @ComposeIsomorphismOf : category morphism.
     Local Hint Extern 1 => eassumption.
@@ -175,24 +180,24 @@
     Qed.
 
     Global Add Parametric Relation : _ Isomorphic
-      reflexivity proved by Isomorphic_refl
-      symmetry proved by Isomorphic_sym
-      transitivity proved by Isomorphic_trans
-        as Isomorphic_rel.
+        reflexivity proved by Isomorphic_refl
+        symmetry proved by Isomorphic_sym
+        transitivity proved by Isomorphic_trans
+          as Isomorphic_rel.
   End Isomorphic.
 
   (* XXX TODO: Automate this better. *)
   Lemma iso_is_epi s d (m : _ s d) : IsIsomorphism m -> IsEpimorphism (C := C) m.
     destruct 1 as [ x [ i0 i1 ] ]; intros z m1 m2 e.
-    transitivity (Compose m1 (Compose m x)); [ rewrite_hyp; autorewrite with morphism | ]; trivial.
-    transitivity (Compose m2 (Compose m x)); [ repeat rewrite <- Associativity | ]; rewrite_hyp; autorewrite with morphism; trivial.
+    transitivity (m1 ∘ (m ∘ x)); [ rewrite_hyp; autorewrite with morphism | ]; trivial.
+    transitivity (m2 ∘ (m ∘ x)); [ repeat rewrite <- Associativity | ]; rewrite_hyp; autorewrite with morphism; trivial.
   Qed.
 
   (* XXX TODO: Automate this better. *)
   Lemma iso_is_mono s d (m : _ s d) : IsIsomorphism m -> IsMonomorphism (C := C) m.
     destruct 1 as [ x [ i0 i1 ] ]; intros z m1 m2 e.
-    transitivity (Compose (Compose x m) m1); [ rewrite_hyp; autorewrite with morphism | ]; trivial.
-    transitivity (Compose (Compose x m) m2); [ repeat rewrite Associativity | ]; rewrite_hyp; autorewrite with morphism; trivial.
+    transitivity ((x ∘ m) ∘ m1); [ rewrite_hyp; autorewrite with morphism | ]; trivial.
+    transitivity ((x ∘ m) ∘ m2); [ repeat rewrite Associativity | ]; rewrite_hyp; autorewrite with morphism; trivial.
   Qed.
 End Category.
 
@@ -200,22 +205,22 @@
 
 Ltac eapply_by_compose H :=
   match goal with
-    | [ |- @eq (@Morphism ?obj ?mor ?C) _ _ ] => eapply (H obj mor C)
-    | [ |- @Compose ?obj ?mor ?C _ _ _ _ _ = _ ] => eapply (H obj mor C)
-    | [ |- _ = @Compose ?obj ?mor ?C _ _ _ _ _ ] => eapply (H obj mor C)
+    | [ |- @eq (@Morphism ?C) _ _ ] => eapply (H C)
+    | [ |- @Compose ?C _ _ _ _ _ = _ ] => eapply (H C)
+    | [ |- _ = @Compose ?C _ _ _ _ _ ] => eapply (H C)
     | _ => eapply H
-    | [ C : @SpecializedCategory ?obj ?mor |- _ ] => eapply (H obj mor C)
-    | [ C : ?T |- _ ] => match eval hnf in T with | @SpecializedCategory ?obj ?mor => eapply (H obj mor C) end
+    | [ C : Category |- _ ] => eapply (H C)
+    | [ C : ?T |- _ ] => match eval hnf in T with | Category => eapply (H C) end
   end.
 
 Ltac solve_isomorphism := destruct_hypotheses;
-  solve [ eauto ] ||
-    match goal with
-      | [ _ : Compose ?x ?x' = Identity _ |- IsIsomorphism ?x ] => solve [ exists x'; hnf; eauto ]
-      | [ _ : Compose ?x ?x' = Identity _ |- Isomorphism ?x ] => solve [ exists x'; hnf; eauto ]
-      | [ _ : Compose ?x ?x' = Identity _ |- IsomorphismOf ?x ] => solve [ exists x'; hnf; eauto ]
-      | [ _ : Compose ?x ?x' = Identity _ |- context[Compose ?x _ = Identity _] ] => solve [ try exists x'; hnf; eauto ]
-    end.
+                         solve [ eauto ] ||
+                               match goal with
+                                 | [ _ : Compose ?x ?x' = Identity _ |- IsIsomorphism ?x ] => solve [ exists x'; hnf; eauto ]
+                                 | [ _ : Compose ?x ?x' = Identity _ |- Isomorphism ?x ] => solve [ exists x'; hnf; eauto ]
+                                 | [ _ : Compose ?x ?x' = Identity _ |- IsomorphismOf ?x ] => solve [ exists x'; hnf; eauto ]
+                                 | [ _ : Compose ?x ?x' = Identity _ |- context[Compose ?x _ = Identity _] ] => solve [ try exists x'; hnf; eauto ]
+                               end.
 
 (* [eapply] the theorem to get a pre/post composed mono/epi, then find the right one by looking
    for an [Identity], then solve the requirement that it's an isomorphism *)
@@ -229,7 +234,7 @@
   [ solve_isomorphism | ].
 
 Section CategoryObjects1.
-  Context `(C : @SpecializedCategory obj).
+  Variable C : Category.
 
   Definition UniqueUpToUniqueIsomorphism' (P : C.(Object) -> Prop) : Prop :=
     forall o, P o -> forall o', P o' -> exists m : C.(Morphism) o o', IsIsomorphism m /\ is_unique m.
@@ -247,7 +252,7 @@
 
     Record TerminalObject :=
       {
-        TerminalObject_Object' : obj;
+        TerminalObject_Object' : C;
         TerminalObject_Morphism : forall o, Morphism C o TerminalObject_Object';
         TerminalObject_Property : forall o, is_unique (TerminalObject_Morphism o)
       }.
@@ -276,7 +281,7 @@
 
     Record InitialObject :=
       {
-        InitialObject_Object' :> obj;
+        InitialObject_Object' :> C;
         InitialObject_Morphism : forall o, Morphism C InitialObject_Object' o;
         InitialObject_Property : forall o, is_unique (InitialObject_Morphism o)
       }.
@@ -296,21 +301,21 @@
   End initial.
 End CategoryObjects1.
 
-Arguments UniqueUpToUniqueIsomorphism {_ C} P.
-Arguments IsInitialObject' {_ C} o.
-Arguments IsInitialObject {_ C} o.
-Arguments IsTerminalObject' {_ C} o.
-Arguments IsTerminalObject {_ C} o.
+Arguments UniqueUpToUniqueIsomorphism {C} P.
+Arguments IsInitialObject' {C} o.
+Arguments IsInitialObject {C} o.
+Arguments IsTerminalObject' {C} o.
+Arguments IsTerminalObject {C} o.
 
 Section CategoryObjects2.
-  Context `(C : @SpecializedCategory obj).
+  Variable C : Category.
 
   Ltac unique := hnf; intros; specialize_all_ways; destruct_sig;
-    unfold is_unique, unique, uniqueness in *;
-      repeat (destruct 1);
-      repeat match goal with
-               | [ x : _ |- _ ] => exists x
-             end; eauto with category; try split; try solve [ etransitivity; eauto with category ].
+                 unfold is_unique, unique, uniqueness in *;
+                 repeat (destruct 1);
+                 repeat match goal with
+                          | [ x : _ |- _ ] => (exists x)
+                        end; eauto with category; try split; try solve [ etransitivity; eauto with category ].
 
   (* The terminal object is unique up to unique isomorphism. *)
   Theorem TerminalObjectUnique : UniqueUpToUniqueIsomorphism (IsTerminalObject (C := C)).
diff -r 4eb6407c6ca3 CategorySchemaEquivalence.v
--- a/CategorySchemaEquivalence.v	Fri May 24 20:09:37 2013 -0400
+++ b/CategorySchemaEquivalence.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import Bool Omega Setoid Program.
-Require Export Schema Category SmallSchema SmallCategory.
+Require Export Schema Category SmallSchema Category.
 Require Import Common EquivalenceClass.
 Require Import NaturalEquivalence ComputableCategory SNaturalEquivalence ComputableSchemaCategory SmallFunctor SmallTranslation.
 
@@ -10,7 +10,7 @@
 Set Universe Polymorphism.
 
 Section Schema_Category_Equivalence.
-  Variable C : SmallCategory.
+  Variable C : Category.
   Variable S : SmallSchema.
 
   Hint Rewrite concatenate_noedges_p concatenate_p_noedges concatenate_addedge.
@@ -66,7 +66,7 @@
                end; clear_paths; repeat rewrite concatenate_associative in *; repeat rewrite sconcatenate_associative in *; try reflexivity || symmetry.
 
   (* TODO: Speed this up, automate this better. *)
-  Definition saturate : SmallCategory.
+  Definition saturate : Category.
     refine {| SObject := S;
       SMorphism := (fun s d => EquivalenceClass (@SPathsEquivalent S s d));
       (* foo := 1; *) (* uncommenting this line gives "Anomaly: uncaught exception Not_found. Please report."  Maybe I should report this?  But I haven't figured out a minimal test case. *)
@@ -108,7 +108,7 @@
 End Schema_Category_Equivalence.
 
 Section CategorySchemaCategory_RoundTrip.
-  Variable C : SmallCategory.
+  Variable C : Category.
 
   Hint Rewrite sconcatenate_noedges_p sconcatenate_p_noedges sconcatenate_addedge.
   Hint Rewrite <- sconcatenate_prepend_equivalent.
diff -r 4eb6407c6ca3 ChainCategory.v
--- a/ChainCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/ChainCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import ProofIrrelevance.
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import NatFacts Subcategory DefinitionSimplification.
 
 Set Implicit Arguments.
@@ -9,8 +9,8 @@
 Set Universe Polymorphism.
 
 Section ChainCat.
-  Definition OmegaCategory : @SpecializedCategory nat.
-    refine (@Build_SpecializedCategory _
+  Definition OmegaCategory : @Category nat.
+    refine (@Build_Category _
                                        le
                                        le_refl
                                        (fun _ _ _ H0 H1 => le_trans H1 H0)
@@ -20,7 +20,7 @@
     abstract (intros; apply proof_irrelevance).
   Defined.
 
-  Let ChainCategory' n : @SpecializedCategory { m | m <= n }.
+  Let ChainCategory' n : @Category { m | m <= n }.
     simpl_definition_by_tac_and_exact (FullSubcategory OmegaCategory (fun m => m <= n)) ltac:(unfold Subcategory in *).
   Defined.
   Definition ChainCategory n := Eval cbv beta iota zeta delta [ChainCategory'] in ChainCategory' n.
diff -r 4eb6407c6ca3 Coend.v
--- a/Coend.v	Fri May 24 20:09:37 2013 -0400
+++ b/Coend.v	Fri Jun 21 15:07:08 2013 -0400
@@ -37,12 +37,12 @@
      ]]
      where the triangles denote induced colimit morphisms.
      *)
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variable C : Category.
+  Variable D : Category.
 
   Let COp := OppositeCategory C.
 
-  Variable F : SpecializedFunctor (COp * C) D.
+  Variable F : Functor (COp * C) D.
 
   Let MorC := @MorphismFunctor _ _ (fun _ : unit => C) tt. (* [((c0, c1) & f : morC c0 c1)], the set of morphisms of C *)
 
@@ -54,8 +54,8 @@
     apply (InducedColimitMap (G := InducedDiscreteFunctor _ (DomainNaturalTransformation _ (fun _ => C) tt))).
     hnf; simpl.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F0 ?G0 ] =>
-        refine (Build_SpecializedNaturalTransformation F0 G0
+      | [ |- NaturalTransformation ?F0 ?G0 ] =>
+        refine (Build_NaturalTransformation F0 G0
           (fun sdf => let s := fst (projT1 sdf) in MorphismOf F (s := (_, s)) (d := (_, s)) (projT2 sdf, Identity (C := C) s))
           _
         )
@@ -68,8 +68,8 @@
     apply (InducedColimitMap (G := InducedDiscreteFunctor _ (CodomainNaturalTransformation _ (fun _ => C) tt))).
     hnf; simpl.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F0 ?G0 ] =>
-        refine (Build_SpecializedNaturalTransformation F0 G0
+      | [ |- NaturalTransformation ?F0 ?G0 ] =>
+        refine (Build_NaturalTransformation F0 G0
           (fun sdf => let d := snd (projT1 sdf) in MorphismOf F (s := (d, _)) (d := (d, _)) (Identity (C := C) d, projT2 sdf))
           _
         )
diff -r 4eb6407c6ca3 CoendFunctor.v
--- a/CoendFunctor.v	Fri May 24 20:09:37 2013 -0400
+++ b/CoendFunctor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -13,7 +13,7 @@
 Local Open Scope type_scope.
 
 Section Coend.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Let COp := OppositeCategory C.
 
@@ -52,9 +52,9 @@
     end.
   Defined.
 
-  Definition CoendFunctor_Index : SpecializedCategory CoendFunctor_Index_Object.
+  Definition CoendFunctor_Index : Category CoendFunctor_Index_Object.
   Proof.
-    refine (@Build_SpecializedCategory _
+    refine (@Build_Category _
                                        CoendFunctor_Index_Morphism
                                        CoendFunctor_Index_Identity
                                        CoendFunctor_Index_Compose
@@ -105,9 +105,9 @@
       | _ => injection H; intros; subst
     end.
 
-  Definition CoendFunctor_Diagram_pre : SpecializedFunctor CoendFunctor_Index (COp * C).
+  Definition CoendFunctor_Diagram_pre : Functor CoendFunctor_Index (COp * C).
   Proof.
-    refine (Build_SpecializedFunctor
+    refine (Build_Functor
       CoendFunctor_Index (COp * C)
       CoendFunctor_Diagram_ObjectOf_pre
       CoendFunctor_Diagram_MorphismOf_pre
@@ -133,12 +133,12 @@
 End Coend.
 
 Section CoendFunctor.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variable C : Category.
+  Variable D : Category.
 
   Let COp := OppositeCategory C.
 
-  Hypothesis HasColimits : forall F : SpecializedFunctor (CoendFunctor_Index C) D, Colimit F.
+  Hypothesis HasColimits : forall F : Functor (CoendFunctor_Index C) D, Colimit F.
 
   Let CoendFunctor_post := ColimitFunctor HasColimits.
 
diff -r 4eb6407c6ca3 CommaCategory.v
--- a/CommaCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/CommaCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,25 +1,22 @@
-Require Import ProofIrrelevance.
-Require Export Category SpecializedCategory Functor ProductCategory.
-Require Import Common Notations InitialTerminalCategory SpecializedCommaCategory DefinitionSimplification.
+Require Import JMeq ProofIrrelevance.
+Require Export Category Functor ProductCategory InitialTerminalCategory.
+Require Import Common Notations FEqualDep.
 
 Set Implicit Arguments.
 
+Generalizable All Variables.
+
 Set Asymmetric Patterns.
 
 Set Universe Polymorphism.
 
+Local Infix "==" := JMeq.
+
+Local Open Scope morphism_scope.
 Local Open Scope category_scope.
 
-Local Ltac fold_functor := idtac. (*
-  change CObject with (fun C => @Object (CObject C) C) in *;
-    change (@SpecializedFunctor) with (fun objC (C : @SpecializedCategory objC) objD (D : @SpecializedCategory objD) => @Functor C D) in *. *)
-
 Section CommaCategory.
-  (* [Definition]s are not sort-polymorphic, and it's too slow to not use
-     [Definition]s, so we might as well use [Category]s rather than [SpecializedCategory]s. *)
-  Variable A : Category.
-  Variable B : Category.
-  Variable C : Category.
+  Variables A B C : Category.
   Variable S : Functor A C.
   Variable T : Functor B C.
 
@@ -35,7 +32,8 @@
          pairs [(g, h)] where [g : α -> α'] and [h : β -> β'] are
          morphisms in [A] and [B] respectively, such that the
          following diagram commutes:
-         [[
+
+<<
                S g
           S α -----> S α'
            |          |
@@ -44,98 +42,122 @@
            V          V
           T β -----> T β'
                T h
-         ]]
+>>
+
      Morphisms are composed by taking [(g, h) ○ (g', h')] to be
-     [(g ○ g', h ○ h')], whenever the latter expression is defined.
+     [(g ∘ g', h ∘ h')], whenever the latter expression is defined.
      The identity morphism on an object [(α, β, f)] is [(Identity α, Identity β)].
-     *)
+   *)
 
   (* By definining all the parts separately, we can make the [Prop]
      Parts of the definition opaque via [abstract].  This speeds things
      up significantly.  We unfold the definitions at the very end with
      [Eval]. *)
-  (* stupid lack of sort-polymorphism in definitions... *)
-  Let CommaCategory_Object' := Eval hnf in CommaSpecializedCategory_ObjectT S T.
-  Let CommaCategory_Object'' : Type.
-    simpl_definition_by_tac_and_exact CommaCategory_Object' ltac:(simpl in *; fold_functor).
-  Defined.
-  Definition CommaCategory_Object := Eval cbv beta delta [CommaCategory_Object''] in CommaCategory_Object''.
+  Definition CommaCategory_Object := { αβ : (A * B)%type & C.(Morphism) (S (fst αβ)) (T (snd αβ)) }.
 
-  Let CommaCategory_Morphism' (XG X'G' : CommaCategory_Object) := Eval hnf in CommaSpecializedCategory_MorphismT (S := S) (T := T) XG X'G'.
-  Let CommaCategory_Morphism'' (XG X'G' : CommaCategory_Object) : Type.
-    simpl_definition_by_tac_and_exact (CommaCategory_Morphism' XG X'G') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
-  Defined.
-  Definition CommaCategory_Morphism (XG X'G' : CommaCategory_Object) := Eval hnf in CommaCategory_Morphism'' XG X'G'.
+  Definition CommaCategory_Morphism (αβf α'β'f' : CommaCategory_Object) :=
+    { gh : (A.(Morphism) (fst (projT1 αβf)) (fst (projT1 α'β'f'))) * (B.(Morphism) (snd (projT1 αβf)) (snd (projT1 α'β'f')))  |
+      T.(MorphismOf) (snd gh) ∘ projT2 αβf = projT2 α'β'f' ∘ S.(MorphismOf) (fst gh)
+    }.
 
-  Let CommaCategory_Compose' s d d' Fα F'α'
-    := Eval hnf in CommaSpecializedCategory_Compose (S := S) (T := T) (s := s) (d := d) (d' := d') Fα F'α'.
-  Let CommaCategory_Compose'' s d d' (Fα : CommaCategory_Morphism d d') (F'α' : CommaCategory_Morphism s d) :
-    CommaCategory_Morphism s d'.
-    simpl_definition_by_tac_and_exact (@CommaCategory_Compose' s d d' Fα F'α') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
-  Defined.
-  Definition CommaCategory_Compose s d d' (Fα : CommaCategory_Morphism d d') (F'α' : CommaCategory_Morphism s d) :
-    CommaCategory_Morphism s d'
-    := Eval hnf in @CommaCategory_Compose'' s d d' Fα F'α'.
-
-  Let CommaCategory_Identity' o := Eval hnf in CommaSpecializedCategory_Identity (S := S) (T := T) o.
-  Let CommaCategory_Identity'' (o : CommaCategory_Object) : CommaCategory_Morphism o o.
-    simpl_definition_by_tac_and_exact (@CommaCategory_Identity' o) ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
-  Defined.
-  Definition CommaCategory_Identity (o : CommaCategory_Object) : CommaCategory_Morphism o o
-    := Eval hnf in @CommaCategory_Identity'' o.
-
-  Global Arguments CommaCategory_Compose _ _ _ _ _ /.
-  Global Arguments CommaCategory_Identity _ /.
-
-  Lemma CommaCategory_Associativity o1 o2 o3 o4 (m1 : CommaCategory_Morphism o1 o2) (m2 : CommaCategory_Morphism o2 o3) (m3 : CommaCategory_Morphism o3 o4) :
-    CommaCategory_Compose (CommaCategory_Compose m3 m2) m1 = CommaCategory_Compose m3 (CommaCategory_Compose m2 m1).
-    abstract apply (CommaSpecializedCategory_Associativity (S := S) (T := T) m1 m2 m3).
+  Lemma CommaCategory_Morphism_eq αβf α'β'f' αβf2 α'β'f'2
+        (M : CommaCategory_Morphism αβf α'β'f')
+        (N : CommaCategory_Morphism αβf2 α'β'f'2) :
+    αβf = αβf2
+    -> α'β'f' = α'β'f'2
+    -> proj1_sig M == proj1_sig N
+    -> M == N.
+    clear; intros.
+    subst.
+    JMeq_eq.
+    simpl_eq'.
+    assumption.
   Qed.
 
-  Lemma CommaCategory_LeftIdentity (a b : CommaCategory_Object) (f : CommaCategory_Morphism a b) :
-    CommaCategory_Compose (CommaCategory_Identity b) f = f.
+  Definition CommaCategory_Compose s d d'
+             (gh : CommaCategory_Morphism d d') (g'h' : CommaCategory_Morphism s d)
+  : CommaCategory_Morphism s d'.
+    exists (Compose (C := A * B) (proj1_sig gh) (proj1_sig g'h')).
+    simpl in *.
+    abstract (
+        hnf in *;
+        destruct_head sig;
+        destruct_head sigT;
+        destruct_head prod;
+        repeat (reflexivity
+                  || rewrite FCompositionOf
+                  || try_associativity rewrite_hyp)
+      ).
+  Defined.
+
+  Global Arguments CommaCategory_Compose _ _ _ _ _ / .
+
+  Definition CommaCategory_Identity o : CommaCategory_Morphism o o.
+    exists (Identity (C := A * B) (projT1 o)).
+    abstract (
+        simpl; autorewrite with category; reflexivity
+      ).
+  Defined.
+
+  Global Arguments CommaCategory_Identity _ / .
+
+  Local Ltac comma_t :=
+    repeat (
+        let H := fresh in intro H; destruct H as [ ]
+      );
+    destruct_hypotheses;
+    simpl in *;
+    simpl_eq;
+    autorewrite with category;
+    f_equal;
+    try reflexivity.
+
+  Lemma CommaCategory_Associativity : forall o1 o2 o3 o4 (m1 : CommaCategory_Morphism o1 o2) (m2 : CommaCategory_Morphism o2 o3) (m3 : CommaCategory_Morphism o3 o4),
+                                        CommaCategory_Compose (CommaCategory_Compose m3 m2) m1 =
+                                        CommaCategory_Compose m3 (CommaCategory_Compose m2 m1).
   Proof.
-    abstract apply (CommaSpecializedCategory_LeftIdentity (S := S) (T := T) f).
+    intros.
+    abstract (
+        apply sig_eq;
+        simpl;
+        repeat rewrite Associativity;
+        reflexivity
+      ).
   Qed.
 
-  Lemma CommaCategory_RightIdentity (a b : CommaCategory_Object) (f : CommaCategory_Morphism a b) :
-    CommaCategory_Compose f (CommaCategory_Identity a) = f.
+  Lemma CommaCategory_LeftIdentity : forall a b (f : CommaCategory_Morphism a b),
+                                       CommaCategory_Compose (CommaCategory_Identity b) f = f.
   Proof.
-    abstract apply (CommaSpecializedCategory_RightIdentity (S := S) (T := T) f).
+    abstract comma_t.
   Qed.
 
-  Definition CommaCategory : Category.
-    refine (@Build_SpecializedCategory CommaCategory_Object CommaCategory_Morphism
-      CommaCategory_Identity
-      CommaCategory_Compose
-      _ _ _
-    );
-    subst_body;
-    abstract (apply CommaCategory_Associativity || apply CommaCategory_LeftIdentity || apply CommaCategory_RightIdentity).
-  Defined.
+  Lemma CommaCategory_RightIdentity : forall a b (f : CommaCategory_Morphism a b),
+                                        CommaCategory_Compose f (CommaCategory_Identity a) = f.
+  Proof.
+    abstract comma_t.
+  Qed.
+
+  Definition CommaCategory : Category
+    := @Build_Category CommaCategory_Object
+                       CommaCategory_Morphism
+                       CommaCategory_Identity
+                       CommaCategory_Compose
+                       CommaCategory_Associativity
+                       CommaCategory_LeftIdentity
+                       CommaCategory_RightIdentity.
 End CommaCategory.
 
-Hint Unfold CommaCategory_Compose CommaCategory_Identity CommaCategory_Morphism CommaCategory_Object : category.
-
-Arguments CommaCategory [A B C] S T.
-
-Local Notation "S ↓ T" := (CommaCategory S T).
+Hint Unfold CommaCategory_Compose CommaCategory_Identity : category.
 
 Section SliceCategory.
-  Variable A : Category.
-  Variable C : Category.
+  Variables A C : Category.
   Variable a : C.
   Variable S : Functor A C.
-  Let B := TerminalCategory.
 
-  Definition SliceCategory_Functor : Functor B C.
-    refine {| ObjectOf := (fun _ => a);
-      MorphismOf := (fun _ _ _ => Identity a)
-    |}; abstract (intros; auto with morphism).
-  Defined.
+  Definition SliceCategory := CommaCategory S (FunctorFromTerminal C a).
+  Definition CosliceCategory := CommaCategory (FunctorFromTerminal C a) S.
 
-  Definition SliceCategory := CommaCategory S SliceCategory_Functor.
-  Definition CosliceCategory := CommaCategory SliceCategory_Functor S.
+(* [x ↓ F] is a coslice category; [F ↓ x] is a slice category; [x ↓ F] deals with morphisms [x -> F y]; [F ↓ x] has morphisms [F y -> x] *)
 End SliceCategory.
 
 Section SliceCategoryOver.
@@ -151,3 +173,22 @@
 
   Definition ArrowCategory := CommaCategory (IdentityFunctor C) (IdentityFunctor C).
 End ArrowCategory.
+
+Notation "C / a" := (@SliceCategoryOver C a) : category_scope.
+Notation "a \ C" := (@CosliceCategoryOver C a) : category_scope.
+
+Definition CC_Functor' C D := Functor C D.
+Coercion CC_FunctorFromTerminal' (C : Category) (x : C) : CC_Functor' TerminalCategory C := FunctorFromTerminal C x.
+Arguments CC_Functor' / .
+Arguments CC_FunctorFromTerminal' / .
+
+(* Set some notations for printing *)
+Notation "x ↓ F" := (CosliceCategory x F) : category_scope.
+Notation "F ↓ x" := (SliceCategory x F) : category_scope.
+Notation "S ↓ T" := (CommaCategory S T) : category_scope.
+(* set the notation for parsing *)
+Notation "S ↓ T" := (CommaCategory (S : CC_Functor' _ _)
+                                   (T : CC_Functor' _ _)) : category_scope.
+(*Set Printing All.
+Check (fun `(C : @Category objC) `(D : @Category objD) `(E : @Category objE) (S : Functor C D) (T : Functor E D) => (S ↓ T)%category).
+Check (fun `(D : @Category objD) `(E : @Category objE) (S : Functor E D) (x : D) => (x ↓ S)%category).*)
diff -r 4eb6407c6ca3 CommaCategoryFunctorProperties.v
--- a/CommaCategoryFunctorProperties.v	Fri May 24 20:09:37 2013 -0400
+++ b/CommaCategoryFunctorProperties.v	Fri Jun 21 15:07:08 2013 -0400
@@ -14,7 +14,7 @@
 
 Local Ltac slice_t :=
   repeat match goal with
-           | [ |- @eq (SpecializedFunctor _ _) _ _ ] => apply Functor_eq; intros
+           | [ |- @eq (Functor _ _) _ _ ] => apply Functor_eq; intros
            | [ |- @JMeq (sig _) _ (sig _) _ ] => apply sig_JMeq
            | [ |- (_ -> _) = (_ -> _) ] => apply f_type_equal
            | _ => progress (intros; simpl in *; trivial; subst_body)
@@ -28,16 +28,16 @@
              apply (@f_equal_JMeq _ _ _ _ x y f g)
            | _ =>
              progress (
-               destruct_type @CommaSpecializedCategory_Object;
-               destruct_type @CommaSpecializedCategory_Morphism;
+               destruct_type @CommaCategory_Object;
+               destruct_type @CommaCategory_Morphism;
                destruct_sig
              )
          end.
 
 Section FCompositionOf.
-  Context `(A : @SpecializedCategory objA).
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
+  Variable A : Category.
+  Variable B : Category.
+  Variable C : Category.
 
   Lemma CommaCategoryInducedFunctor_FCompositionOf s d d'
         (m1 : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d)
diff -r 4eb6407c6ca3 CommaCategoryInducedFunctors.v
--- a/CommaCategoryInducedFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/CommaCategoryInducedFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq.
-Require Export CommaCategoryProjection SmallCat Duals FunctorCategory.
+Require Export CommaCategoryProjection Cat Duals FunctorCategory.
 Require Import Common Notations DefinitionSimplification DiscreteCategory FEqualDep DefinitionSimplification ExponentialLaws FunctorProduct ProductLaws CanonicalStructureSimplification.
 
 Set Implicit Arguments.
@@ -32,8 +32,8 @@
   repeat intro; simpl in *;
   repeat quick_step;
   simpl in *;
-  destruct_head_hnf @CommaSpecializedCategory_Morphism;
-  destruct_head_hnf @CommaSpecializedCategory_Object;
+  destruct_head_hnf @CommaCategory_Morphism;
+  destruct_head_hnf @CommaCategory_Object;
   destruct_sig;
   subst_body;
   unfold Object in *;
@@ -76,30 +76,30 @@
   repeat induced_step.
 
 Section CommaCategoryInducedFunctor.
-  Context `(A : @SpecializedCategory objA).
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
+  Variable A : Category.
+  Variable B : Category.
+  Variable C : Category.
 
   Let CommaCategoryInducedFunctor_ObjectOf s d (m : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d)
       (x : fst s ↓ snd s) : (fst d ↓ snd d)
     := existT _
               (projT1 x)
               (Compose ((snd m) (snd (projT1 x))) (Compose (projT2 x) ((fst m) (fst (projT1 x))))) :
-         CommaSpecializedCategory_ObjectT (fst d) (snd d).
+         CommaCategory_ObjectT (fst d) (snd d).
 
   Let CommaCategoryInducedFunctor_MorphismOf s d m s0 d0 (m0 : Morphism (fst s ↓ snd s) s0 d0) :
     Morphism (fst d ↓ snd d)
              (@CommaCategoryInducedFunctor_ObjectOf s d m s0)
              (@CommaCategoryInducedFunctor_ObjectOf s d m d0).
-    refine (_ : CommaSpecializedCategory_MorphismT _ _); subst_body; simpl in *.
+    refine (_ : CommaCategory_MorphismT _ _); subst_body; simpl in *.
     exists (projT1 m0);
       simpl in *; clear.
     abstract induced_anihilate.
   Defined.
 
   Let CommaCategoryInducedFunctor' s d (m : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d) :
-    SpecializedFunctor (fst s ↓ snd s) (fst d ↓ snd d).
-    refine (Build_SpecializedFunctor (fst s ↓ snd s) (fst d ↓ snd d)
+    Functor (fst s ↓ snd s) (fst d ↓ snd d).
+    refine (Build_Functor (fst s ↓ snd s) (fst d ↓ snd d)
                                      (@CommaCategoryInducedFunctor_ObjectOf s d m)
                                      (@CommaCategoryInducedFunctor_MorphismOf s d m)
                                      _
@@ -111,44 +111,44 @@
   Defined.
 
   Let CommaCategoryInducedFunctor'' s d (m : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d) :
-    SpecializedFunctor (fst s ↓ snd s) (fst d ↓ snd d).
+    Functor (fst s ↓ snd s) (fst d ↓ snd d).
     simpl_definition_by_tac_and_exact (@CommaCategoryInducedFunctor' s d m) ltac:(subst_body; eta_red).
   Defined.
 
   Definition CommaCategoryInducedFunctor s d (m : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d) :
-    SpecializedFunctor (fst s ↓ snd s) (fst d ↓ snd d)
+    Functor (fst s ↓ snd s) (fst d ↓ snd d)
     := Eval cbv beta iota zeta delta [CommaCategoryInducedFunctor''] in @CommaCategoryInducedFunctor'' s d m.
 End CommaCategoryInducedFunctor.
 
 Section SliceCategoryInducedFunctor.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Section Slice_Coslice.
-    Context `(D : @SpecializedCategory objD).
+    Variable D : Category.
 
     (* TODO(jgross): See if this can be recast as an exponential law functor about how Cat ^ 1 is isomorphic to Cat, or something *)
     Definition SliceCategoryInducedFunctor_NT s d (m : Morphism D s d) :
-      SpecializedNaturalTransformation (FunctorFromTerminal D s)
+      NaturalTransformation (FunctorFromTerminal D s)
                                        (FunctorFromTerminal D d).
       exists (fun _ : unit => m).
       simpl; intros; clear;
       abstract (autorewrite with category; reflexivity).
     Defined.
 
-    Variable F : SpecializedFunctor C D.
+    Variable F : Functor C D.
     Variable a : D.
 
     Section Slice.
-      Local Notation "F ↓ A" := (SliceSpecializedCategory A F) : category_scope.
+      Local Notation "F ↓ A" := (SliceCategory A F) : category_scope.
 
-      Let SliceCategoryInducedFunctor' F' a' (m : Morphism D a a') (T : SpecializedNaturalTransformation F' F) :
-        SpecializedFunctor (F ↓ a) (F' ↓ a').
+      Let SliceCategoryInducedFunctor' F' a' (m : Morphism D a a') (T : NaturalTransformation F' F) :
+        Functor (F ↓ a) (F' ↓ a').
         functor_simpl_abstract_trailing_props (CommaCategoryInducedFunctor (s := (F, FunctorFromTerminal D a))
                                                                        (d := (F', FunctorFromTerminal D a'))
                                                                        (T, @SliceCategoryInducedFunctor_NT a a' m)).
       Defined.
-      Definition SliceCategoryInducedFunctor F' a' (m : Morphism D a a') (T : SpecializedNaturalTransformation F' F) :
-        SpecializedFunctor (F ↓ a) (F' ↓ a')
+      Definition SliceCategoryInducedFunctor F' a' (m : Morphism D a a') (T : NaturalTransformation F' F) :
+        Functor (F ↓ a) (F' ↓ a')
         := Eval hnf in @SliceCategoryInducedFunctor' F' a' m T.
 
       Definition SliceCategoryNTInducedFunctor F' T := @SliceCategoryInducedFunctor F' a (Identity _) T.
@@ -156,16 +156,16 @@
     End Slice.
 
     Section Coslice.
-      Local Notation "A ↓ F" := (CosliceSpecializedCategory A F) : category_scope.
+      Local Notation "A ↓ F" := (CosliceCategory A F) : category_scope.
 
-      Let CosliceCategoryInducedFunctor' F' a' (m : Morphism D a' a) (T : SpecializedNaturalTransformation F F') :
-        SpecializedFunctor (a ↓ F) (a' ↓ F').
+      Let CosliceCategoryInducedFunctor' F' a' (m : Morphism D a' a) (T : NaturalTransformation F F') :
+        Functor (a ↓ F) (a' ↓ F').
         functor_simpl_abstract_trailing_props (CommaCategoryInducedFunctor (s := (FunctorFromTerminal D a, F))
                                                                        (d := (FunctorFromTerminal D a', F'))
                                                                        (@SliceCategoryInducedFunctor_NT a' a m, T)).
       Defined.
-      Definition CosliceCategoryInducedFunctor F' a' (m : Morphism D a' a) (T : SpecializedNaturalTransformation F F') :
-        SpecializedFunctor (a ↓ F) (a' ↓ F')
+      Definition CosliceCategoryInducedFunctor F' a' (m : Morphism D a' a) (T : NaturalTransformation F F') :
+        Functor (a ↓ F) (a' ↓ F')
         := Eval hnf in @CosliceCategoryInducedFunctor' F' a' m T.
 
       Definition CosliceCategoryNTInducedFunctor F' T := @CosliceCategoryInducedFunctor F' a (Identity _) T.
@@ -173,52 +173,52 @@
     End Coslice.
   End Slice_Coslice.
 
-  Definition SliceCategoryOverInducedFunctor a a' (m : Morphism C a a') : SpecializedFunctor (C / a) (C / a')
+  Definition SliceCategoryOverInducedFunctor a a' (m : Morphism C a a') : Functor (C / a) (C / a')
     := Eval hnf in SliceCategoryMorphismInducedFunctor _ _ _ m.
-  Definition CosliceCategoryOverInducedFunctor a a' (m : Morphism C a' a) : SpecializedFunctor (a \ C) (a' \ C)
+  Definition CosliceCategoryOverInducedFunctor a a' (m : Morphism C a' a) : Functor (a \ C) (a' \ C)
     := Eval hnf in CosliceCategoryMorphismInducedFunctor _ _ _ m.
 End SliceCategoryInducedFunctor.
 
-Section LocallySmallCatOverInducedFunctor.
+Section LocallyCatOverInducedFunctor.
   (* Let's do this from scratch so we get better polymorphism *)
-  Let C := LocallySmallCat.
-  (*Local Opaque LocallySmallCat.*)
+  Let C := LocallyCat.
+  (*Local Opaque LocallyCat.*)
 
-  Let LocallySmallCatOverInducedFunctor_ObjectOf a a' (m : Morphism C a a') : C / a -> C / a'.
+  Let LocallyCatOverInducedFunctor_ObjectOf a a' (m : Morphism C a a') : C / a -> C / a'.
     refine (fun x => existT (fun αβ : _ * unit => Morphism C (fst αβ) a')
                             (projT1 x)
                             _ :
-                       CommaSpecializedCategory_ObjectT (IdentityFunctor C) (FunctorFromTerminal C a')).
+                       CommaCategory_ObjectT (IdentityFunctor C) (FunctorFromTerminal C a')).
     functor_simpl_abstract_trailing_props (Compose m (projT2 x)).
   Defined.
-  Let OverLocallySmallCatInducedFunctor_ObjectOf a a' (m : Morphism C a' a) : a \ C -> a' \ C.
+  Let OverLocallyCatInducedFunctor_ObjectOf a a' (m : Morphism C a' a) : a \ C -> a' \ C.
     refine (fun x => existT (fun αβ : unit * _ => Morphism C a' (snd αβ))
                             (projT1 x)
                             _ :
-                       CommaSpecializedCategory_ObjectT (FunctorFromTerminal C a') (IdentityFunctor C)).
+                       CommaCategory_ObjectT (FunctorFromTerminal C a') (IdentityFunctor C)).
     functor_simpl_abstract_trailing_props (Compose (projT2 x) m).
   Defined.
 
-  Let LocallySmallCatOverInducedFunctor_MorphismOf a a' (m : Morphism C a a') s d (m0 : Morphism (C / a) s d) :
-    Morphism (C / a') (@LocallySmallCatOverInducedFunctor_ObjectOf _ _ m s) (@LocallySmallCatOverInducedFunctor_ObjectOf _ _ m d).
-    refine (_ : CommaSpecializedCategory_MorphismT _ _).
+  Let LocallyCatOverInducedFunctor_MorphismOf a a' (m : Morphism C a a') s d (m0 : Morphism (C / a) s d) :
+    Morphism (C / a') (@LocallyCatOverInducedFunctor_ObjectOf _ _ m s) (@LocallyCatOverInducedFunctor_ObjectOf _ _ m d).
+    refine (_ : CommaCategory_MorphismT _ _).
     exists (projT1 m0).
       subst_body; simpl in *; clear.
     abstract anihilate.
   Defined.
-  Let OverLocallySmallCatInducedFunctor_MorphismOf a a' (m : Morphism C a' a) s d (m0 : Morphism (a \ C) s d) :
-    Morphism (a' \ C) (@OverLocallySmallCatInducedFunctor_ObjectOf _ _ m s) (@OverLocallySmallCatInducedFunctor_ObjectOf _ _ m d).
-    refine (_ : CommaSpecializedCategory_MorphismT _ _).
+  Let OverLocallyCatInducedFunctor_MorphismOf a a' (m : Morphism C a' a) s d (m0 : Morphism (a \ C) s d) :
+    Morphism (a' \ C) (@OverLocallyCatInducedFunctor_ObjectOf _ _ m s) (@OverLocallyCatInducedFunctor_ObjectOf _ _ m d).
+    refine (_ : CommaCategory_MorphismT _ _).
     exists (projT1 m0);
       subst_body; simpl in *; clear.
     abstract anihilate.
   Defined.
 
-  Local Opaque LocallySmallCat.
-  Let LocallySmallCatOverInducedFunctor'' a a' (m : Morphism C a a') : SpecializedFunctor (C / a) (C / a').
-    refine (Build_SpecializedFunctor (C / a) (C / a')
-                                     (@LocallySmallCatOverInducedFunctor_ObjectOf _ _ m)
-                                     (@LocallySmallCatOverInducedFunctor_MorphismOf _ _ m)
+  Local Opaque LocallyCat.
+  Let LocallyCatOverInducedFunctor'' a a' (m : Morphism C a a') : Functor (C / a) (C / a').
+    refine (Build_Functor (C / a) (C / a')
+                                     (@LocallyCatOverInducedFunctor_ObjectOf _ _ m)
+                                     (@LocallyCatOverInducedFunctor_MorphismOf _ _ m)
                                      _
                                      _
            );
@@ -226,10 +226,10 @@
     (* admit. *)
     abstract anihilate.
   Defined.
-  Let OverLocallySmallCatInducedFunctor'' a a' (m : Morphism C a' a) : SpecializedFunctor (a \ C) (a' \ C).
-    refine (Build_SpecializedFunctor (a \ C) (a' \ C)
-                                     (@OverLocallySmallCatInducedFunctor_ObjectOf _ _ m)
-                                     (@OverLocallySmallCatInducedFunctor_MorphismOf _ _ m)
+  Let OverLocallyCatInducedFunctor'' a a' (m : Morphism C a' a) : Functor (a \ C) (a' \ C).
+    refine (Build_Functor (a \ C) (a' \ C)
+                                     (@OverLocallyCatInducedFunctor_ObjectOf _ _ m)
+                                     (@OverLocallyCatInducedFunctor_MorphismOf _ _ m)
                                      _
                                      _
            );
@@ -237,16 +237,16 @@
     (* admit. *) abstract anihilate.
   Defined.
 
-  Let LocallySmallCatOverInducedFunctor''' a a' (m : Morphism C a a') : SpecializedFunctor (C / a) (C / a').
-    simpl_definition_by_tac_and_exact (@LocallySmallCatOverInducedFunctor'' _ _ m) ltac:(subst_body; eta_red).
+  Let LocallyCatOverInducedFunctor''' a a' (m : Morphism C a a') : Functor (C / a) (C / a').
+    simpl_definition_by_tac_and_exact (@LocallyCatOverInducedFunctor'' _ _ m) ltac:(subst_body; eta_red).
   Defined.
-  Let OverLocallySmallCatInducedFunctor''' a a' (m : Morphism C a' a) : SpecializedFunctor (a \ C) (a' \ C).
-    simpl_definition_by_tac_and_exact (@OverLocallySmallCatInducedFunctor'' _ _ m) ltac:(subst_body; eta_red).
+  Let OverLocallyCatInducedFunctor''' a a' (m : Morphism C a' a) : Functor (a \ C) (a' \ C).
+    simpl_definition_by_tac_and_exact (@OverLocallyCatInducedFunctor'' _ _ m) ltac:(subst_body; eta_red).
   Defined.
 
-  Definition LocallySmallCatOverInducedFunctor a a' (m : Morphism C a a') : SpecializedFunctor (C / a) (C / a')
-    := Eval cbv beta iota zeta delta [LocallySmallCatOverInducedFunctor'''] in @LocallySmallCatOverInducedFunctor''' _ _ m.
-  Definition OverLocallySmallCatInducedFunctor a a' (m : Morphism C a' a) : SpecializedFunctor (a \ C) (a' \ C)
-    := Eval cbv beta iota zeta delta [OverLocallySmallCatInducedFunctor'''] in @OverLocallySmallCatInducedFunctor''' _ _ m.
-  Local Transparent LocallySmallCat.
-End LocallySmallCatOverInducedFunctor.
+  Definition LocallyCatOverInducedFunctor a a' (m : Morphism C a a') : Functor (C / a) (C / a')
+    := Eval cbv beta iota zeta delta [LocallyCatOverInducedFunctor'''] in @LocallyCatOverInducedFunctor''' _ _ m.
+  Definition OverLocallyCatInducedFunctor a a' (m : Morphism C a' a) : Functor (a \ C) (a' \ C)
+    := Eval cbv beta iota zeta delta [OverLocallyCatInducedFunctor'''] in @OverLocallyCatInducedFunctor''' _ _ m.
+  Local Transparent LocallyCat.
+End LocallyCatOverInducedFunctor.
diff -r 4eb6407c6ca3 CommaCategoryProjection.v
--- a/CommaCategoryProjection.v	Fri May 24 20:09:37 2013 -0400
+++ b/CommaCategoryProjection.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCommaCategory ProductCategory.
+Require Export CommaCategory ProductCategory.
 Require Import Common Notations.
 
 Set Implicit Arguments.
@@ -12,14 +12,14 @@
 Local Open Scope category_scope.
 
 Section CommaCategory.
-  Context `(A : @SpecializedCategory objA).
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
-  Variable S : SpecializedFunctor A C.
-  Variable T : SpecializedFunctor B C.
+  Variable A : Category.
+  Variable B : Category.
+  Variable C : Category.
+  Variable S : Functor A C.
+  Variable T : Functor B C.
 
-  Definition CommaCategoryProjection : SpecializedFunctor (S ↓ T) (A * B)
-    := Build_SpecializedFunctor (S ↓ T) (A * B)
+  Definition CommaCategoryProjection : Functor (S ↓ T) (A * B)
+    := Build_Functor (S ↓ T) (A * B)
                                 (@projT1 _ _)
                                 (fun _ _ m => proj1_sig m)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -27,34 +27,34 @@
 End CommaCategory.
 
 Section SliceCategory.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
 
   Local Arguments ComposeFunctors' / .
 
-  Definition ArrowCategoryProjection : SpecializedFunctor (ArrowSpecializedCategory A) A
+  Definition ArrowCategoryProjection : Functor (ArrowCategory A) A
     := Eval simpl in ComposeFunctors' fst_Functor (CommaCategoryProjection _ (IdentityFunctor A)).
 
-  Definition SliceCategoryOverProjection (a : A) : SpecializedFunctor (A / a) A
+  Definition SliceCategoryOverProjection (a : A) : Functor (A / a) A
     := Eval simpl in ComposeFunctors' fst_Functor (CommaCategoryProjection (IdentityFunctor A) _).
 
-  Definition CosliceCategoryOverProjection (a : A) : SpecializedFunctor (a \ A) A
+  Definition CosliceCategoryOverProjection (a : A) : Functor (a \ A) A
     := ComposeFunctors' snd_Functor (CommaCategoryProjection _ (IdentityFunctor A)).
 
   Section Slice_Coslice.
-    Context `(C : @SpecializedCategory objC).
+    Variable C : Category.
     Variable a : C.
-    Variable S : SpecializedFunctor A C.
+    Variable S : Functor A C.
 
     Section Slice.
-      Definition SliceCategoryProjection : SpecializedFunctor (S ↓ a) A
+      Definition SliceCategoryProjection : Functor (S ↓ a) A
         := Eval simpl in ComposeFunctors' fst_Functor (CommaCategoryProjection S (FunctorFromTerminal C a)).
     End Slice.
 
     Section Coslice.
-      Definition CosliceCategoryProjection : SpecializedFunctor (a ↓ S) A
+      Definition CosliceCategoryProjection : Functor (a ↓ S) A
         := Eval simpl in ComposeFunctors' snd_Functor (CommaCategoryProjection (FunctorFromTerminal C a) S).
       Check CosliceCategoryProjection.
-      Eval simpl in SpecializedFunctor (a ↓ S) A.
+      Eval simpl in Functor (a ↓ S) A.
     End Coslice.
   End Slice_Coslice.
 End SliceCategory.
diff -r 4eb6407c6ca3 CommaCategoryProjectionFunctors.v
--- a/CommaCategoryProjectionFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/CommaCategoryProjectionFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -32,8 +32,8 @@
   repeat intro; simpl in *;
   repeat quick_step;
   simpl in *;
-  destruct_head_hnf @CommaSpecializedCategory_Morphism;
-  destruct_head_hnf @CommaSpecializedCategory_Object;
+  destruct_head_hnf @CommaCategory_Morphism;
+  destruct_head_hnf @CommaCategory_Object;
   destruct_sig;
   subst_body;
   unfold Object in *;
@@ -76,23 +76,23 @@
   repeat induced_step.
 
 Section CommaCategoryProjectionFunctor.
-  Context `(A : LocallySmallSpecializedCategory objA).
-  Context `(B : LocallySmallSpecializedCategory objB).
-  Context `(C : SpecializedCategory objC).
+  Context `(A : LocallyCategory objA).
+  Context `(B : LocallyCategory objB).
+  Variable C : Category.
 
   Definition CommaCategoryProjectionFunctor_ObjectOf (ST : (OppositeCategory (C ^ A)) * (C ^ B)) :
-    LocallySmallCat / (A * B : LocallySmallSpecializedCategory _)
+    LocallyCat / (A * B : LocallyCategory _)
     := let S := (fst ST) in
        let T := (snd ST) in
        existT _
-              ((S ↓ T : LocallySmallSpecializedCategory _) : LocallySmallCategory, tt)
+              ((S ↓ T : LocallyCategory _) : LocallyCategory, tt)
               (CommaCategoryProjection S T) :
-         CommaSpecializedCategory_ObjectT (IdentityFunctor _)
-                                          (FunctorFromTerminal LocallySmallCat
-                                                               (A * B : LocallySmallSpecializedCategory _)).
+         CommaCategory_ObjectT (IdentityFunctor _)
+                                          (FunctorFromTerminal LocallyCat
+                                                               (A * B : LocallyCategory _)).
 
   Definition CommaCategoryProjectionFunctor_MorphismOf s d (m : Morphism ((OppositeCategory (C ^ A)) * (C ^ B)) s d) :
-    Morphism (LocallySmallCat / (A * B : LocallySmallSpecializedCategory _))
+    Morphism (LocallyCat / (A * B : LocallyCategory _))
              (CommaCategoryProjectionFunctor_ObjectOf s)
              (CommaCategoryProjectionFunctor_ObjectOf d).
     hnf in *; constructor; simpl in *.
@@ -100,7 +100,7 @@
     abstract (
         simpl;
         functor_eq;
-        destruct_head_hnf @CommaSpecializedCategory_Morphism;
+        destruct_head_hnf @CommaCategory_Morphism;
         destruct_sig;
         reflexivity
       ).
@@ -119,10 +119,10 @@
   Qed.
 
   Definition CommaCategoryProjectionFunctor :
-    SpecializedFunctor ((OppositeCategory (C ^ A)) * (C ^ B))
-                       (LocallySmallCat / (A * B : LocallySmallSpecializedCategory _)).
-    refine (Build_SpecializedFunctor ((OppositeCategory (C ^ A)) * (C ^ B))
-                                     (LocallySmallCat / (A * B : LocallySmallSpecializedCategory _))
+    Functor ((OppositeCategory (C ^ A)) * (C ^ B))
+                       (LocallyCat / (A * B : LocallyCategory _)).
+    refine (Build_Functor ((OppositeCategory (C ^ A)) * (C ^ B))
+                                     (LocallyCat / (A * B : LocallyCategory _))
                                      CommaCategoryProjectionFunctor_ObjectOf
                                      CommaCategoryProjectionFunctor_MorphismOf
                                      _
@@ -134,34 +134,34 @@
 End CommaCategoryProjectionFunctor.
 
 Section SliceCategoryProjectionFunctor.
-  Context `(C : LocallySmallSpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
+  Context `(C : LocallyCategory objC).
+  Variable D : Category.
 
   Local Arguments ExponentialLaw4Functor_Inverse_ObjectOf_ObjectOf / .
   Local Arguments ComposeFunctors / .
-  Local Arguments LocallySmallCatOverInducedFunctor / .
+  Local Arguments LocallyCatOverInducedFunctor / .
   (*Local Arguments ProductLaw1Functor / . *)
   Local Arguments CommaCategoryProjectionFunctor / .
   Local Arguments SwapFunctor / .
   Local Arguments ExponentialLaw1Functor_Inverse / .
   Local Arguments IdentityFunctor / .
 (*
-  Let ArrowCategoryProjection' : SpecializedFunctor (ArrowSpecializedCategory A) A
+  Let ArrowCategoryProjection' : Functor (ArrowCategory A) A
     := ComposeFunctors fst_Functor (CommaCategoryProjection _ (IdentityFunctor A)).
-  Let ArrowCategoryProjection'' : SpecializedFunctor (ArrowSpecializedCategory A) A. functor_simpl_abstract_trailing_props ArrowCategoryProjection'. Defined.
-  Definition ArrowCategoryProjection : SpecializedFunctor (ArrowSpecializedCategory A) A := Eval hnf in ArrowCategoryProjection''.
+  Let ArrowCategoryProjection'' : Functor (ArrowCategory A) A. functor_simpl_abstract_trailing_props ArrowCategoryProjection'. Defined.
+  Definition ArrowCategoryProjection : Functor (ArrowCategory A) A := Eval hnf in ArrowCategoryProjection''.
 *)
 
   Definition SliceCategoryProjectionFunctor_pre_pre'
-    := Eval hnf in (LocallySmallCatOverInducedFunctor (ProductLaw1Functor C : Morphism LocallySmallCat (C * 1 : LocallySmallSpecializedCategory _) C)).
+    := Eval hnf in (LocallyCatOverInducedFunctor (ProductLaw1Functor C : Morphism LocallyCat (C * 1 : LocallyCategory _) C)).
 
-  Definition SliceCategoryProjectionFunctor_pre_pre : SpecializedFunctor (LocallySmallCat / (C * 1 : LocallySmallSpecializedCategory _)) (LocallySmallCat / C).
+  Definition SliceCategoryProjectionFunctor_pre_pre : Functor (LocallyCat / (C * 1 : LocallyCategory _)) (LocallyCat / C).
     functor_abstract_trailing_props SliceCategoryProjectionFunctor_pre_pre'.
   Defined.
 
   Arguments SliceCategoryProjectionFunctor_pre_pre / .
 
-(*  Arguments Build_CommaSpecializedCategory_Object' / .
+(*  Arguments Build_CommaCategory_Object' / .
 
   Eval simpl in SliceCategoryProjectionFunctor_pre_pre'.
 
@@ -177,43 +177,43 @@
   Local Ltac refine_right_compose_functor_by_abstract F :=
     let H := fresh in pose_functor_by_abstract H F; refine (ComposeFunctors _ H); clear H.*)
 
-  Definition SliceCategoryProjectionFunctor_pre' : ((LocallySmallCat / C) ^ (D * (OppositeCategory (D ^ C)))).
+  Definition SliceCategoryProjectionFunctor_pre' : ((LocallyCat / C) ^ (D * (OppositeCategory (D ^ C)))).
     refine (ComposeFunctors _ ((ExponentialLaw1Functor_Inverse D) * IdentityFunctor (OppositeCategory (D ^ C)))).
     refine (ComposeFunctors _ (SwapFunctor _ _)).
-    refine (ComposeFunctors _ (CommaCategoryProjectionFunctor (C : LocallySmallSpecializedCategory _) (1 : LocallySmallSpecializedCategory _) D)).
+    refine (ComposeFunctors _ (CommaCategoryProjectionFunctor (C : LocallyCategory _) (1 : LocallyCategory _) D)).
     (*
     refine_right_compose_functor_by_abstract ((ExponentialLaw1Functor_Inverse D) * IdentityFunctor (OppositeCategory (D ^ C)))%functor.
     refine_right_compose_functor_by_abstract (SwapFunctor (D ^ 1)%functor (OppositeCategory (D ^ C)%functor)).
-    refine_right_compose_functor_by_abstract (CommaCategoryProjectionFunctor (C : LocallySmallSpecializedCategory _) (1 : LocallySmallSpecializedCategory _) D).*)
+    refine_right_compose_functor_by_abstract (CommaCategoryProjectionFunctor (C : LocallyCategory _) (1 : LocallyCategory _) D).*)
     let F := (eval hnf in SliceCategoryProjectionFunctor_pre_pre) in
     exact F.
-(* (LocallySmallCatOverInducedFunctor (ProductLaw1Functor C : Morphism LocallySmallCat (C * 1 : LocallySmallSpecializedCategory _) (C : LocallySmallSpecializedCategory _))). *)
+(* (LocallyCatOverInducedFunctor (ProductLaw1Functor C : Morphism LocallyCat (C * 1 : LocallyCategory _) (C : LocallyCategory _))). *)
   Defined.
 
-  Definition SliceCategoryProjectionFunctor_pre'' : ((LocallySmallCat / C) ^ (D * (OppositeCategory (D ^ C)))).
+  Definition SliceCategoryProjectionFunctor_pre'' : ((LocallyCat / C) ^ (D * (OppositeCategory (D ^ C)))).
     functor_abstract_trailing_props SliceCategoryProjectionFunctor_pre'.
   Defined.
 
   Definition SliceCategoryProjectionFunctor_pre := Eval hnf in SliceCategoryProjectionFunctor_pre''.
 
-  Definition SliceCategoryProjectionFunctor' : (((LocallySmallCat / C) ^ D) ^ (OppositeCategory (D ^ C))).
+  Definition SliceCategoryProjectionFunctor' : (((LocallyCat / C) ^ D) ^ (OppositeCategory (D ^ C))).
     refine ((ExponentialLaw4Functor_Inverse _ _ _) _).
     let F := (eval hnf in SliceCategoryProjectionFunctor_pre) in
     exact F.
   Defined.
 
-  Definition SliceCategoryProjectionFunctor'' : (((LocallySmallCat / C) ^ D) ^ (OppositeCategory (D ^ C))).
+  Definition SliceCategoryProjectionFunctor'' : (((LocallyCat / C) ^ D) ^ (OppositeCategory (D ^ C))).
     functor_abstract_trailing_props SliceCategoryProjectionFunctor'.
   Defined.
 
-  Definition SliceCategoryProjectionFunctor : ((LocallySmallCat / C) ^ D) ^ (OppositeCategory (D ^ C))
+  Definition SliceCategoryProjectionFunctor : ((LocallyCat / C) ^ D) ^ (OppositeCategory (D ^ C))
     := Eval cbv beta iota zeta delta [SliceCategoryProjectionFunctor''] in SliceCategoryProjectionFunctor''.
 
-  Definition CosliceCategoryProjectionFunctor : ((LocallySmallCat / C) ^ (OppositeCategory D)) ^ (D ^ C).
+  Definition CosliceCategoryProjectionFunctor : ((LocallyCat / C) ^ (OppositeCategory D)) ^ (D ^ C).
     refine ((ExponentialLaw4Functor_Inverse _ _ _) _).
     refine (ComposeFunctors _ ((OppositeFunctor (ExponentialLaw1Functor_Inverse D)) * IdentityFunctor (D ^ C))).
-    refine (ComposeFunctors _ (CommaCategoryProjectionFunctor (1 : LocallySmallSpecializedCategory _) (C : LocallySmallSpecializedCategory _) D)).
-    refine (LocallySmallCatOverInducedFunctor _).
+    refine (ComposeFunctors _ (CommaCategoryProjectionFunctor (1 : LocallyCategory _) (C : LocallyCategory _) D)).
+    refine (LocallyCatOverInducedFunctor _).
     refine (ComposeFunctors _ (SwapFunctor _ _)).
     exact (ProductLaw1Functor _).
   Defined.
diff -r 4eb6407c6ca3 ComputableCategory.v
--- a/ComputableCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/ComputableCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -9,19 +9,18 @@
 
 Section ComputableCategory.
   Variable I : Type.
-  Variable Index2Object : I -> Type.
-  Variable Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i).
+  Variable Index2Cat : I -> Category.
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
-  Definition ComputableCategory : @SpecializedCategory I.
-    refine (@Build_SpecializedCategory _
-                                       (fun C D : I => SpecializedFunctor C D)
-                                       (fun o : I => IdentityFunctor o)
-                                       (fun C D E : I => ComposeFunctors (C := C) (D := D) (E := E))
-                                       _
-                                       _
-                                       _);
+  Definition ComputableCategory : Category.
+    refine (@Build_Category I
+                            (fun C D : I => Functor C D)
+                            (fun o : I => IdentityFunctor o)
+                            (fun C D E : I => ComposeFunctors (C := C) (D := D) (E := E))
+                            _
+                            _
+                            _);
     abstract functor_eq.
   Defined.
 End ComputableCategory.
diff -r 4eb6407c6ca3 ComputableGraphCategory.v
--- a/ComputableGraphCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/ComputableGraphCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Graph GraphTranslation.
+Require Export Category Graph GraphTranslation.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -14,8 +14,8 @@
 
   Local Coercion Index2Graph : I >-> Graph.
 
-  Definition ComputableGraphCategory : @SpecializedCategory I.
-    refine (@Build_SpecializedCategory _
+  Definition ComputableGraphCategory : @Category I.
+    refine (@Build_Category _
                                        (fun C D : I => GraphTranslation C D)
                                        (fun o : I => IdentityGraphTranslation o)
                                        (fun C D E : I => ComposeGraphTranslations (C := C) (D := D) (E := E))
diff -r 4eb6407c6ca3 ComputableSchemaCategory.v
--- a/ComputableSchemaCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/ComputableSchemaCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -22,8 +22,8 @@
   Hint Rewrite ComposeSmallTranslationsAssociativity.
 
   (* XXX TODO: Automate this better. *)
-  Definition ComputableSchemaCategory : @SpecializedCategory O (fun s d : O => EquivalenceClass (@SmallTranslationsEquivalent s d)).
-    refine (Build_SpecializedCategory (fun s d : O => EquivalenceClass (@SmallTranslationsEquivalent s d))
+  Definition ComputableSchemaCategory : @Category O (fun s d : O => EquivalenceClass (@SmallTranslationsEquivalent s d)).
+    refine (Build_Category (fun s d : O => EquivalenceClass (@SmallTranslationsEquivalent s d))
       (fun o => @classOf _ _ (@SmallTranslationsEquivalent_rel _ _)
         (IdentitySmallTranslation o))
       (fun (s d d' : O) F G => @apply2_to_class _ _ _ _ _ _ (@SmallTranslationsEquivalent_rel _ _)
diff -r 4eb6407c6ca3 Correspondences.v
--- a/Correspondences.v	Fri May 24 20:09:37 2013 -0400
+++ b/Correspondences.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import ProofIrrelevance FunctionalExtensionality.
-Require Export SpecializedCategory Functor SetCategory ProductCategory Duals BoolCategory.
+Require Export Category Functor SetCategory ProductCategory Duals BoolCategory.
 Require Import Common Notations Subcategory.
 
 Set Implicit Arguments.
@@ -72,12 +72,12 @@
      is a functor
      [[M: C^{op} × C' → Set.]]
      *)
-  Context `(C : @SpecializedCategory objC).
-  Context `(C' : @SpecializedCategory objC').
+  Variable C : Category.
+  Context `(C' : @Category objC').
 
   Let COp := OppositeCategory C.
 
-  Variable M : SpecializedFunctor (COp * C') TypeCat. (* the correspondence *)
+  Variable M : Functor (COp * C') TypeCat. (* the correspondence *)
 
   (*
      If [M] is a correspondence from [C] to [C'], we can define a new
@@ -120,8 +120,8 @@
   (* TODO: Figure out how to get Coq to do automatic type inference
      here, and simplify this proof *)
   (* TODO(jgross): Rewrite fg_equal_in using typeclasses? for speed *)
-  Definition CorrespondenceCategory : @SpecializedCategory (C + C')%type.
-    refine (@Build_SpecializedCategory _
+  Definition CorrespondenceCategory : @Category (C + C')%type.
+    refine (@Build_Category _
                                        CorrespondenceCategory_Morphism
                                        CorrespondenceCategory_Identity
                                        CorrespondenceCategory_Compose
@@ -166,10 +166,10 @@
     ).
   Defined.
 
-  Definition CorrespondenceCategoryFunctor : SpecializedFunctor (C ★^{M} C') ([1]).
+  Definition CorrespondenceCategoryFunctor : Functor (C ★^{M} C') ([1]).
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           CorrespondenceCategoryFunctor_ObjectOf
           CorrespondenceCategoryFunctor_MorphismOf
           _
@@ -191,8 +191,8 @@
      [F : M → [1] ], we can define [C = F⁻¹ {0}], [C' = F⁻¹ {1}], and a
      correspondence [M : C → C'] by the formula
      [M (X, Y) = Hom_{M} (X, Y)]. *)
-  Context `(M : @SpecializedCategory objM).
-  Variable F : SpecializedFunctor M ([1]).
+  Variable M : Category.
+  Variable F : Functor M ([1]).
 
   (* Comments after these two are for if we want to use [ChainCategory] instead of [BoolCat]. *)
   Definition CorrespondenceCategory0 := FullSubcategory M (fun x => F x = false). (* proj1_sig (F x) = 0).*)
@@ -202,11 +202,11 @@
   Let C' := CorrespondenceCategory1.
   Let COp := OppositeCategory C.
 
-  Definition Correspondence : SpecializedFunctor (COp * C') TypeCat.
+  Definition Correspondence : Functor (COp * C') TypeCat.
     subst_body.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun cc' => Morphism M (proj1_sig (fst cc')) (proj1_sig (snd cc')))
           (fun s d m => fun m0 => Compose (C := M) (snd m) (Compose m0 (fst m)))
           _
diff -r 4eb6407c6ca3 DataMigrationFunctorExamples.v
--- a/DataMigrationFunctorExamples.v	Fri May 24 20:09:37 2013 -0400
+++ b/DataMigrationFunctorExamples.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,7 +1,7 @@
 Require Import JMeq.
 Require Export String.
 Require Export DataMigrationFunctors PathsCategory PathsCategoryFunctors.
-Require Import Notations Common Limits SetLimits SetColimits FEqualDep FunctorCategory DefinitionSimplification SetLimits SetColimits SpecializedCommaCategory CommaCategoryFunctors.
+Require Import Notations Common Limits SetLimits SetColimits FEqualDep FunctorCategory DefinitionSimplification SetLimits SetColimits CommaCategory CommaCategoryFunctors.
 
 Set Implicit Arguments.
 
@@ -256,9 +256,9 @@
         | _ => Empty_set
       end.
 
-    Example C_Category_Ex22 : LocallySmallSpecializedCategory _ := PathsCategory C_Edges_Ex22.
-    Example D_Category_Ex22 : LocallySmallSpecializedCategory _ := PathsCategory D_Edges_Ex22.
-    Example E_Category_Ex22 : LocallySmallSpecializedCategory _ := PathsCategory E_Edges_Ex22.
+    Example C_Category_Ex22 : LocallyCategory _ := PathsCategory C_Edges_Ex22.
+    Example D_Category_Ex22 : LocallyCategory _ := PathsCategory D_Edges_Ex22.
+    Example E_Category_Ex22 : LocallyCategory _ := PathsCategory E_Edges_Ex22.
 
     Example F_Functor_Ex22_ObjectOf (x : C_Objects_Ex22) : D_Objects_Ex22 :=
       match x with
@@ -278,13 +278,13 @@
         | V_Ex22_E => U_Ex22_D
       end.
 
-    Example F_Functor_Ex22 : SpecializedFunctor C_Category_Ex22 D_Category_Ex22.
+    Example F_Functor_Ex22 : Functor C_Category_Ex22 D_Category_Ex22.
     Proof.
       apply (@FunctorFromPaths _ _ _ _ F_Functor_Ex22_ObjectOf).
       fill_unique_paths_functor.
     Defined.
 
-    Example G_Functor_Ex22 : SpecializedFunctor E_Category_Ex22 D_Category_Ex22.
+    Example G_Functor_Ex22 : Functor E_Category_Ex22 D_Category_Ex22.
     Proof.
       apply (@FunctorFromPaths _ _ _ _ G_Functor_Ex22_ObjectOf).
       fill_unique_paths_functor.
@@ -307,7 +307,7 @@
           | U_Ex22_D => Id_Ex221
         end.
 
-      Example  δ_Functor_Ex221 : SpecializedFunctor D_Category_Ex22 SetCat.
+      Example  δ_Functor_Ex221 : Functor D_Category_Ex22 SetCat.
       Proof.
         apply (@FunctorFromPaths _ _ _ _ δ_Functor_Ex221_ObjectOf).
         eliminate_useless_paths_functor_cases;
@@ -495,7 +495,7 @@
           | T2_Ex22_C => T2_Id_Ex222
         end.
 
-      Example γ_Functor_Ex222 : SpecializedFunctor C_Category_Ex22 SetCat.
+      Example γ_Functor_Ex222 : Functor C_Category_Ex22 SetCat.
       Proof.
         apply (@FunctorFromPaths _ _ _ _ γ_Functor_Ex222_ObjectOf).
         eliminate_useless_paths_functor_cases;
@@ -559,7 +559,7 @@
                                     D_Category_Ex22
                                     TypeCat
                                     F_Functor_Ex22
-                                    (fun (g : SpecializedFunctorToType _) d => @TypeLimit
+                                    (fun (g : FunctorToType _) d => @TypeLimit
                                                                                  _
                                                                                  _
                                                                                  (RightPushforwardAlong_pre_Functor
@@ -567,14 +567,14 @@
                                                                                     D_Category_Ex22
                                                                                     TypeCat
                                                                                     F_Functor_Ex22
-                                                                                    (g : SpecializedFunctorToType _)
+                                                                                    (g : FunctorToType _)
                                                                                     d))).
 
       Let U_Ex22_D_ex := Eval hnf in (ObjectOf (@RightPushforwardAlong C_Category_Ex22
                                                                        D_Category_Ex22
                                                                        TypeCat
                                                                        F_Functor_Ex22
-                                                                       (fun (g : SpecializedFunctorToType _) d => @TypeLimit
+                                                                       (fun (g : FunctorToType _) d => @TypeLimit
                                                                                                                     _
                                                                                                                     _
                                                                                                                     (RightPushforwardAlong_pre_Functor
@@ -582,16 +582,16 @@
                                                                                                                        D_Category_Ex22
                                                                                                                        TypeCat
                                                                                                                        F_Functor_Ex22
-                                                                                                                       (g : SpecializedFunctorToType _)
+                                                                                                                       (g : FunctorToType _)
                                                                                                                        d)))
-                                               ((γ_Functor_Ex222 : SpecializedFunctorToSet _) : SpecializedFunctorToType _)
+                                               ((γ_Functor_Ex222 : FunctorToSet _) : FunctorToType _)
                                                U_Ex22_D).
 
       Let Π_F__γ := (@RightPushforwardAlong C_Category_Ex22
                                             D_Category_Ex22
                                             TypeCat
                                             F_Functor_Ex22
-                                            (fun (g : SpecializedFunctorToType _) d => @TypeLimit
+                                            (fun (g : FunctorToType _) d => @TypeLimit
                                                                                          _
                                                                                          _
                                                                                          (RightPushforwardAlong_pre_Functor
@@ -599,9 +599,9 @@
                                                                                             D_Category_Ex22
                                                                                             TypeCat
                                                                                             F_Functor_Ex22
-                                                                                            (g : SpecializedFunctorToType _)
+                                                                                            (g : FunctorToType _)
                                                                                             d)))
-                      ((γ_Functor_Ex222 : SpecializedFunctorToSet _) : SpecializedFunctorToType _).
+                      ((γ_Functor_Ex222 : FunctorToSet _) : FunctorToType _).
 
       Section Π_F__γ__T1_Ex222_cleanup.
         Example Π_F__γ__U_Ex222' :=
@@ -613,7 +613,7 @@
           Eval cbv beta iota zeta delta [Π_F__γ__U_Ex222'
                                            Object
                                            C_Edges_Ex22
-                                           SliceSpecializedCategory_Functor
+                                           SliceCategory_Functor
                                            RightPushforwardAlong_pre_Functor
                                            γ_Functor_Ex222] in Π_F__γ__U_Ex222'.
 
@@ -630,7 +630,7 @@
           assert (f : focus Π_F__γ__U_Ex222''') by constructor.
           unfold Π_F__γ__U_Ex222''' in f.
           simpl in f.
-          unfold CommaSpecializedCategory_ObjectT in *; simpl in *.
+          unfold CommaCategory_ObjectT in *; simpl in *.
           unfold γ_Functor_Ex222_ObjectOf in *.
           simpl in *.
           unfold F_Functor_Ex22_ObjectOf in *.
@@ -646,29 +646,29 @@
           subst_body; hnf.
           intro f.
           match type of f with
-            | focus ({ S0 : forall c : CommaSpecializedCategory_Object ?A ?B, @?C c |
+            | focus ({ S0 : forall c : CommaCategory_Object ?A ?B, @?C c |
                        @?D S0 }) =>
-              clear f; assert (f : focus ({ S0 : forall c : CommaSpecializedCategory_ObjectT A B,
-                                                   C (Build_CommaSpecializedCategory_Object A B c) |
-                                            D (fun c => S0 (CommaSpecializedCategory_Object_Member c)) }))
+              clear f; assert (f : focus ({ S0 : forall c : CommaCategory_ObjectT A B,
+                                                   C (Build_CommaCategory_Object A B c) |
+                                            D (fun c => S0 (CommaCategory_Object_Member c)) }))
                        by constructor;
-              unfold CommaSpecializedCategory_ObjectT in f; simpl in f
+              unfold CommaCategory_ObjectT in f; simpl in f
           end.
           match type of f with
             | focus ({ S0 : ?ST |
-                       forall (c c' : CommaSpecializedCategory_Object ?A ?B)
-                              (g : CommaSpecializedCategory_Morphism (CommaSpecializedCategory_Object_Member c)
-                                                                     (CommaSpecializedCategory_Object_Member c')),
+                       forall (c c' : CommaCategory_Object ?A ?B)
+                              (g : CommaCategory_Morphism (CommaCategory_Object_Member c)
+                                                                     (CommaCategory_Object_Member c')),
                          @?D S0 c c' g }) =>
               clear f; assert (f : focus { S0 : ST |
-                                           forall (c c' : CommaSpecializedCategory_ObjectT A B)
-                                                  (g : CommaSpecializedCategory_MorphismT c c'),
+                                           forall (c c' : CommaCategory_ObjectT A B)
+                                                  (g : CommaCategory_MorphismT c c'),
                                              D S0
-                                               (Build_CommaSpecializedCategory_Object A B c)
-                                               (Build_CommaSpecializedCategory_Object A B c')
-                                               (Build_CommaSpecializedCategory_Morphism c c' g) })
+                                               (Build_CommaCategory_Object A B c)
+                                               (Build_CommaCategory_Object A B c')
+                                               (Build_CommaCategory_Morphism c c' g) })
                        by constructor;
-              unfold CommaSpecializedCategory_ObjectT, CommaSpecializedCategory_MorphismT in f;
+              unfold CommaCategory_ObjectT, CommaCategory_MorphismT in f;
               simpl in f
           end.
           match type of f with
@@ -839,24 +839,24 @@
         destruct x.
         simpl in *.
         match type of x with
-          | forall c : CommaSpecializedCategory_Object ?A ?B, @?f c =>
-            assert (x' : forall c : CommaSpecializedCategory_ObjectT A B,
-                      f (Build_CommaSpecializedCategory_Object A B c))
+          | forall c : CommaCategory_Object ?A ?B, @?f c =>
+            assert (x' : forall c : CommaCategory_ObjectT A B,
+                      f (Build_CommaCategory_Object A B c))
         end.
         admit.
         match type of e with
-          | forall (c c' : CommaSpecializedCategory_Object ?A ?B) (g : CommaSpecializedCategory_Morphism (CommaSpecializedCategory_Object_Member c)
-               (CommaSpecializedCategory_Object_Member c')),
+          | forall (c c' : CommaCategory_Object ?A ?B) (g : CommaCategory_Morphism (CommaCategory_Object_Member c)
+               (CommaCategory_Object_Member c')),
               @?f c c' g =>
-            clear e; assert (e : forall (c c' : CommaSpecializedCategory_ObjectT A B) (g : CommaSpecializedCategory_MorphismT c c'),
-                                    f (Build_CommaSpecializedCategory_Object A B c) (Build_CommaSpecializedCategory_Object A B c')
-                                      (Build_CommaSpecializedCategory_Morphism c c' g)) by admit
+            clear e; assert (e : forall (c c' : CommaCategory_ObjectT A B) (g : CommaCategory_MorphismT c c'),
+                                    f (Build_CommaCategory_Object A B c) (Build_CommaCategory_Object A B c')
+                                      (Build_CommaCategory_Morphism c c' g)) by admit
         end.
         match type of x with
-          | forall c : CommaSpecializedCategory_Object ?A ?B, @?f c =>
+          | forall c : CommaCategory_Object ?A ?B, @?f c =>
             match type of e with
               | appcontext e'[x] =>
-                let e'' := context e'[fun c : CommaSpecializedCategory_Object A B => x' (CommaSpecializedCategory_Object_Member c)] in
+                let e'' := context e'[fun c : CommaCategory_Object A B => x' (CommaCategory_Object_Member c)] in
                 assert (e''' : e'') by admit
             end
         end.
@@ -864,10 +864,10 @@
         rename e''' into e'.
         simpl in *.
         match type of x with
-          | forall c : CommaSpecializedCategory_Object ?A ?B, @?f c =>
+          | forall c : CommaCategory_Object ?A ?B, @?f c =>
             repeat match type of e' with
                      | appcontext e''[x] =>
-                       let e''' := context e''[fun c : CommaSpecializedCategory_Object A B => x' (CommaSpecializedCategory_Object_Member c)] in
+                       let e''' := context e''[fun c : CommaCategory_Object A B => x' (CommaCategory_Object_Member c)] in
                        clear e'; assert (e' : e''') by admit
                    end
         end.
@@ -875,8 +875,8 @@
         rename x' into x, e' into e.
         simpl in *.
         compute in x.
-        unfold CommaSpecializedCategory_ObjectT in e; simpl in *.
-        unfold CommaSpecializedCategory_MorphismT in *; simpl in *.
+        unfold CommaCategory_ObjectT in e; simpl in *.
+        unfold CommaCategory_MorphismT in *; simpl in *.
         Print existT.
         Print exist.
         match type of e with
@@ -1064,13 +1064,13 @@
 
         compute in e.
         assert (forall
-        c : CommaSpecializedCategory_ObjectT
+        c : CommaCategory_ObjectT
               {|
               ObjectOf := fun _ : unit => U_Ex22_D;
               MorphismOf := fun _ _ _ : unit => NoEdges;
-              FCompositionOf' := SliceSpecializedCategory_Functor_subproof
+              FCompositionOf' := SliceCategory_Functor_subproof
                                    D_Category_Ex22 U_Ex22_D;
-              FIdentityOf' := SliceSpecializedCategory_Functor_subproof0
+              FIdentityOf' := SliceCategory_Functor_subproof0
                                 D_Category_Ex22 U_Ex22_D |} F_Functor_Ex22,
       match snd (projT1 ( c)) with
       | SSN_Ex22_C => SSN
@@ -1088,7 +1088,7 @@
 
         Require Import InitialTerminalCategory ChainCategory DiscreteCategoryFunctors.
 
-        Definition F objC (C : SpecializedCategory objC) : SpecializedFunctor C TerminalCategory.
+        Definition F objC (C : Category objC) : Functor C TerminalCategory.
           clear.
           eexists; intros; simpl; eauto.
           Grab Existential Variables.
@@ -1096,20 +1096,20 @@
           intros; simpl; eauto.
         Defined.
 
-        Let Π_F_C objC (C : @LocallySmallSpecializedCategory objC) := Eval hnf in
+        Let Π_F_C objC (C : @LocallyCategory objC) := Eval hnf in
                              (@RightPushforwardAlong C
-                                                     (TerminalCategory : LocallySmallSpecializedCategory _)
+                                                     (TerminalCategory : LocallyCategory _)
                                                      TypeCat
                                                      (@F objC C)
-                                                     (fun (g : SpecializedFunctorToType _) d => @TypeLimit
+                                                     (fun (g : FunctorToType _) d => @TypeLimit
                                                                                                   _
                                                                                                   _
                                                                                                   (RightPushforwardAlong_pre_Functor
                                                                                                      C
-                                                                                                     (TerminalCategory : LocallySmallSpecializedCategory _)
+                                                                                                     (TerminalCategory : LocallyCategory _)
                                                                                                      TypeCat
                                                                                                      (@F objC C)
-                                                                                                     (g : SpecializedFunctorToType _)
+                                                                                                     (g : FunctorToType _)
                                                                                                      d))).
 
         Let Π_F_C'_ObjectOf objC C := Eval hnf in @ObjectOf _ _ _ _ (@Π_F_C objC C).
@@ -1117,7 +1117,7 @@
 
         Require Import ProofIrrelevance.
 
-        Definition Functor_01_0 : SpecializedFunctor [0] [1].
+        Definition Functor_01_0 : Functor [0] [1].
           clear.
           eexists (fun _ => exist _ 0 _) _;
             intros; compute; try apply proof_irrelevance.
@@ -1126,7 +1126,7 @@
           intros; compute; constructor; trivial.
         Defined.
 
-        Definition Functor_01_1 : SpecializedFunctor [0] [1].
+        Definition Functor_01_1 : Functor [0] [1].
           clear.
           eexists (fun _ => exist _ 1 _) _;
             intros; compute; try apply proof_irrelevance.
@@ -1136,23 +1136,23 @@
         Defined.
 
         Let Π_F_01_F F := Eval hnf in
-                           (@RightPushforwardAlong ([0] : LocallySmallSpecializedCategory _)
-                                                   ([1] : LocallySmallSpecializedCategory _)
+                           (@RightPushforwardAlong ([0] : LocallyCategory _)
+                                                   ([1] : LocallyCategory _)
                                                    TypeCat
                                                    F
-                                                   (fun (g : SpecializedFunctorToType _) d => @TypeLimit
+                                                   (fun (g : FunctorToType _) d => @TypeLimit
                                                                                                 _
                                                                                                 _
                                                                                                 (RightPushforwardAlong_pre_Functor
-                                                                                                   ([0] : LocallySmallSpecializedCategory _)
-                                                                                                   ([1] : LocallySmallSpecializedCategory _)
+                                                                                                   ([0] : LocallyCategory _)
+                                                                                                   ([1] : LocallyCategory _)
                                                                                                    TypeCat
                                                                                                    F
-                                                                                                   (g : SpecializedFunctorToType _)
+                                                                                                   (g : FunctorToType _)
                                                                                                    d))).
 
-        Example Π_F_01_F_ObjectOf (F : SpecializedFunctor [0] [1]) (x : SpecializedFunctor [0] TypeCat) := Eval hnf in (@ObjectOf _ _ _ _ (Π_F_01_F F) x).
-        Example Π_F_01_F_MorphismOf (F : SpecializedFunctor [0] [1]) (s d : SpecializedFunctor [0] TypeCat) m m' := Eval simpl in (@MorphismOf _ _ _ _ (Π_F_01_F F) s d m m').
+        Example Π_F_01_F_ObjectOf (F : Functor [0] [1]) (x : Functor [0] TypeCat) := Eval hnf in (@ObjectOf _ _ _ _ (Π_F_01_F F) x).
+        Example Π_F_01_F_MorphismOf (F : Functor [0] [1]) (s d : Functor [0] TypeCat) m m' := Eval simpl in (@MorphismOf _ _ _ _ (Π_F_01_F F) s d m m').
 
         Example Π_F_01_F_ObjectOf_ObjectOf F x (x' : [1]%category) := Eval hnf in (@Π_F_01_F_ObjectOf F x x').
         Example Π_F_01_F_ObjectOf_MorhismOf F x s d (m : Morphism [1] s d) := Eval simpl in (MorphismOf (@Π_F_01_F_ObjectOf F x) m).
@@ -1164,39 +1164,39 @@
           hnf in *;
             simpl in *; unfold Object in *; simpl in *.
           match goal with
-            | [ f := { S0 : forall c : CommaSpecializedCategory_Object ?A ?B, @?C c |
+            | [ f := { S0 : forall c : CommaCategory_Object ?A ?B, @?C c |
                        @?D S0 } |- _ ] =>
-              clear f; pose ({ S0 : forall c : CommaSpecializedCategory_ObjectT A B, C (Build_CommaSpecializedCategory_Object A B c) |
-                               D (fun c => S0 (CommaSpecializedCategory_Object_Member c)) }) as f; unfold CommaSpecializedCategory_ObjectT in *;
+              clear f; pose ({ S0 : forall c : CommaCategory_ObjectT A B, C (Build_CommaCategory_Object A B c) |
+                               D (fun c => S0 (CommaCategory_Object_Member c)) }) as f; unfold CommaCategory_ObjectT in *;
               simpl in *
           end.
           match goal with
             | [ f := { S0 : ?ST |
-                       forall (c c' : CommaSpecializedCategory_Object ?A ?B)
-                              (g : CommaSpecializedCategory_Morphism (CommaSpecializedCategory_Object_Member c)
-                                                                     (CommaSpecializedCategory_Object_Member c')),
+                       forall (c c' : CommaCategory_Object ?A ?B)
+                              (g : CommaCategory_Morphism (CommaCategory_Object_Member c)
+                                                                     (CommaCategory_Object_Member c')),
                          @?D S0 c c' g } |- _ ] =>
               clear f; pose ({ S0 : ST |
-                               forall (c c' : CommaSpecializedCategory_ObjectT A B)
-                                      (g : CommaSpecializedCategory_MorphismT c c'),
+                               forall (c c' : CommaCategory_ObjectT A B)
+                                      (g : CommaCategory_MorphismT c c'),
                                  D S0
-                                   (Build_CommaSpecializedCategory_Object A B c)
-                                   (Build_CommaSpecializedCategory_Object A B c')
-                                   (Build_CommaSpecializedCategory_Morphism c c' g) }) as f;
-              unfold CommaSpecializedCategory_ObjectT, CommaSpecializedCategory_MorphismT in f;
+                                   (Build_CommaCategory_Object A B c)
+                                   (Build_CommaCategory_Object A B c')
+                                   (Build_CommaCategory_Morphism c c' g) }) as f;
+              unfold CommaCategory_ObjectT, CommaCategory_MorphismT in f;
               simpl in f
           end.
           exact f.
         Defined.
 
-        Example Π_F_01_F_ObjectOf_ObjectOf (F : SpecializedFunctor [0] [1]) (x : SpecializedFunctor [0] TypeCat) (x' : [1]%category) :
+        Example Π_F_01_F_ObjectOf_ObjectOf (F : Functor [0] [1]) (x : Functor [0] TypeCat) (x' : [1]%category) :
           typeof (Π_F_01_F_ObjectOf_ObjectOf F x x').
         Proof.
           hnf in *; simpl in *.
           assert (Hf : focus (Π_F_01_F_ObjectOf_ObjectOf F x x')) by constructor.
           unfold Π_F_01_F_ObjectOf_ObjectOf in Hf; simpl in Hf.
           revert Hf; clear; intro.
-          unfold CommaSpecializedCategory_ObjectT, CommaSpecializedCategory_MorphismT in Hf.
+          unfold CommaCategory_ObjectT, CommaCategory_MorphismT in Hf.
           simpl in Hf.
           match type of Hf with
             | focus ({ S0 : forall c : { ab : unit * ?A & @?B ab },
@@ -1260,7 +1260,7 @@
           exact (@id _).
         Defined.
 
-        Definition f_to_functor (f : [0]%category -> TypeCat) : SpecializedFunctor [0] TypeCat.
+        Definition f_to_functor (f : [0]%category -> TypeCat) : Functor [0] TypeCat.
         Proof.
           revert f; clear; intro.
           exists f (f_to_functor_MorphismOf f); intros; simpl; destruct_sig; simpl in *;
@@ -1331,26 +1331,26 @@
           match goal with
             | [ f := { S0 : ?ST |
                        forall (c c' : ?C)
-                              (g : CommaSpecializedCategory_Morphism c c'),
+                              (g : CommaCategory_Morphism c c'),
                          @?D S0 c c' g } |- _ ] =>
               clear f; pose ({ S0 : ST |
                                forall (c c' : C)
-                                      (g : CommaSpecializedCategory_MorphismT (Build_CommaSpecializedCategory_Object c) (CommaSpecializedCategory_Object_Member c')),
-                                 D S0 c c' (Build_CommaSpecializedCategory_Morphism c c' g) }) (*as f;
-              unfold CommaSpecializedCategory_MorphismT in f; simpl in * *)
+                                      (g : CommaCategory_MorphismT (Build_CommaCategory_Object c) (CommaCategory_Object_Member c')),
+                                 D S0 c c' (Build_CommaCategory_Morphism c c' g) }) (*as f;
+              unfold CommaCategory_MorphismT in f; simpl in * *)
           end.
                                                                       c')),
 
 
                          @?D S0 c c' g
-                              (g : CommaSpecializedCategory_Morphism (CommaSpecializedCategory_Object_Member c)
-                                                                     (CommaSpecializedCategory_Object_Member c')),
+                              (g : CommaCategory_Morphism (CommaCategory_Object_Member c)
+                                                                     (CommaCategory_Object_Member c')),
 
-                        D (exist _ (Build_CommaSpecializedCategory_Object A B (proj1_sig S0)) (proj2_sig S0))
-                          (Build_CommaSpecializedCategory_Object A' B' c)
-                          (Build_CommaSpecializedCategory_Object A' B' c')
-                          (Build_CommaSpecializedCategory_Morphism (Build_CommaSpecializedCategory_Object A' B' c)
-                                                                   (Build_CommaSpecializedCategory_Object A' B' c')
+                        D (exist _ (Build_CommaCategory_Object A B (proj1_sig S0)) (proj2_sig S0))
+                          (Build_CommaCategory_Object A' B' c)
+                          (Build_CommaCategory_Object A' B' c')
+                          (Build_CommaCategory_Morphism (Build_CommaCategory_Object A' B' c)
+                                                                   (Build_CommaCategory_Object A' B' c')
                                                                    g) }) as f'
 
           compute in f.
@@ -1377,21 +1377,21 @@
         clear Π_F_C'
         compute in Π_F_C''.
 
-        ((γ_Functor_Ex222 : SpecializedFunctorToSet _) : SpecializedFunctorToType _).
+        ((γ_Functor_Ex222 : FunctorToSet _) : FunctorToType _).
 
         Goal Π_F__γ__U_Ex222''''.
         hnf.
         subst_body.
         eexists.
         intros.
-        destruct_head @CommaSpecializedCategory_Object.
+        destruct_head @CommaCategory_Object.
         simpl in *.
         destruct_sig.
         simpl in *.
         destruct_head prod.
         simpl in *.
         destruct_head unit.
-        destruct_head @CommaSpecializedCategory_Morphism.
+        destruct_head @CommaCategory_Morphism.
         simpl in *.
         destruct_sig; destruct_head prod; destruct_head unit; simpl in *.
         destruct c0; simpl in *.
@@ -1404,13 +1404,13 @@
             destruct_head C_Objects_Ex22;
               simpl in *.
 
-            CommaSpecializedCategory_Object
+            CommaCategory_Object
               {|
                 ObjectOf := fun _ : unit => U_Ex22_D;
                 MorphismOf := fun _ _ _ : unit => NoEdges;
-                FCompositionOf' := SliceSpecializedCategory_Functor_subproof
+                FCompositionOf' := SliceCategory_Functor_subproof
                                      D_Category_Ex22 U_Ex22_D;
-                FIdentityOf' := SliceSpecializedCategory_Functor_subproof0
+                FIdentityOf' := SliceCategory_Functor_subproof0
                                   D_Category_Ex22 U_Ex22_D |} F_Functor_Ex22.
 
             Example Π_F__γ__U_Ex222''''' : Type.
@@ -1420,12 +1420,12 @@
               - intros a b; hnf in *.
                 destruct a as [ x e ].
                 simpl in *.
-                pose proof (fun c : CommaSpecializedCategory_ObjectT _ _ => x c) as x'; simpl in x'.
+                pose proof (fun c : CommaCategory_ObjectT _ _ => x c) as x'; simpl in x'.
                 move x' at top.
                 compute in x'.
-                pose proof (fun (c c' : CommaSpecializedCategory_ObjectT _ _) (g : CommaSpecializedCategory_MorphismT _ _) => e c c' g) as e'.
+                pose proof (fun (c c' : CommaCategory_ObjectT _ _) (g : CommaCategory_MorphismT _ _) => e c c' g) as e'.
                 move e' at top; simpl in e'.
-                unfold CommaSpecializedCategory_ObjectT, CommaSpecializedCategory_MorphismT in e'; simpl in e'.
+                unfold CommaCategory_ObjectT, CommaCategory_MorphismT in e'; simpl in e'.
                 unfold C_Edges_Ex22, D_Edges_Ex22 in *; simpl in *.
                 pattern x in e'.
                 Check e'.
@@ -1462,24 +1462,24 @@
                     Eval hnf in Π_F__γ__U_Ex222''''.
 
 
-                    change (CommaSpecializedCategory_Object_Member ?x) with x in *.
+                    change (CommaCategory_Object_Member ?x) with x in *.
                     match goal with
                       | [ f := context G'[let (x) := ?y in x] |- _ ] => idtac
                     end.
                     change (let (x) := ?y in x) with y in *.
                     unfold C_Objects_Ex22 in *.
                     simpl in f.
-                    change (CommaSpecializedCategory_Object ?A ?B) with (CommaSpecializedCategory_ObjectT A B) in f.
+                    change (CommaCategory_Object ?A ?B) with (CommaCategory_ObjectT A B) in f.
 
-                    Check (CommaSpecializedCategory_ObjectT _ _ : CommaSpecializedCategory_Object
+                    Check (CommaCategory_ObjectT _ _ : CommaCategory_Object
                                                                     {|
                                                                       ObjectOf := fun _ : unit => U_Ex22_D;
                                                                       MorphismOf := fun _ _ _ : unit => NoEdges;
-                                                                      FCompositionOf' := SliceSpecializedCategory_Functor_subproof
+                                                                      FCompositionOf' := SliceCategory_Functor_subproof
                                                                                            D_Category_Ex22 U_Ex22_D;
-                                                                      FIdentityOf' := SliceSpecializedCategory_Functor_subproof0
+                                                                      FIdentityOf' := SliceCategory_Functor_subproof0
                                                                                         D_Category_Ex22 U_Ex22_D |} F_Functor_Ex22).
-                    unfold CommaSpecializedCategory_Object.
+                    unfold CommaCategory_Object.
                     unfold Object, C_Edges_Ex22 in f.
 
                     Print Π_F__γ__U_Ex222'.
diff -r 4eb6407c6ca3 DataMigrationFunctors.v
--- a/DataMigrationFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/DataMigrationFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -30,14 +30,14 @@
                       || (abstract apply CosliceCategoryInducedFunctor_FIdentityOf)
                       || (abstract apply SliceCategoryInducedFunctor_FCompositionOf)
                       || (abstract apply CosliceCategoryInducedFunctor_FCompositionOf)*)
-                      || apply CommaSpecializedCategory_Morphism_eq
+                      || apply CommaCategory_Morphism_eq
                       || apply f_equal
                       || (apply f_equal2; try reflexivity; [])
                       || apply sigT_eq (* simpl_eq *)
                       || (progress nt_eq)
                       || (progress functor_eq)
-                      || (progress (destruct_head_hnf @CommaSpecializedCategory_Object;
-                                    destruct_head_hnf @CommaSpecializedCategory_Morphism;
+                      || (progress (destruct_head_hnf @CommaCategory_Object;
+                                    destruct_head_hnf @CommaCategory_Morphism;
                                     destruct_sig)));
                    simpl; trivial.
 
@@ -52,27 +52,27 @@
            | _ => progress simpl_eq'
            | _ => progress (apply Functor_eq || apply Functor_JMeq)
            | _ => progress (apply NaturalTransformation_eq || apply NaturalTransformation_JMeq)
-           | _ => progress apply CommaSpecializedCategory_Morphism_eq
-           | _ => progress (destruct_head_hnf @CommaSpecializedCategory_Object; destruct_head_hnf @CommaSpecializedCategory_Morphism)
+           | _ => progress apply CommaCategory_Morphism_eq
+           | _ => progress (destruct_head_hnf @CommaCategory_Object; destruct_head_hnf @CommaCategory_Morphism)
            | _ => progress destruct_sig
            | _ => progress autorewrite with morphism
            | _ => progress repeat rewrite FIdentityOf
-           | [ |- @eq (LaxCosliceSpecializedCategory_Object _ _) _ _ ] => expand; apply f_equal
-           | [ |- @eq (LaxCosliceSpecializedCategory_Morphism _ _) _ _ ] => expand; apply f_equal
-           | [ |- @eq (LaxSliceSpecializedCategory_Object _ _) _ _ ] => expand; apply f_equal
-           | [ |- @eq (LaxSliceSpecializedCategory_Morphism _ _) _ _ ] => expand; apply f_equal
+           | [ |- @eq (LaxCosliceCategory_Object _ _) _ _ ] => expand; apply f_equal
+           | [ |- @eq (LaxCosliceCategory_Morphism _ _) _ _ ] => expand; apply f_equal
+           | [ |- @eq (LaxSliceCategory_Object _ _) _ _ ] => expand; apply f_equal
+           | [ |- @eq (LaxSliceCategory_Morphism _ _) _ _ ] => expand; apply f_equal
            | _ => progress (apply functional_extensionality_dep_JMeq; intros)
          end.
 
 Section DataMigrationFunctors.
-  Context `(C : LocallySmallSpecializedCategory objC).
-  Context `(D : LocallySmallSpecializedCategory objD).
-  Context `(S : SpecializedCategory objS).
+  Context `(C : LocallyCategory objC).
+  Context `(D : LocallyCategory objD).
+  Variable S : Category.
 
   Section Δ.
     Definition PullbackAlongFunctor : ((S ^ C) ^ (S ^ D)) ^ (D ^ C)
       := (ExponentialLaw4Functor_Inverse _ _ _) (FunctorialComposition _ _ _).
-    Definition PullbackAlong (F : SpecializedFunctor C D) : (S ^ C) ^ (S ^ D)
+    Definition PullbackAlong (F : Functor C D) : (S ^ C) ^ (S ^ D)
       := Eval hnf in PullbackAlongFunctor F.
     (*
 
@@ -90,7 +90,7 @@
       Implicit Arguments exist [A P].
       unfold IdentityFunctor, IdentityNaturalTransformation, NTComposeF, NTComposeT, ComposeFunctors in Hf.
       simpl in Hf.
-    Let PullbackAlong' (F : SpecializedFunctor C D) : { PA : (S ^ C) ^ (S ^ D) | PullbackAlongFunctor' F = PA }.
+    Let PullbackAlong' (F : Functor C D) : { PA : (S ^ C) ^ (S ^ D) | PullbackAlongFunctor' F = PA }.
       pre_abstract_trailing_props.
       functor_simpl_abstract_trailing_props_with_equality.
     Defined.
@@ -103,14 +103,14 @@
       nt_simpl_abstract_trailing_props_with_equality.
     Defined.
 
-    Definition PullbackAlong (F : SpecializedFunctor C D) : (S ^ C) ^ (S ^ D)
+    Definition PullbackAlong (F : Functor C D) : (S ^ C) ^ (S ^ D)
       := Eval hnf in proj1_sig (PullbackAlong' F).
 
     Let PullbackAlongFunctor'' : ((S ^ C) ^ (S ^ D)) ^ (D ^ C).
       hnf.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            PullbackAlong
                                            (fun s d m => proj1_sig (@PullbackAlongFunctor'_MorphismOf' s d m))
                                            _
@@ -231,7 +231,7 @@
       unfold Morphism, FunctorCategory at 1.
       functor_simpl_abstract_trailing_props_with_equality.
     Defined.
-    Definition PullbackAlong (F : SpecializedFunctor C D) : SpecializedFunctor (S ^ D) (S ^ C)
+    Definition PullbackAlong (F : Functor C D) : Functor (S ^ D) (S ^ C)
       := Eval hnf in PullbackAlong' F.
     Print PullbackAlong.
       := ComposeFunctors (FunctorialComposition C D S)
@@ -246,8 +246,8 @@
   Eval simpl in PullbackAlongFunctor.
 
   Section Π.
-    Local Notation "A ↓ F" := (CosliceSpecializedCategory A F).
-    (*Local Notation "C / c" := (@SliceSpecializedCategoryOver _ _ C c).*)
+    Local Notation "A ↓ F" := (CosliceCategory A F).
+    (*Local Notation "C / c" := (@SliceCategoryOver _ _ C c).*)
 
     (** Quoting David Spivak in "Functorial Data Migration":
        Definition 2.1.2. Let [F : C -> D] be a morphism of schemas and
@@ -276,17 +276,17 @@
      *)
 
     Section pre_functorial.
-      Variable F : SpecializedFunctor C D.
+      Variable F : Functor C D.
 
       (* Define [ɣ ○ (π^F d)] *)
-      Definition RightPushforwardAlong_pre_pre_Functor (g : S ^ C) (d : D) : SpecializedFunctor (d ↓ F) S.
+      Definition RightPushforwardAlong_pre_pre_Functor (g : S ^ C) (d : D) : Functor (d ↓ F) S.
         refine (ComposeFunctors (ComposeFunctors g (projT2 (CosliceCategoryProjectionFunctor C D F d))) _).
-        unfold CosliceSpecializedCategory, SliceSpecializedCategory_Functor, Object; simpl.
+        unfold CosliceCategory, SliceCategory_Functor, Object; simpl.
         refine (CommaCategoryInducedFunctor (s := (_, F)) (d := (_, F)) (_, IdentityNaturalTransformation F)).
         simpl.
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            refine (Build_SpecializedNaturalTransformation F G
+          | [ |- NaturalTransformation ?F ?G ] =>
+            refine (Build_NaturalTransformation F G
                                                            (fun _ => Identity _)
                                                            _)
         end;
@@ -294,17 +294,17 @@
       Defined.
 
       (*Variable HasLimits : forall g d, Limit (RightPushforwardAlong_pre_Functor g d).*)
-      (* Variable HasLimits : forall d (F' : SpecializedFunctor (d ↓ F) S), Limit F'.*)
+      (* Variable HasLimits : forall d (F' : Functor (d ↓ F) S), Limit F'.*)
 
       Let Index2Cat d := d ↓ F.
 
-      Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
 
       (*Let HasLimits' (C0 : CAT ⇑ S) : Limit (projT2 C0)
         := HasLimits (projT2 C0). *)
 
       Let RightPushforwardAlong_pre_curried_ObjectOf_pre (g : S ^ C) (d : D) : CAT ⇑ S
-        := existT _ (tt, _) (RightPushforwardAlong_pre_pre_Functor g d) : LaxCosliceSpecializedCategory_ObjectT Index2Cat S.
+        := existT _ (tt, _) (RightPushforwardAlong_pre_pre_Functor g d) : LaxCosliceCategory_ObjectT Index2Cat S.
 
       Let RightPushforwardAlong_pre_curried_ObjectOf (gd : (S ^ C) * D) : CAT ⇑ S
         := RightPushforwardAlong_pre_curried_ObjectOf_pre (fst gd) (snd gd).
@@ -318,8 +318,8 @@
         simpl;
         unfold Object, Morphism, GeneralizeFunctor.
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            let F' := eval hnf in F in let G' := eval hnf in G in change (SpecializedNaturalTransformation F' G')
+          | [ |- NaturalTransformation ?F ?G ] =>
+            let F' := eval hnf in F in let G' := eval hnf in G in change (NaturalTransformation F' G')
         end.
         exists (fun c : d' ↓ F => m (snd (projT1 c)) : Morphism S _ _).
         simpl;
@@ -332,7 +332,7 @@
         Morphism (CAT ⇑ S) (RightPushforwardAlong_pre_curried_ObjectOf gd) (RightPushforwardAlong_pre_curried_ObjectOf g'd')
         := @RightPushforwardAlong_pre_curried_MorphismOf_pre (fst gd) (snd gd) (fst g'd') (snd g'd') (fst m) (snd m).
 
-      Lemma RightPushforwardAlong_pre_curried_FCompositionOf (s d d' : SpecializedFunctor C S * LSObject D)
+      Lemma RightPushforwardAlong_pre_curried_FCompositionOf (s d d' : Functor C S * LSObject D)
             (m1 : Morphism ((S ^ C)%functor * D) s d)
             (m2 : Morphism ((S ^ C)%functor * D) d d') :
         RightPushforwardAlong_pre_curried_MorphismOf (Compose m2 m1) =
@@ -346,7 +346,7 @@
         Time anihilate.
       Qed.
 
-      Lemma RightPushforwardAlong_pre_curried_FIdentityOf (o : SpecializedFunctor C S * LSObject D) :
+      Lemma RightPushforwardAlong_pre_curried_FIdentityOf (o : Functor C S * LSObject D) :
         RightPushforwardAlong_pre_curried_MorphismOf (Identity o) =
         Identity (RightPushforwardAlong_pre_curried_ObjectOf o).
       Proof.
@@ -357,10 +357,10 @@
         Time anihilate.
       Qed.
 
-      Definition RightPushforwardAlong_pre_curried : SpecializedFunctor ((S ^ C) * D) (CAT ⇑ S).
+      Definition RightPushforwardAlong_pre_curried : Functor ((S ^ C) * D) (CAT ⇑ S).
         match goal with
-          | [ |- SpecializedFunctor ?F ?G ] =>
-            exact (Build_SpecializedFunctor F G
+          | [ |- Functor ?F ?G ] =>
+            exact (Build_Functor F G
                                             RightPushforwardAlong_pre_curried_ObjectOf
                                             RightPushforwardAlong_pre_curried_MorphismOf
                                             RightPushforwardAlong_pre_curried_FCompositionOf
@@ -373,13 +373,13 @@
     Section functorial.
       Let Index2Cat (dF : D * (D ^ C)) := (fst dF) ↓ (snd dF).
 
-      Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
 
       Let RightPushforwardAlongFunctor_pre_curried_ObjectOf (dgF : D * ((S ^ C) * (OppositeCategory (D ^ C)))) : CAT ⇑ S
         := let d := fst dgF in
            let g := fst (snd dgF) in
            let F := snd (snd dgF) in
-           existT _ (tt, (d, F)) (projT2 (RightPushforwardAlong_pre_curried F (g, d))) : LaxCosliceSpecializedCategory_ObjectT Index2Cat S.
+           existT _ (tt, (d, F)) (projT2 (RightPushforwardAlong_pre_curried F (g, d))) : LaxCosliceCategory_ObjectT Index2Cat S.
 
       Definition RightPushforwardAlongFunctor_pre_curried_MorphismOf dgF d'g'F' (m : Morphism (D * ((S ^ C) * (OppositeCategory (D ^ C)))) dgF d'g'F') :
         Morphism (CAT ⇑ S)
@@ -396,13 +396,13 @@
         let mg := constr:(fst (snd m)) in
         let mF := constr:(snd (snd m)) in
         exists (tt, CosliceCategoryInducedFunctor d' d md mF);
-          exists (fun c : CommaSpecializedCategory_Object (SliceSpecializedCategory_Functor D d') F'
+          exists (fun c : CommaCategory_Object (SliceCategory_Functor D d') F'
                   => mg (snd (projT1 c)));
           simpl; subst_body;
           abstract (
               intros;
-              destruct_head @CommaSpecializedCategory_Object;
-              destruct_head @CommaSpecializedCategory_Morphism;
+              destruct_head @CommaCategory_Object;
+              destruct_head @CommaCategory_Morphism;
               destruct_sig;
               destruct_head_hnf @prod;
               simpl in *;
@@ -410,10 +410,10 @@
             ).
       Defined.
 
-      Definition RightPushforwardAlongFunctor_pre_curried : SpecializedFunctor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) (CAT ⇑ S).
+      Definition RightPushforwardAlongFunctor_pre_curried : Functor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) (CAT ⇑ S).
         match goal with
-          | [ |- SpecializedFunctor ?C ?D ] =>
-            refine (Build_SpecializedFunctor C D
+          | [ |- Functor ?C ?D ] =>
+            refine (Build_Functor C D
                                              RightPushforwardAlongFunctor_pre_curried_ObjectOf
                                              RightPushforwardAlongFunctor_pre_curried_MorphismOf
                                              _
@@ -425,16 +425,16 @@
     End functorial.
 
     Section applied.
-      Variable HasLimits : forall (F : SpecializedFunctor C D) d (F' : SpecializedFunctor (d ↓ F) S), Limit F'.
+      Variable HasLimits : forall (F : Functor C D) d (F' : Functor (d ↓ F) S), Limit F'.
 
       Let Index2Cat (dF : D * (D ^ C)) := (fst dF) ↓ (snd dF).
 
-      Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
 
       Let HasLimits' (C0 : CAT ⇑ S) : Limit (projT2 C0)
         := HasLimits (projT2 C0).
 
-      Definition RightPushforwardAlongFunctor_curried : SpecializedFunctor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) S
+      Definition RightPushforwardAlongFunctor_curried : Functor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) S
         := ComposeFunctors (InducedLimitFunctor HasLimits') RightPushforwardAlongFunctor_pre_curried.
 
       Definition RightPushforwardAlongFunctor : ((S ^ D) ^ (S ^ C)) ^ (OppositeCategory (D ^ C))
@@ -454,7 +454,7 @@
         Check ExponentialLaw4Functor_Inverse_ObjectOf _. *)
 
       (*
-      Definition RightPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+      Definition RightPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
         NaturalTransformation
           (ComposeFunctors (RightPushforwardAlong_pre_Functor s c) (IdentityFunctor _))
           (RightPushforwardAlong_pre_Functor d c).
@@ -462,8 +462,8 @@
         eapply (NTComposeT (ComposeFunctorsAssociator1 _ _ _) _).
         Grab Existential Variables.
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            refine (Build_SpecializedNaturalTransformation F G
+          | [ |- NaturalTransformation ?F ?G ] =>
+            refine (Build_NaturalTransformation F G
                                                            (fun x => m (snd (projT1 x)))
                                                            _
                    )
@@ -479,7 +479,7 @@
             ).
       Defined.
 
-      Definition RightPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+      Definition RightPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
         Morphism S ((RightPushforwardAlong_ObjectOf s) c) ((RightPushforwardAlong_ObjectOf d) c).
       Proof.
         simpl; subst_body; simpl.
@@ -487,8 +487,8 @@
         exact (@RightPushforwardAlong_MorphismOf_ComponentsOf_Pre s d m c).
       Defined.
 
-      Definition RightPushforwardAlong_MorphismOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) :
-        SpecializedNaturalTransformation (RightPushforwardAlong_ObjectOf s) (RightPushforwardAlong_ObjectOf d).
+      Definition RightPushforwardAlong_MorphismOf (s d : S ^ C) (m : NaturalTransformation s d) :
+        NaturalTransformation (RightPushforwardAlong_ObjectOf s) (RightPushforwardAlong_ObjectOf d).
       Proof.
         exists (@RightPushforwardAlong_MorphismOf_ComponentsOf s d m).
         rename d into t.
@@ -504,8 +504,8 @@
   End Π.
 
   Section Σ.
-    Local Notation "F ↓ A" := (SliceSpecializedCategory A F).
-    (*Local Notation "C / c" := (@SliceSpecializedCategoryOver _ _ C c).*)
+    Local Notation "F ↓ A" := (SliceCategory A F).
+    (*Local Notation "C / c" := (@SliceCategoryOver _ _ C c).*)
 
     (** Quoting David Spivak in "Functorial Data Migration":
        Definition 2.1.3. Let [F : C -> D] be a morphism of schemas and
@@ -535,21 +535,21 @@
      *)
 
     Section pre_functorial.
-      Variable F : SpecializedFunctor C D.
+      Variable F : Functor C D.
 
       Let Index2Cat d := F ↓ d.
 
-      Local Notation "'CAT' ⇓ D" := (@LaxSliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇓ D" := (@LaxSliceCategory _ _ Index2Cat _ D).
 
       (* Define [ɣ ○ (π_F d)] *)
-      Definition LeftPushforwardAlong_pre_pre_Functor (g : S ^ C) (d : D) : SpecializedFunctor (F ↓ d) S.
+      Definition LeftPushforwardAlong_pre_pre_Functor (g : S ^ C) (d : D) : Functor (F ↓ d) S.
         refine (ComposeFunctors (ComposeFunctors g (projT2 (SliceCategoryProjectionFunctor C D F d))) _).
-        unfold SliceSpecializedCategory, SliceSpecializedCategory_Functor, Object; simpl.
+        unfold SliceCategory, SliceCategory_Functor, Object; simpl.
         refine (CommaCategoryInducedFunctor (s := (F, _)) (d := (F, _)) (IdentityNaturalTransformation F, _)).
         simpl.
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            refine (Build_SpecializedNaturalTransformation F G
+          | [ |- NaturalTransformation ?F ?G ] =>
+            refine (Build_NaturalTransformation F G
                                                            (fun _ => Identity _)
                                                            _)
         end;
@@ -557,7 +557,7 @@
       Defined.
 
       Let LeftPushforwardAlong_pre_curried_ObjectOf_pre (g : S ^ C) (d : D) : CAT ⇓ S
-        := existT _ (_, tt) (LeftPushforwardAlong_pre_pre_Functor g d) : LaxSliceSpecializedCategory_ObjectT _ Index2Cat S.
+        := existT _ (_, tt) (LeftPushforwardAlong_pre_pre_Functor g d) : LaxSliceCategory_ObjectT _ Index2Cat S.
 
       Let LeftPushforwardAlong_pre_curried_ObjectOf (gd : (S ^ C) * D) : CAT ⇓ S
         := LeftPushforwardAlong_pre_curried_ObjectOf_pre (fst gd) (snd gd).
@@ -571,8 +571,8 @@
         simpl;
         unfold Object, Morphism, GeneralizeFunctor.
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            let F' := eval hnf in F in let G' := eval hnf in G in change (SpecializedNaturalTransformation F' G')
+          | [ |- NaturalTransformation ?F ?G ] =>
+            let F' := eval hnf in F in let G' := eval hnf in G in change (NaturalTransformation F' G')
         end.
         exists (fun c : F ↓ d => m (fst (projT1 c)) : Morphism S _ _).
         simpl;
@@ -585,7 +585,7 @@
         Morphism (CAT ⇓ S) (LeftPushforwardAlong_pre_curried_ObjectOf gd) (LeftPushforwardAlong_pre_curried_ObjectOf g'd')
         := @LeftPushforwardAlong_pre_curried_MorphismOf_pre (fst gd) (snd gd) (fst g'd') (snd g'd') (fst m) (snd m).
 
-      Lemma LeftPushforwardAlong_pre_curried_FCompositionOf (s d d' : SpecializedFunctor C S * LSObject D)
+      Lemma LeftPushforwardAlong_pre_curried_FCompositionOf (s d d' : Functor C S * LSObject D)
             (m1 : Morphism ((S ^ C)%functor * D) s d)
             (m2 : Morphism ((S ^ C)%functor * D) d d') :
         LeftPushforwardAlong_pre_curried_MorphismOf (Compose m2 m1) =
@@ -599,7 +599,7 @@
         Time anihilate.
       Qed.
 
-      Lemma LeftPushforwardAlong_pre_curried_FIdentityOf (o : SpecializedFunctor C S * LSObject D) :
+      Lemma LeftPushforwardAlong_pre_curried_FIdentityOf (o : Functor C S * LSObject D) :
         LeftPushforwardAlong_pre_curried_MorphismOf (Identity o) =
         Identity (LeftPushforwardAlong_pre_curried_ObjectOf o).
       Proof.
@@ -610,10 +610,10 @@
         Time anihilate.
       Qed.
 
-      Definition LeftPushforwardAlong_pre_curried : SpecializedFunctor ((S ^ C) * D) (CAT ⇓ S).
+      Definition LeftPushforwardAlong_pre_curried : Functor ((S ^ C) * D) (CAT ⇓ S).
         match goal with
-          | [ |- SpecializedFunctor ?F ?G ] =>
-            exact (Build_SpecializedFunctor F G
+          | [ |- Functor ?F ?G ] =>
+            exact (Build_Functor F G
                                             LeftPushforwardAlong_pre_curried_ObjectOf
                                             LeftPushforwardAlong_pre_curried_MorphismOf
                                             LeftPushforwardAlong_pre_curried_FCompositionOf
@@ -626,13 +626,13 @@
     Section functorial.
       Let Index2Cat (dF : D * (D ^ C)) := (snd dF) ↓ (fst dF).
 
-      Local Notation "'CAT' ⇓ D" := (@LaxSliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇓ D" := (@LaxSliceCategory _ _ Index2Cat _ D).
 
       Let LeftPushforwardAlongFunctor_pre_curried_ObjectOf (dgF : D * ((S ^ C) * (OppositeCategory (D ^ C)))) : CAT ⇓ S
         := let d := fst dgF in
            let g := fst (snd dgF) in
            let F := snd (snd dgF) in
-           existT _ ((d, F), tt) (projT2 (LeftPushforwardAlong_pre_curried F (g, d))) : LaxSliceSpecializedCategory_ObjectT _ Index2Cat S.
+           existT _ ((d, F), tt) (projT2 (LeftPushforwardAlong_pre_curried F (g, d))) : LaxSliceCategory_ObjectT _ Index2Cat S.
 
       Definition LeftPushforwardAlongFunctor_pre_curried_MorphismOf dgF d'g'F' (m : Morphism (D * ((S ^ C) * (OppositeCategory (D ^ C)))) dgF d'g'F') :
         Morphism (CAT ⇓ S)
@@ -649,13 +649,13 @@
         let mg := constr:(fst (snd m)) in
         let mF := constr:(snd (snd m)) in
         exists (SliceCategoryInducedFunctor d d' md mF, tt);
-          exists (fun c : CommaSpecializedCategory_Object F (SliceSpecializedCategory_Functor D d)
+          exists (fun c : CommaCategory_Object F (SliceCategory_Functor D d)
                   => mg (fst (projT1 c)));
           simpl; subst_body;
           abstract (
               intros;
-              destruct_head @CommaSpecializedCategory_Object;
-              destruct_head @CommaSpecializedCategory_Morphism;
+              destruct_head @CommaCategory_Object;
+              destruct_head @CommaCategory_Morphism;
               destruct_sig;
               destruct_head_hnf @prod;
               simpl in *;
@@ -663,10 +663,10 @@
             ).
       Defined.
 
-      Definition LeftPushforwardAlongFunctor_pre_curried : SpecializedFunctor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) (CAT ⇓ S).
+      Definition LeftPushforwardAlongFunctor_pre_curried : Functor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) (CAT ⇓ S).
         match goal with
-          | [ |- SpecializedFunctor ?C ?D ] =>
-            refine (Build_SpecializedFunctor C D
+          | [ |- Functor ?C ?D ] =>
+            refine (Build_Functor C D
                                              LeftPushforwardAlongFunctor_pre_curried_ObjectOf
                                              LeftPushforwardAlongFunctor_pre_curried_MorphismOf
                                              _
@@ -678,16 +678,16 @@
     End functorial.
 
     Section applied.
-      Variable HasColimits : forall (F : SpecializedFunctor C D) d (F' : SpecializedFunctor (F ↓ d) S), Colimit F'.
+      Variable HasColimits : forall (F : Functor C D) d (F' : Functor (F ↓ d) S), Colimit F'.
 
       Let Index2Cat (dF : D * (D ^ C)) := (snd dF) ↓ (fst dF).
 
-      Local Notation "'CAT' ⇓ D" := (@LaxSliceSpecializedCategory _ _ Index2Cat _ D).
+      Local Notation "'CAT' ⇓ D" := (@LaxSliceCategory _ _ Index2Cat _ D).
 
       Let HasColimits' (C0 : CAT ⇓ S) : Colimit (projT2 C0)
         := HasColimits (projT2 C0).
 
-      Definition LeftPushforwardAlongFunctor_curried : SpecializedFunctor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) S
+      Definition LeftPushforwardAlongFunctor_curried : Functor (D * ((S ^ C) * (OppositeCategory (D ^ C)))) S
         := ComposeFunctors (InducedColimitFunctor HasColimits') LeftPushforwardAlongFunctor_pre_curried.
 
       Definition LeftPushforwardAlongFunctor : ((S ^ D) ^ (S ^ C)) ^ (OppositeCategory (D ^ C))
diff -r 4eb6407c6ca3 DataMigrationFunctorsAdjoint.v
--- a/DataMigrationFunctorsAdjoint.v	Fri May 24 20:09:37 2013 -0400
+++ b/DataMigrationFunctorsAdjoint.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,6 +1,6 @@
 Require Import FunctionalExtensionality.
 Require Export Adjoint DataMigrationFunctors.
-Require Import Common Notations FunctorCategory LimitFunctors AdjointUnit SpecializedCommaCategory InducedLimitFunctors CommaCategoryFunctors (* NaturalTransformation Hom Duals CommaCategoryFunctors SetLimits SetColimits LimitFunctors LimitFunctorTheorems InducedLimitFunctors DefinitionSimplification FEqualDep CommaCategoryFunctorProperties*).
+Require Import Common Notations FunctorCategory LimitFunctors AdjointUnit CommaCategory InducedLimitFunctors CommaCategoryFunctors (* NaturalTransformation Hom Duals CommaCategoryFunctors SetLimits SetColimits LimitFunctors LimitFunctorTheorems InducedLimitFunctors DefinitionSimplification FEqualDep CommaCategoryFunctorProperties*).
 
 Set Implicit Arguments.
 
@@ -12,18 +12,18 @@
 (*
 Section coslice_initial.
   (* TODO(jgross): This should go elsewhere *)
-  Definition CosliceSpecializedCategory_InitialObject objC C objD D (F : @SpecializedFunctor objC C objD D) x :
-    { o : _ & @InitialObject _ (CosliceSpecializedCategory (F x) F) o }.
+  Definition CosliceCategory_InitialObject objC C objD D (F : @Functor objC C objD D) x :
+    { o : _ & @InitialObject _ (CosliceCategory (F x) F) o }.
     unfold Object; simpl;
     match goal with
-      | [ |- { o : CommaSpecializedCategory_Object ?A ?B & _ } ] =>
-        (exists (existT _ (tt, _) (@Identity _ D _) : CommaSpecializedCategory_ObjectT A B))
+      | [ |- { o : CommaCategory_Object ?A ?B & _ } ] =>
+        (exists (existT _ (tt, _) (@Identity _ D _) : CommaCategory_ObjectT A B))
     end;
     simpl.
     intro o'; hnf; simpl.
     destruct o' as [ [ [ ? o' ] m' ] ]; simpl in *.
     evar (T : Type); evar (t : T); subst T.
-    eexists (Build_CommaSpecializedCategory_Morphism _ _ _).
+    eexists (Build_CommaCategory_Morphism _ _ _).
     instantiate (1 := t); simpl in *.
 
     {αβ : unit * objC & Morphism D (F x) (F (snd αβ))}
@@ -38,15 +38,15 @@
 End coslice_initial.
 *)
 Section DataMigrationFunctorsAdjoint.
-  Variables C D : LocallySmallCategory.
+  Variables C D : LocallyCategory.
   Variables S : Category.
 
-  Variable F : SpecializedFunctor C D.
+  Variable F : Functor C D.
 
   Let Δ_F := PullbackAlong C D S F.
 
   Section ΔΠ.
-    (*Variable HasLimits' : forall (F : SpecializedFunctor C D), Limit F.
+    (*Variable HasLimits' : forall (F : Functor C D), Limit F.
     Goal forall F (M := TerminalMorphism_Morphism (HasLimits' F)), True.
     intros.
     hnf in M.
@@ -64,19 +64,19 @@
                                        _
                                        (tt, c)
                                        (Identity (F c)) :
-                                       CommaSpecializedCategory_ObjectT (SliceSpecializedCategory_Functor D (F c)) F)).
+                                       CommaCategory_ObjectT (SliceCategory_Functor D (F c)) F)).
     (* simpl in *.
       unfold Limit, LimitObject in *;
         intro_from_universal_objects;
         intro_universal_morphisms;
         simpl in *.
       match goal with
-        | [ m : SpecializedNaturalTransformation _ _ |- _ ] =>
+        | [ m : NaturalTransformation _ _ |- _ ] =>
           let X := constr:(ComponentsOf m) in
           match type of X with
             | forall _ : ?T, _ => match eval hnf in T with
-                                    | CommaSpecializedCategory_Object ?A ?B =>
-                                      let x := constr:(CommaSpecializedCategory_ObjectT A B) in
+                                    | CommaCategory_Object ?A ?B =>
+                                      let x := constr:(CommaCategory_ObjectT A B) in
                                       match eval hnf in x with
                                         | @sigT (?A' * ?B') ?C' =>
                                           let a := fresh in
@@ -86,7 +86,7 @@
                                           match B' with | unit => pose tt as b | _ => evar (b : B') end;
                                           evar (c : C' (a, b));
                                           let X' := fresh in
-                                          pose (X ((@existT _ _ (a, b) c) : CommaSpecializedCategory_ObjectT A B)) as X';
+                                          pose (X ((@existT _ _ (a, b) c) : CommaCategory_ObjectT A B)) as X';
                                             subst a b c;
                                             simpl in X';
                                             apply X'
@@ -116,7 +116,7 @@
       match goal with
         | [ |- ?A = _ ] =>
           match A with
-            | context A'[@Build_SpecializedNaturalTransformation ?objC ?C ?objD ?D ?F ?G ?CO _] =>
+            | context A'[@Build_NaturalTransformation ?objC ?C ?objD ?D ?F ?G ?CO _] =>
               (* make evars for ComponentsOf, Commutes *)
               let ComT := fresh in
               let Com' := fresh in
@@ -134,7 +134,7 @@
                   reflexivity
                 | ];
                 (* do transitivity with the evar'd components *)
-                let A'' := context A'[@Build_SpecializedNaturalTransformation objC C objD D F G CO' Com'] in
+                let A'' := context A'[@Build_NaturalTransformation objC C objD D F G CO' Com'] in
                 transitivity A''; subst Com' CO'
           end
       end.
@@ -151,7 +151,7 @@
       simpl in *.
       present_spfunctor.
       present_spcategory.
-      Print CommaSpecializedCategory_Object.
+      Print CommaCategory_Object.
       Check fun (g : (S ^ C)%functor) (d : D) =>
               (RightPushforwardAlong_pre_Functor C D S F g d).
       match goal with
@@ -198,13 +198,13 @@
              (tt, d) (Identity (F d)))
       assert True.
       -
-        unfold CosliceSpecializedCategory, CommaSpecializedCategory, Object in x; simpl in x.
+        unfold CosliceCategory, CommaCategory, Object in x; simpl in x.
         unfold Object in x.
       asser
       Check (TerminalMorphism_Morphism (HasLimits G (F d)) (existT (fun αβ : unit * LSObject C => Morphism D (F d) (F (snd αβ)))
            (tt, d) (Identity (F d)))).
       match goal with
-        | [ |- context G[@Build_CommaSpecializedCategory_Object ?objA ?A ?objB ?C ?objC ?C
+        | [ |- context G[@Build_CommaCategory_Object ?objA ?A ?objB ?C ?objC ?C
       Set Printing All.
       assert (Compose
      ((TerminalMorphism_Morphism (HasLimits G (F d)))
@@ -213,7 +213,7 @@
      (TerminalProperty_Morphism (HasLimits G (F d))
         (TerminalMorphism_Object (HasLimits G (F s)))
         {|
-        ComponentsOf' := fun x : CosliceSpecializedCategory (F d) F =>
+        ComponentsOf' := fun x : CosliceCategory (F d) F =>
                          (TerminalMorphism_Morphism (HasLimits G (F s)))
                            (existT
                                                   (fun αβ : unit * LSObject C =>
@@ -234,7 +234,7 @@
       t with
       expand.
       match goal with
-        | [ |- @eq (@SpecializedNaturalTransformation ?objC ?C ?objD ?D ?F ?G) _ _ ] =>
+        | [ |- @eq (@NaturalTransformation ?objC ?C ?objD ?D ?F ?G) _ _ ] =>
           apply (@NaturalTransformation_eq objC C objD D F G)
       end.
       nt_eq.
@@ -249,14 +249,14 @@
       etransitivity.
       mat
       - do 3 match goal with
-               | [ |- @Build_SpecializedNaturalTransformation ?objC ?C ?objD ?D ?F ?G _ _ = _ ] => (apply (@NaturalTransformation_eq objC C objD D F G); simpl) || fail 1 "Cannot apply NaturalTransformation_eq"
-               | [ |- appcontext[Build_SpecializedNaturalTransformation] ] => apply f_equal2 || apply f_equal || fail 1 "Cannot apply f_equal"
+               | [ |- @Build_NaturalTransformation ?objC ?C ?objD ?D ?F ?G _ _ = _ ] => (apply (@NaturalTransformation_eq objC C objD D F G); simpl) || fail 1 "Cannot apply NaturalTransformation_eq"
+               | [ |- appcontext[Build_NaturalTransformation] ] => apply f_equal2 || apply f_equal || fail 1 "Cannot apply f_equal"
                | _ => reflexivity
              end.
-        instantiate (1 := (fun x : CosliceSpecializedCategory (F d) F =>
+        instantiate (1 := (fun x : CosliceCategory (F d) F =>
     (TerminalMorphism_Morphism (HasLimits G (F s)))
       {|
-      CommaSpecializedCategory_Object_Member := existT
+      CommaCategory_Object_Member := existT
                                                   (fun αβ : unit * LSObject C =>
                                                   Morphism D
                                                   (F s)
@@ -277,10 +277,10 @@
                                                    [ let H := fresh in intro H; rewrite <- H; simpl; reflexivity | ]
         end.
         Show Existentials.
-        instantiate (1 := (fun x : CosliceSpecializedCategory (F d) F =>
+        instantiate (1 := (fun x : CosliceCategory (F d) F =>
     (TerminalMorphism_Morphism (HasLimits G (F s)))
       {|
-      CommaSpecializedCategory_Object_Member := existT
+      CommaCategory_Object_Member := existT
                                                   (fun αβ : unit * LSObject C =>
                                                   Morphism D
                                                   (F s)
@@ -296,7 +296,7 @@
         Focus 2.
         instantiate (1 := (fun x => (TerminalMorphism_Morphism (HasLimits G (F s)))
      {|
-     CommaSpecializedCategory_Object_Member := existT
+     CommaCategory_Object_Member := existT
                                                  (fun αβ : unit * LSObject C =>
                                                   Morphism D
                                                   (F s)
@@ -320,12 +320,12 @@
 
 
         match goal with
-          | [ |- appcontext[Build_SpecializedNaturalTransformation] ] => try apply f_equal2
+          | [ |- appcontext[Build_NaturalTransformation] ] => try apply f_equal2
 
         end.
 
       match goal with
-        | [ |- appcontext[@Build_SpecializedNaturalTransformation ?objC ?C ?objD ?D ?F ?G ?CO ?Com] ] =>
+        | [ |- appcontext[@Build_NaturalTransformation ?objC ?C ?objD ?D ?F ?G ?CO ?Com] ] =>
           let A := fresh in
           let H := fresh in
           set (A := CO);
@@ -349,8 +349,8 @@
       clear.
 
       match goal with
-        | [ |- appcontext[@Build_SpecializedNaturalTransformation _ ?C _ ?D ?F ?G ?CO ?Com] ] =>
-          assert (forall com, @Build_SpecializedNaturalTransformation _ C _ D F G CO com = @Build_SpecializedNaturalTransformation _ C _ D F G CO com)
+        | [ |- appcontext[@Build_NaturalTransformation _ ?C _ ?D ?F ?G ?CO ?Com] ] =>
+          assert (forall com, @Build_NaturalTransformation _ C _ D F G CO com = @Build_NaturalTransformation _ C _ D F G CO com)
                  by reflexivity
       end.
       assert True.
@@ -386,7 +386,7 @@
                                  ((TerminalMorphism_Morphism
                                      (HasLimits G (F s)))
                                     {|
-                                    CommaSpecializedCategory_Object_Member := existT
+                                    CommaCategory_Object_Member := existT
                                                   (fun αβ : unit * LSObject C =>
                                                   Morphism D
                                                   (F s)
@@ -404,7 +404,7 @@
       assert (
           forall e2 e2',
 {|
-        ComponentsOf' := fun c : CosliceSpecializedCategory (F d) F =>
+        ComponentsOf' := fun c : CosliceCategory (F d) F =>
                          Compose
                            (Compose
                               (Compose
@@ -417,7 +417,7 @@
                                  ((TerminalMorphism_Morphism
                                      (HasLimits G (F s)))
                                     {|
-                                    CommaSpecializedCategory_Object_Member := @existT
+                                    CommaCategory_Object_Member := @existT
                                                   (unit * LSObject C)
                                                   (fun αβ : unit * LSObject C =>
                                                   Morphism D
@@ -431,7 +431,7 @@
         Commutes' := e2 |}
 =
 {|
-        ComponentsOf' := fun c : CosliceSpecializedCategory (F d) F =>
+        ComponentsOf' := fun c : CosliceCategory (F d) F =>
                          Compose
                            (Compose
                               (Compose
@@ -444,7 +444,7 @@
                                  ((TerminalMorphism_Morphism
                                      (HasLimits G (F s)))
                                     {|
-                                    CommaSpecializedCategory_Object_Member := @existT
+                                    CommaCategory_Object_Member := @existT
                                                   (unit * LSObject C)
                                                   (fun αβ : unit * LSObject C =>
                                                   Morphism D
@@ -457,7 +457,7 @@
                               (TerminalMorphism_Object (HasLimits G (F s))));
         Commutes' := e2 |}).
             {|
-              ComponentsOf' := fun c : CosliceSpecializedCategory (F d) F =>
+              ComponentsOf' := fun c : CosliceCategory (F d) F =>
                                  Compose
                                    (Compose
                                       (Compose
@@ -470,7 +470,7 @@
                                          ((TerminalMorphism_Morphism
                                              (HasLimits G (F s)))
                                             {|
-                                              CommaSpecializedCategory_Object_Member := existT
+                                              CommaCategory_Object_Member := existT
                                                                                           (fun αβ : unit * LSObject C =>
                                                                                              Morphism D
                                                                                                       (F s)
@@ -482,7 +482,7 @@
                                       (TerminalMorphism_Object (HasLimits G (F s))));
               Commutes' := e2 |} =
             {|
-              ComponentsOf' := fun c : CosliceSpecializedCategory (F d) F =>
+              ComponentsOf' := fun c : CosliceCategory (F d) F =>
                                  Compose
                                    (Compose
                                       (Compose
@@ -495,7 +495,7 @@
                                          ((TerminalMorphism_Morphism
                                              (HasLimits G (F s)))
                                             {|
-                                              CommaSpecializedCategory_Object_Member := existT
+                                              CommaCategory_Object_Member := existT
                                                                                           (fun αβ : unit * LSObject C =>
                                                                                              Morphism D
                                                                                                       (F s)
@@ -545,10 +545,10 @@
               (IdentityFunctor (S ^ C)%functor).
 
       let t := type of H1 in
-      pose ((@existT _ _ (H, H0) H1) : CommaSpecializedCategory_ObjectT (SliceSpecializedCategory_Functor D (F c)) F) as H2.
+      pose ((@existT _ _ (H, H0) H1) : CommaCategory_ObjectT (SliceCategory_Functor D (F c)) F) as H2.
       Set Printing All.
 
-      Print CommaSpecializedCategory_Object.
+      Print CommaCategory_Object.
       unfold
       apply X.
       intro_universal_properties.
@@ -605,8 +605,8 @@
   End Δ.
 
   Section Π.
-    Local Notation "A ↓ F" := (CosliceSpecializedCategory A F).
-    (*Local Notation "C / c" := (@SliceSpecializedCategoryOver _ _ C c).*)
+    Local Notation "A ↓ F" := (CosliceCategory A F).
+    (*Local Notation "C / c" := (@SliceCategoryOver _ _ C c).*)
 
     (** Quoting David Spivak in "Functorial Data Migration":
        Definition 2.1.2. Let [F : C -> D] be a morphism of schemas and
@@ -635,20 +635,20 @@
        *)
 
     (* Define [ɣ ○ (π^F d)] *)
-    Definition RightPushforwardAlong_pre_Functor (g : S ^ C) (d : D) : SpecializedFunctor (d ↓ F) S
+    Definition RightPushforwardAlong_pre_Functor (g : S ^ C) (d : D) : Functor (d ↓ F) S
       := ComposeFunctors g (projT2 (CosliceCategoryProjectionFunctor C D F d)).
 
     Variable HasLimits : forall g d, Limit (RightPushforwardAlong_pre_Functor g d).
 
     Let Index2Cat d := d ↓ F.
 
-    Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
+    Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
 
     Let RightPushforwardAlong_ObjectOf_ObjectOf (g : S ^ C) d := LimitObject (HasLimits g d).
 
     Let RightPushforwardAlong_ObjectOf_MorphismOf_Pre' (g : S ^ C) s d (m : Morphism D s d) :
-      {F0 : unit * SpecializedFunctor (d ↓ F) (s ↓ F) &
-        SpecializedNaturalTransformation
+      {F0 : unit * Functor (d ↓ F) (s ↓ F) &
+        NaturalTransformation
         (ComposeFunctors (RightPushforwardAlong_pre_Functor g s) (snd F0))
         (RightPushforwardAlong_pre_Functor g d) }.
       exists (tt, fst (proj1_sig (MorphismOf (CosliceCategoryProjectionFunctor C D F) m))).
@@ -657,15 +657,15 @@
           simpl;
             unfold Object, Morphism, GeneralizeFunctor.
       match goal with
-        | [ |- SpecializedNaturalTransformation (ComposeFunctors (ComposeFunctors ?g ?h) ?i) _ ] =>
+        | [ |- NaturalTransformation (ComposeFunctors (ComposeFunctors ?g ?h) ?i) _ ] =>
           eapply (NTComposeT _ (ComposeFunctorsAssociator1 g h i))
       end.
       Grab Existential Variables.
       eapply (NTComposeF (IdentityNaturalTransformation g) _).
       Grab Existential Variables.
       match goal with
-        | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ C : _ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun _ => Identity (C := C) _)
             _
           )
@@ -680,20 +680,20 @@
 
     Let RightPushforwardAlong_ObjectOf_MorphismOf_Pre'' (g : S ^ C) s d (m : Morphism D s d) :
       Morphism (CAT ⇑ S)
-      (existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : @LaxCosliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
-      (existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : @LaxCosliceSpecializedCategory_ObjectT _ _ Index2Cat _ _).
+      (existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : @LaxCosliceCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : @LaxCosliceCategory_ObjectT _ _ Index2Cat _ _).
     Proof.
       hnf.
       match goal with
-        | [ |- LaxCosliceSpecializedCategory_Morphism ?a ?b ] =>
-          exact (RightPushforwardAlong_ObjectOf_MorphismOf_Pre' g _ _ m : @LaxCosliceSpecializedCategory_MorphismT _ _ _ _ _ a b)
+        | [ |- LaxCosliceCategory_Morphism ?a ?b ] =>
+          exact (RightPushforwardAlong_ObjectOf_MorphismOf_Pre' g _ _ m : @LaxCosliceCategory_MorphismT _ _ _ _ _ a b)
       end.
     Defined.
 
     Definition RightPushforwardAlong_ObjectOf_MorphismOf_Pre (g : S ^ C) s d (m : Morphism D s d) :
       Morphism (CAT ⇑ S)
-      (existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : @LaxCosliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
-      (existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : @LaxCosliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : @LaxCosliceCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : @LaxCosliceCategory_ObjectT _ _ Index2Cat _ _)
       := Eval cbv beta iota zeta delta [RightPushforwardAlong_ObjectOf_MorphismOf_Pre' RightPushforwardAlong_ObjectOf_MorphismOf_Pre'' RightPushforwardAlong_ObjectOf_ObjectOf Index2Cat] in
         @RightPushforwardAlong_ObjectOf_MorphismOf_Pre'' g s d m.
 
@@ -704,7 +704,7 @@
                     || (abstract apply CosliceCategoryInducedFunctor_FIdentityOf)
                     || (abstract apply SliceCategoryInducedFunctor_FCompositionOf)
                     || (abstract apply CosliceCategoryInducedFunctor_FCompositionOf)
-                    || apply CommaSpecializedCategory_Morphism_eq
+                    || apply CommaCategory_Morphism_eq
                     || apply f_equal
                     || (apply f_equal2; try reflexivity; [])
                     || apply sigT_eq (* simpl_eq *)
@@ -725,7 +725,7 @@
 
     Lemma RightPushforwardAlong_ObjectOf_FCompositionOf_Pre (g : S ^ C) s d d' (m1 : Morphism D s d) (m2 : Morphism D d d') :
       RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ (Compose m2 m1) =
-      Compose (C := LaxCosliceSpecializedCategory _ _) (RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m2) (RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m1).
+      Compose (C := LaxCosliceCategory _ _) (RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m2) (RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m1).
     Proof.
       (* For speed temporarily *)
     Admitted. (*
@@ -736,7 +736,7 @@
 
     Lemma RightPushforwardAlong_ObjectOf_FIdentityOf_Pre (g : S ^ C) x :
       RightPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ (Identity x) =
-      Identity (C := LaxCosliceSpecializedCategory _ _) _.
+      Identity (C := LaxCosliceCategory _ _) _.
     Proof.
       unfold RightPushforwardAlong_ObjectOf_MorphismOf_Pre.
       Time pre_anihilate.
@@ -747,8 +747,8 @@
       Morphism S (RightPushforwardAlong_ObjectOf_ObjectOf g s) (RightPushforwardAlong_ObjectOf_ObjectOf g d).
       subst RightPushforwardAlong_ObjectOf_ObjectOf RightPushforwardAlong_ObjectOf_MorphismOf_Pre' RightPushforwardAlong_ObjectOf_MorphismOf_Pre''; simpl.
       apply (InducedLimitFunctor_MorphismOf (Index2Cat := Index2Cat) (D := S)
-        (s := existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : LaxCosliceSpecializedCategory_ObjectT _ _)
-        (d := existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : LaxCosliceSpecializedCategory_ObjectT _ _)
+        (s := existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : LaxCosliceCategory_ObjectT _ _)
+        (d := existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : LaxCosliceCategory_ObjectT _ _)
         (HasLimits g s)
         (HasLimits g d)
       ); simpl in *.
@@ -763,7 +763,7 @@
       unfold Object in X.
       simpl in X.
 (*      Set Printing All. *)
-      refine (Build_SpecializedFunctor D S
+      refine (Build_Functor D S
         (@RightPushforwardAlong_ObjectOf_ObjectOf g)
         (@RightPushforwardAlong_ObjectOf_MorphismOf g)
         _
@@ -779,9 +779,9 @@
                 intros s d d' m1 m2;
                   etransitivity;
                   try apply (InducedLimitFunctor_FCompositionOf (Index2Cat := Index2Cat) (D := S)
-                    (s := existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : LaxCosliceSpecializedCategory_ObjectT _ _)
-                    (d := existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : LaxCosliceSpecializedCategory_ObjectT _ _)
-                    (d' := existT _ (tt, d') (RightPushforwardAlong_pre_Functor g d') : LaxCosliceSpecializedCategory_ObjectT _ _)
+                    (s := existT _ (tt, s) (RightPushforwardAlong_pre_Functor g s) : LaxCosliceCategory_ObjectT _ _)
+                    (d := existT _ (tt, d) (RightPushforwardAlong_pre_Functor g d) : LaxCosliceCategory_ObjectT _ _)
+                    (d' := existT _ (tt, d') (RightPushforwardAlong_pre_Functor g d') : LaxCosliceCategory_ObjectT _ _)
                     (HasLimits g s)
                     (HasLimits g d)
                     (HasLimits g d')
@@ -792,7 +792,7 @@
                   intros x;
                     etransitivity;
                     try apply (InducedLimitFunctor_FIdentityOf (Index2Cat := Index2Cat) (D := S)
-                      (existT _ (tt, x) (RightPushforwardAlong_pre_Functor g x) : LaxCosliceSpecializedCategory_ObjectT _ _)
+                      (existT _ (tt, x) (RightPushforwardAlong_pre_Functor g x) : LaxCosliceCategory_ObjectT _ _)
                       (HasLimits g x)
                     ); []
               ];
@@ -801,7 +801,7 @@
             ).
     Defined.
 
-    Definition RightPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+    Definition RightPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
       NaturalTransformation
       (ComposeFunctors (RightPushforwardAlong_pre_Functor s c) (IdentityFunctor _))
       (RightPushforwardAlong_pre_Functor d c).
@@ -809,8 +809,8 @@
       eapply (NTComposeT (ComposeFunctorsAssociator1 _ _ _) _).
       Grab Existential Variables.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun x => m (snd (projT1 x)))
             _
           )
@@ -826,7 +826,7 @@
       ).
     Defined.
 
-    Definition RightPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+    Definition RightPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
       Morphism S ((RightPushforwardAlong_ObjectOf s) c) ((RightPushforwardAlong_ObjectOf d) c).
     Proof.
       simpl; subst_body; simpl.
@@ -834,8 +834,8 @@
       exact (@RightPushforwardAlong_MorphismOf_ComponentsOf_Pre s d m c).
     Defined.
 
-    Definition RightPushforwardAlong_MorphismOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) :
-      SpecializedNaturalTransformation (RightPushforwardAlong_ObjectOf s) (RightPushforwardAlong_ObjectOf d).
+    Definition RightPushforwardAlong_MorphismOf (s d : S ^ C) (m : NaturalTransformation s d) :
+      NaturalTransformation (RightPushforwardAlong_ObjectOf s) (RightPushforwardAlong_ObjectOf d).
     Proof.
       exists (@RightPushforwardAlong_MorphismOf_ComponentsOf s d m).
       present_spnt.
@@ -897,8 +897,8 @@
 (*
 
     change (diagonal_functor_morphism_of C D) with (@ΔMor _ _ _ _ C D) in *;
-    change (MorphismOf (DiagonalSpecializedFunctor C D)) with (@ΔMor _ _ _ _ C D) in *;
-    change (ObjectOf (DiagonalSpecializedFunctor C D)) with (@Δ _ _ _ _ C D) in *;
+    change (MorphismOf (DiagonalFunctor C D)) with (@ΔMor _ _ _ _ C D) in *;
+    change (ObjectOf (DiagonalFunctor C D)) with (@Δ _ _ _ _ C D) in *;
     change InitialMorphism_Object with colimo in *;
     change InitialMorphism_Morphism with φ in *;
     change @InitialProperty_Morphism with @unique_m in *.*)
@@ -1049,8 +1049,8 @@
       apply (NTComposeF
       exists (ComponentsOf' m).
       assert (forall
-         c0 : CommaSpecializedCategory_Object
-                (SliceSpecializedCategory_Functor D c) F,
+         c0 : CommaCategory_Object
+                (SliceCategory_Functor D c) F,
        CMorphism S
          (ObjectOf
             (ComposeFunctors (RightPushforwardAlong_pre_Functor s c)
@@ -1059,8 +1059,8 @@
       present_spnt.
       intro c0; destruct c0 as [ [ [ [] c0 ] cm ] ]; simpl in *.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun _ => Identity _)
             _
           )
@@ -1078,14 +1078,14 @@
 
 
       Set Printing All.
-      Check LaxCosliceSpecializedCategory_Object LSObject LSMorphism
+      Check LaxCosliceCategory_Object LSObject LSMorphism
                 LSUnderlyingCategory S.
       specialize (s0 (fun c => (HasLimits (projT2 c)))).
       specialize
 
       simpl in *.
       S ^ D.
-      refine (Build_SpecializedFunctor D S
+      refine (Build_Functor D S
         (@RightPushforwardAlong_ObjectOf_ObjectOf g)
         (@RightPushforwardAlong_ObjectOf_MorphismOf g)
         _
@@ -1114,7 +1114,7 @@
       snt_eq
       destruct x; try reflexivity.
 
-      assert (forall (C : SmallCategory) D (F G : Functor C D) (T : SmallNaturalTransformation F G), T = {| SComponentsOf := SComponentsOf T; SCommutes := SCommutes T |}).
+      assert (forall (C : Category) D (F G : Functor C D) (T : SmallNaturalTransformation F G), T = {| SComponentsOf := SComponentsOf T; SCommutes := SCommutes T |}).
       let T := fresh in intros ? ? ? ? T; destruct T; simpl; reflexivity.
       symmetry.
       rewrite (H .
@@ -1125,10 +1125,10 @@
 *)
     Defined.
 
-    Definition RightPushforwardAlong : SpecializedFunctor (S ^ C) (S ^ D).
+    Definition RightPushforwardAlong : Functor (S ^ C) (S ^ D).
       match goal with
-          | [ |- SpecializedFunctor ?C ?D ] =>
-            refine (Build_SpecializedFunctor
+          | [ |- Functor ?C ?D ] =>
+            refine (Build_Functor
                       C D
                       RightPushforwardAlong_ObjectOf
                       RightPushforwardAlong_MorphismOf
@@ -1156,8 +1156,8 @@
   End Π.
 
   Section Σ.
-    Local Notation "F ↓ A" := (SliceSpecializedCategory A F).
-    (*Local Notation "C / c" := (@SliceSpecializedCategoryOver _ _ C c).*)
+    Local Notation "F ↓ A" := (SliceCategory A F).
+    (*Local Notation "C / c" := (@SliceCategoryOver _ _ C c).*)
 
     (** Quoting David Spivak in "Functorial Data Migration":
        Definition 2.1.3. Let [F : C -> D] be a morphism of schemas and
@@ -1187,20 +1187,20 @@
        *)
 
     (* Define [ɣ ○ (π_F d)] *)
-    Definition LeftPushforwardAlong_pre_Functor (g : S ^ C) (d : D) : SpecializedFunctor (F ↓ d) S
+    Definition LeftPushforwardAlong_pre_Functor (g : S ^ C) (d : D) : Functor (F ↓ d) S
       := ComposeFunctors g (projT2 (SliceCategoryProjectionFunctor C D F d)).
 
     Variable HasColimits : forall g d, Colimit (LeftPushforwardAlong_pre_Functor g d).
 
     Let Index2Cat d := F ↓ d.
 
-    Local Notation "'CAT' ⇓ D" := (@LaxSliceSpecializedCategory _ _ Index2Cat _ D).
+    Local Notation "'CAT' ⇓ D" := (@LaxSliceCategory _ _ Index2Cat _ D).
 
     Let LeftPushforwardAlong_ObjectOf_ObjectOf (g : S ^ C) d := ColimitObject (HasColimits g d).
 
     Let LeftPushforwardAlong_ObjectOf_MorphismOf_Pre' (g : S ^ C) s d (m : Morphism D s d) :
-      {F0 : SpecializedFunctor (F ↓ s) (F ↓ d) * unit &
-        SpecializedNaturalTransformation
+      {F0 : Functor (F ↓ s) (F ↓ d) * unit &
+        NaturalTransformation
         (LeftPushforwardAlong_pre_Functor g s)
         (ComposeFunctors (LeftPushforwardAlong_pre_Functor g d) (fst F0)) }.
       exists (fst (proj1_sig (MorphismOf (SliceCategoryProjectionFunctor C D F) m)), tt).
@@ -1209,15 +1209,15 @@
           simpl;
             unfold Object, Morphism, GeneralizeFunctor.
       match goal with
-        | [ |- SpecializedNaturalTransformation _ (ComposeFunctors (ComposeFunctors ?g ?h) ?i) ] =>
+        | [ |- NaturalTransformation _ (ComposeFunctors (ComposeFunctors ?g ?h) ?i) ] =>
           eapply (NTComposeT (ComposeFunctorsAssociator2 g h i) _)
       end.
       Grab Existential Variables.
       eapply (NTComposeF (IdentityNaturalTransformation g) _).
       Grab Existential Variables.
       match goal with
-        | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ C : _ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun _ => Identity (C := C) _)
             _
           )
@@ -1232,20 +1232,20 @@
 
     Let LeftPushforwardAlong_ObjectOf_MorphismOf_Pre'' (g : S ^ C) s d (m : Morphism D s d) :
       Morphism (CAT ⇓ S)
-      (existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : @LaxSliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
-      (existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : @LaxSliceSpecializedCategory_ObjectT _ _ Index2Cat _ _).
+      (existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : @LaxSliceCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : @LaxSliceCategory_ObjectT _ _ Index2Cat _ _).
     Proof.
       hnf.
       match goal with
-        | [ |- LaxSliceSpecializedCategory_Morphism ?a ?b ] =>
-          exact (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre' g _ _ m : @LaxSliceSpecializedCategory_MorphismT _ _ _ _ _ a b)
+        | [ |- LaxSliceCategory_Morphism ?a ?b ] =>
+          exact (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre' g _ _ m : @LaxSliceCategory_MorphismT _ _ _ _ _ a b)
       end.
     Defined.
 
     Definition LeftPushforwardAlong_ObjectOf_MorphismOf_Pre (g : S ^ C) s d (m : Morphism D s d) :
       Morphism (CAT ⇓ S)
-      (existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : @LaxSliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
-      (existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : @LaxSliceSpecializedCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : @LaxSliceCategory_ObjectT _ _ Index2Cat _ _)
+      (existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : @LaxSliceCategory_ObjectT _ _ Index2Cat _ _)
       := Eval cbv beta iota zeta delta [LeftPushforwardAlong_ObjectOf_MorphismOf_Pre' LeftPushforwardAlong_ObjectOf_MorphismOf_Pre'' LeftPushforwardAlong_ObjectOf_ObjectOf Index2Cat] in
         @LeftPushforwardAlong_ObjectOf_MorphismOf_Pre'' g s d m.
 
@@ -1256,7 +1256,7 @@
                     || (abstract apply CosliceCategoryInducedFunctor_FIdentityOf)
                     || (abstract apply SliceCategoryInducedFunctor_FCompositionOf)
                     || (abstract apply CosliceCategoryInducedFunctor_FCompositionOf)
-                    || apply CommaSpecializedCategory_Morphism_eq
+                    || apply CommaCategory_Morphism_eq
                     || apply f_equal
                     || (apply f_equal2; try reflexivity; [])
                     || apply sigT_eq (* simpl_eq *)
@@ -1277,7 +1277,7 @@
 
     Lemma LeftPushforwardAlong_ObjectOf_FCompositionOf_Pre (g : S ^ C) s d d' (m1 : Morphism D s d) (m2 : Morphism D d d') :
       LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ (Compose m2 m1) =
-      Compose (C := LaxSliceSpecializedCategory _ _ _) (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m2) (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m1).
+      Compose (C := LaxSliceCategory _ _ _) (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m2) (LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ m1).
     Proof.
       (* For speed temporarily *)
     Admitted. (*
@@ -1288,7 +1288,7 @@
 
     Lemma LeftPushforwardAlong_ObjectOf_FIdentityOf_Pre (g : S ^ C) x :
       LeftPushforwardAlong_ObjectOf_MorphismOf_Pre g _ _ (Identity x) =
-      Identity (C := LaxSliceSpecializedCategory _ _ _) _.
+      Identity (C := LaxSliceCategory _ _ _) _.
     Proof.
       unfold LeftPushforwardAlong_ObjectOf_MorphismOf_Pre.
       Time pre_anihilate.
@@ -1299,8 +1299,8 @@
       Morphism S (LeftPushforwardAlong_ObjectOf_ObjectOf g s) (LeftPushforwardAlong_ObjectOf_ObjectOf g d).
       subst LeftPushforwardAlong_ObjectOf_ObjectOf LeftPushforwardAlong_ObjectOf_MorphismOf_Pre' LeftPushforwardAlong_ObjectOf_MorphismOf_Pre''; simpl.
       apply (InducedColimitFunctor_MorphismOf (Index2Cat := Index2Cat) (D := S)
-        (s := existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : LaxSliceSpecializedCategory_ObjectT _ _ _)
-        (d := existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : LaxSliceSpecializedCategory_ObjectT _ _ _)
+        (s := existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : LaxSliceCategory_ObjectT _ _ _)
+        (d := existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : LaxSliceCategory_ObjectT _ _ _)
         (HasColimits g s)
         (HasColimits g d)
       ); simpl in *.
@@ -1310,7 +1310,7 @@
     Hint Resolve LeftPushforwardAlong_ObjectOf_FIdentityOf_Pre LeftPushforwardAlong_ObjectOf_FCompositionOf_Pre.
 
     Definition LeftPushforwardAlong_ObjectOf (g : S ^ C) : S ^ D.
-      refine (Build_SpecializedFunctor
+      refine (Build_Functor
                 D
                 S
                 (@LeftPushforwardAlong_ObjectOf_ObjectOf g)
@@ -1331,9 +1331,9 @@
                 try apply (InducedColimitFunctor_FCompositionOf
                              (Index2Cat := Index2Cat)
                              (D := S)
-                             (s := existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : LaxSliceSpecializedCategory_ObjectT _ _ _)
-                             (d := existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : LaxSliceSpecializedCategory_ObjectT _ _ _)
-                             (d' := existT _ (d', tt) (LeftPushforwardAlong_pre_Functor g d') : LaxSliceSpecializedCategory_ObjectT _ _ _)
+                             (s := existT _ (s, tt) (LeftPushforwardAlong_pre_Functor g s) : LaxSliceCategory_ObjectT _ _ _)
+                             (d := existT _ (d, tt) (LeftPushforwardAlong_pre_Functor g d) : LaxSliceCategory_ObjectT _ _ _)
+                             (d' := existT _ (d', tt) (LeftPushforwardAlong_pre_Functor g d') : LaxSliceCategory_ObjectT _ _ _)
                              (HasColimits g s)
                              (HasColimits g d)
                              (HasColimits g d')
@@ -1346,7 +1346,7 @@
                 try apply (InducedColimitFunctor_FIdentityOf
                              (Index2Cat := Index2Cat)
                              (D := S)
-                             (existT _ (x, tt) (LeftPushforwardAlong_pre_Functor g x) : LaxSliceSpecializedCategory_ObjectT _ _ _)
+                             (existT _ (x, tt) (LeftPushforwardAlong_pre_Functor g x) : LaxSliceCategory_ObjectT _ _ _)
                              (HasColimits g x)
                           ); []
               ];
@@ -1355,7 +1355,7 @@
           ).
     Defined.
 
-    Definition LeftPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+    Definition LeftPushforwardAlong_MorphismOf_ComponentsOf_Pre (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
       NaturalTransformation
         (LeftPushforwardAlong_pre_Functor s c)
         (ComposeFunctors (LeftPushforwardAlong_pre_Functor d c) (IdentityFunctor _)).
@@ -1363,8 +1363,8 @@
       eapply (NTComposeT _ (ComposeFunctorsAssociator2 _ _ _)).
       Grab Existential Variables.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation
                     F
                     G
                     (fun x => m (fst (projT1 x)))
@@ -1382,7 +1382,7 @@
           ).
     Defined.
 
-    Definition LeftPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) (c : D) :
+    Definition LeftPushforwardAlong_MorphismOf_ComponentsOf (s d : S ^ C) (m : NaturalTransformation s d) (c : D) :
       Morphism S ((LeftPushforwardAlong_ObjectOf s) c) ((LeftPushforwardAlong_ObjectOf d) c).
     Proof.
       simpl; subst_body; simpl.
@@ -1390,8 +1390,8 @@
       exact (@LeftPushforwardAlong_MorphismOf_ComponentsOf_Pre s d m c).
     Defined.
 
-    Definition LeftPushforwardAlong_MorphismOf (s d : S ^ C) (m : SpecializedNaturalTransformation s d) :
-      SpecializedNaturalTransformation (LeftPushforwardAlong_ObjectOf s) (LeftPushforwardAlong_ObjectOf d).
+    Definition LeftPushforwardAlong_MorphismOf (s d : S ^ C) (m : NaturalTransformation s d) :
+      NaturalTransformation (LeftPushforwardAlong_ObjectOf s) (LeftPushforwardAlong_ObjectOf d).
     Proof.
       exists (@LeftPushforwardAlong_MorphismOf_ComponentsOf s d m).
       present_spnt.
@@ -1401,10 +1401,10 @@
       admit.
     Defined.
 
-    Definition LeftPushforwardAlong : SpecializedFunctor (S ^ C) (S ^ D).
+    Definition LeftPushforwardAlong : Functor (S ^ C) (S ^ D).
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor
                     C D
                     LeftPushforwardAlong_ObjectOf
                     LeftPushforwardAlong_MorphismOf
diff -r 4eb6407c6ca3 DecidableComputableCategory.v
--- a/DecidableComputableCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/DecidableComputableCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,11 +10,11 @@
 
 Section ComputableCategory.
   Variable I : Type.
-  Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+  Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
-  Let eq_dec_on_cat `(C : @SpecializedCategory objC) := forall x y : objC, {x = y} + {x <> y}.
+  Let eq_dec_on_cat `(C : @Category objC) := forall x y : objC, {x = y} + {x <> y}.
 
-  Definition ComputableCategoryDec := @SpecializedCategory_sigT_obj _ (@ComputableCategory _ _ Index2Cat) (fun C => eq_dec_on_cat C).
+  Definition ComputableCategoryDec := @Category_sigT_obj _ (@ComputableCategory _ _ Index2Cat) (fun C => eq_dec_on_cat C).
 End ComputableCategory.
diff -r 4eb6407c6ca3 DecidableDiscreteCategory.v
--- a/DecidableDiscreteCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/DecidableDiscreteCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import Common Notations.
 
 Set Implicit Arguments.
@@ -25,12 +25,12 @@
   Let DiscreteCategoryDec_Morphism s d : Set := if s == d then unit else Empty_set.
 
   Local Ltac simpl_eq_dec := subst_body; simpl in *; intros;
-  (*    unfold eq_b in *;*)
-    repeat match goal with
-             | [ _ : context[eq_dec ?a ?b] |- _ ] => destruct (eq_dec a b); try contradict_by_transitivity
-             | [ |- context[eq_dec ?a ?b] ] => destruct (eq_dec a b); try contradict_by_transitivity
-           end;
-      auto.
+                             (*    unfold eq_b in *;*)
+                             repeat match goal with
+                                      | [ _ : context[eq_dec ?a ?b] |- _ ] => destruct (eq_dec a b); try contradict_by_transitivity
+                                      | [ |- context[eq_dec ?a ?b] ] => destruct (eq_dec a b); try contradict_by_transitivity
+                                    end;
+                             auto.
 
   Definition DiscreteCategoryDec_Compose (s d d' : O) (m : DiscreteCategoryDec_Morphism d d') (m' : DiscreteCategoryDec_Morphism s d) :
     DiscreteCategoryDec_Morphism s d'.
@@ -41,14 +41,15 @@
     simpl_eq_dec.
   Defined.
 
-  Definition DiscreteCategoryDec : @SpecializedCategory O.
-    refine (@Build_SpecializedCategory _
-                                       DiscreteCategoryDec_Morphism
-                                       DiscreteCategoryDec_Identity
-                                       DiscreteCategoryDec_Compose
-                                       _
-                                       _
-                                       _);
+  Definition DiscreteCategoryDec : Category.
+    refine (@Build_Category O
+                            DiscreteCategoryDec_Morphism
+                            DiscreteCategoryDec_Identity
+                            DiscreteCategoryDec_Compose
+                            _
+                            _
+                            _);
+
     abstract (
         unfold DiscreteCategoryDec_Compose, DiscreteCategoryDec_Identity;
         simpl_eq_dec
diff -r 4eb6407c6ca3 DecidableDiscreteCategoryFunctors.v
--- a/DecidableDiscreteCategoryFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/DecidableDiscreteCategoryFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality Eqdep_dec ProofIrrelevance JMeq.
-Require Export DiscreteCategoryFunctors DecidableDiscreteCategory DecidableSetCategory DecidableComputableCategory DecidableSmallCat.
+Require Export DiscreteCategoryFunctors DecidableDiscreteCategory DecidableSetCategory DecidableComputableCategory DecidableCat.
 Require Import Common Adjoint.
 
 Set Implicit Arguments.
@@ -19,8 +19,8 @@
 Section Obj.
   Local Ltac build_ob_functor Index2Object :=
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun C' => existT _ (Index2Object (projT1 C')) (projT2 C'))
           (fun _ _ F => ObjectOf F)
           _
@@ -32,9 +32,9 @@
   Section type.
     Variable I : Type.
     Variable Index2Object : I -> Type.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctorDec : SpecializedFunctor (@ComputableCategoryDec _ _ Index2Cat) TypeCatDec.
+    Definition ObjectFunctorDec : Functor (@ComputableCategoryDec _ _ Index2Cat) TypeCatDec.
       build_ob_functor Index2Object.
     Defined.
   End type.
@@ -42,9 +42,9 @@
   Section set.
     Variable I : Type.
     Variable Index2Object : I -> Set.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctorToSetDec : SpecializedFunctor (@ComputableCategoryDec _ _ Index2Cat) SetCatDec.
+    Definition ObjectFunctorToSetDec : Functor (@ComputableCategoryDec _ _ Index2Cat) SetCatDec.
       build_ob_functor Index2Object.
     Defined.
   End set.
@@ -52,9 +52,9 @@
   Section prop.
     Variable I : Type.
     Variable Index2Object : I -> Prop.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctorToPropDec : SpecializedFunctor (@ComputableCategoryDec _ _ Index2Cat) PropCatDec.
+    Definition ObjectFunctorToPropDec : Functor (@ComputableCategoryDec _ _ Index2Cat) PropCatDec.
       build_ob_functor Index2Object.
     Defined.
   End prop.
@@ -66,7 +66,7 @@
 
 Section InducedFunctor.
   Variable O : Type.
-  Context `(O' : @SpecializedCategory obj).
+  Context `(O' : @Category obj).
   Variable f : O -> O'.
   Variable eq_dec : forall x y : O, {x = y} + {x <> y}.
 
@@ -101,10 +101,10 @@
 
   Local Arguments Compose {obj} [C s d d'] / _ _ : rename, simpl nomatch.
 
-  Definition InducedDiscreteFunctorDec : SpecializedFunctor (DiscreteCategoryDec eq_dec) O'.
+  Definition InducedDiscreteFunctorDec : Functor (DiscreteCategoryDec eq_dec) O'.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           f
           InducedDiscreteFunctorDec_MorphismOf
           _
@@ -143,11 +143,11 @@
   Hint Unfold DiscreteCategoryDec_Identity.
   Hint Unfold eq_rect_r eq_rect eq_sym.
 
-  Definition DiscreteFunctorDec : SpecializedFunctor TypeCatDec LocallySmallCatDec.
-    refine (Build_SpecializedFunctor TypeCatDec LocallySmallCatDec
+  Definition DiscreteFunctorDec : Functor TypeCatDec LocallyCatDec.
+    refine (Build_Functor TypeCatDec LocallyCatDec
       (fun O => existT
-        (fun C : LocallySmallCategory => forall x y : C, {x = y} + {x <> y})
-        (DiscreteCategoryDec (projT2 O) : LocallySmallSpecializedCategory _)
+        (fun C : LocallyCategory => forall x y : C, {x = y} + {x <> y})
+        (DiscreteCategoryDec (projT2 O) : LocallyCategory _)
         (fun x y => projT2 O x y))
       (fun s d f => InducedDiscreteFunctorDec _ f (projT2 s))
       _
@@ -227,11 +227,11 @@
     admit.
   Defined.
 
-  Definition DiscreteSetFunctorDec : SpecializedFunctor SetCatDec SmallCatDec.
-    refine (Build_SpecializedFunctor SetCatDec SmallCatDec
+  Definition DiscreteSetFunctorDec : Functor SetCatDec CatDec.
+    refine (Build_Functor SetCatDec CatDec
       (fun O => existT
-        (fun C : SmallCategory => forall x y : C, {x = y} + {x <> y})
-        (DiscreteCategoryDec (projT2 O) : SmallSpecializedCategory _)
+        (fun C : Category => forall x y : C, {x = y} + {x <> y})
+        (DiscreteCategoryDec (projT2 O) : Category _)
         (fun x y => projT2 O x y))
       (fun s d f => InducedDiscreteFunctorDec _ f (projT2 s))
       _
diff -r 4eb6407c6ca3 DecidableSetCategory.v
--- a/DecidableSetCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/DecidableSetCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -11,7 +11,7 @@
 Section CSet.
   Let eq_dec_on T := forall x y : T, {x = y} + {x <> y}.
 
-  Definition TypeCatDec := SpecializedCategory_sigT_obj TypeCat eq_dec_on.
-  Definition SetCatDec := SpecializedCategory_sigT_obj SetCat eq_dec_on.
-  Definition PropCatDec := SpecializedCategory_sigT_obj PropCat eq_dec_on.
+  Definition TypeCatDec := Category_sigT_obj TypeCat eq_dec_on.
+  Definition SetCatDec := Category_sigT_obj SetCat eq_dec_on.
+  Definition PropCatDec := Category_sigT_obj PropCat eq_dec_on.
 End CSet.
diff -r 4eb6407c6ca3 DecidableSmallCat.v
--- a/DecidableSmallCat.v	Fri May 24 20:09:37 2013 -0400
+++ b/DecidableSmallCat.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SmallCat SigTCategory.
+Require Export Cat SigTCategory.
 
 Set Implicit Arguments.
 
@@ -8,9 +8,9 @@
 
 Set Universe Polymorphism.
 
-Section SmallCat.
-  Let eq_dec_on_cat `(C : @SpecializedCategory objC) := forall x y : objC, {x = y} + {x <> y}.
+Section Cat.
+  Let eq_dec_on_cat `(C : @Category objC) := forall x y : objC, {x = y} + {x <> y}.
 
-  Definition SmallCatDec := SpecializedCategory_sigT_obj SmallCat (fun C => eq_dec_on_cat C).
-  Definition LocallySmallCatDec := SpecializedCategory_sigT_obj LocallySmallCat (fun C => eq_dec_on_cat C).
-End SmallCat.
+  Definition CatDec := Category_sigT_obj Cat (fun C => eq_dec_on_cat C).
+  Definition LocallyCatDec := Category_sigT_obj LocallyCat (fun C => eq_dec_on_cat C).
+End Cat.
diff -r 4eb6407c6ca3 DiscreteCategory.v
--- a/DiscreteCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/DiscreteCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -10,37 +10,38 @@
 Section DCategory.
   Variable O : Type.
 
-  Local Ltac simpl_eq := subst_body; hnf in *; simpl in *; intros; destruct_type @inhabited; simpl in *;
+  Local Ltac simpl_disc :=
+    subst_body; hnf in *; simpl in *; intros; destruct_type @inhabited; simpl in *;
     repeat constructor;
-      repeat subst;
-        auto;
-          simpl_transitivity.
+    repeat subst;
+    auto;
+    simpl_transitivity.
 
   Let DiscreteCategory_Morphism (s d : O) := s = d.
 
   Definition DiscreteCategory_Compose (s d d' : O) (m : DiscreteCategory_Morphism d d') (m' : DiscreteCategory_Morphism s d) :
     DiscreteCategory_Morphism s d'.
-    simpl_eq.
+    simpl_disc.
   Defined.
 
   Definition DiscreteCategory_Identity o : DiscreteCategory_Morphism o o.
-    simpl_eq.
+    simpl_disc.
   Defined.
 
-  Global Arguments DiscreteCategory_Compose [s d d'] m m' /.
-  Global Arguments DiscreteCategory_Identity o /.
+  Global Arguments DiscreteCategory_Compose [s d d'] m m' / .
+  Global Arguments DiscreteCategory_Identity o / .
 
-  Definition DiscreteCategory : @SpecializedCategory O.
-    refine (@Build_SpecializedCategory _
-                                       DiscreteCategory_Morphism
-                                       DiscreteCategory_Identity
-                                       DiscreteCategory_Compose
-                                       _
-                                       _
-                                       _);
-    abstract (
-        unfold DiscreteCategory_Compose, DiscreteCategory_Identity;
-        simpl_eq
-      ).
+  Definition DiscreteCategory : Category.
+    refine (@Build_Category O
+                            DiscreteCategory_Morphism
+                            DiscreteCategory_Identity
+                            DiscreteCategory_Compose
+                            _
+                            _
+                            _);
+    simpl;
+    compute;
+    repeat (intros [] || intro);
+    reflexivity.
   Defined.
 End DCategory.
diff -r 4eb6407c6ca3 DiscreteCategoryFunctors.v
--- a/DiscreteCategoryFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/DiscreteCategoryFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export DiscreteCategory Functor SetCategory ComputableCategory SmallCat NaturalTransformation.
+Require Export DiscreteCategory Functor SetCategory ComputableCategory Cat NaturalTransformation.
 Require Import Common Adjoint.
 
 Set Implicit Arguments.
@@ -12,7 +12,7 @@
 
 Section FunctorFromDiscrete.
   Variable O : Type.
-  Context `(D : @SpecializedCategory objD).
+  Variable D : Category.
   Variable objOf : O -> objD.
 
   Let FunctorFromDiscrete_MorphismOf s d (m : Morphism (DiscreteCategory O) s d) : Morphism D (objOf s) (objOf d)
@@ -20,7 +20,7 @@
          | eq_refl => Identity _
        end.
 
-  Definition FunctorFromDiscrete : SpecializedFunctor (DiscreteCategory O) D.
+  Definition FunctorFromDiscrete : Functor (DiscreteCategory O) D.
   Proof.
     refine {| ObjectOf := objOf; MorphismOf := FunctorFromDiscrete_MorphismOf |};
       abstract (
@@ -33,8 +33,8 @@
 Section Obj.
   Local Ltac build_ob_functor Index2Object :=
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun C' => Index2Object C')
           (fun _ _ F => ObjectOf F)
           _
@@ -46,9 +46,9 @@
   Section type.
     Variable I : Type.
     Variable Index2Object : I -> Type.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctor : SpecializedFunctor (@ComputableCategory _ _ Index2Cat) TypeCat.
+    Definition ObjectFunctor : Functor (@ComputableCategory _ _ Index2Cat) TypeCat.
       build_ob_functor Index2Object.
     Defined.
   End type.
@@ -56,9 +56,9 @@
   Section set.
     Variable I : Type.
     Variable Index2Object : I -> Set.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctorToSet : SpecializedFunctor (@ComputableCategory _ _ Index2Cat) SetCat.
+    Definition ObjectFunctorToSet : Functor (@ComputableCategory _ _ Index2Cat) SetCat.
       build_ob_functor Index2Object.
     Defined.
   End set.
@@ -66,9 +66,9 @@
   Section prop.
     Variable I : Type.
     Variable Index2Object : I -> Prop.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition ObjectFunctorToProp : SpecializedFunctor (@ComputableCategory _ _ Index2Cat) PropCat.
+    Definition ObjectFunctorToProp : Functor (@ComputableCategory _ _ Index2Cat) PropCat.
       build_ob_functor Index2Object.
     Defined.
   End prop.
@@ -81,8 +81,8 @@
 Section Mor.
   Local Ltac build_mor_functor Index2Object Index2Cat :=
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun C' => { sd : Index2Object C' * Index2Object C' & (Index2Cat C').(Morphism) (fst sd) (snd sd) } )
           (fun _ _ F => (fun sdm =>
             existT (fun sd => Morphism _ (fst sd) (snd sd))
@@ -101,9 +101,9 @@
   Section type.
     Variable I : Type.
     Variable Index2Object : I -> Type.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition MorphismFunctor : SpecializedFunctor (@ComputableCategory _ _ Index2Cat) TypeCat.
+    Definition MorphismFunctor : Functor (@ComputableCategory _ _ Index2Cat) TypeCat.
       build_mor_functor Index2Object Index2Cat.
     Defined.
   End type.
@@ -114,8 +114,8 @@
 Section dom_cod.
   Local Ltac build_dom_cod fst_snd :=
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun _ => (fun sdf => fst_snd _ _ (projT1 sdf)))
           _
         )
@@ -125,13 +125,13 @@
   Section type.
     Variable I : Type.
     Variable Index2Object : I -> Type.
-    Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+    Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-    Definition DomainNaturalTransformation : SpecializedNaturalTransformation (@MorphismFunctor _ _ Index2Cat) ObjectFunctor.
+    Definition DomainNaturalTransformation : NaturalTransformation (@MorphismFunctor _ _ Index2Cat) ObjectFunctor.
       build_dom_cod @fst.
     Defined.
 
-    Definition CodomainNaturalTransformation : SpecializedNaturalTransformation (@MorphismFunctor _ _ Index2Cat) ObjectFunctor.
+    Definition CodomainNaturalTransformation : NaturalTransformation (@MorphismFunctor _ _ Index2Cat) ObjectFunctor.
       build_dom_cod @snd.
     Defined.
   End type.
@@ -139,13 +139,13 @@
 
 Section InducedFunctor.
   Variable O : Type.
-  Context `(O' : @SpecializedCategory obj).
+  Context `(O' : @Category obj).
   Variable f : O -> O'.
 
-  Definition InducedDiscreteFunctor : SpecializedFunctor (DiscreteCategory O) O'.
+  Definition InducedDiscreteFunctor : Functor (DiscreteCategory O) O'.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           f
           (fun s d_ m => match m return _ with eq_refl => Identity (f s) end)
           _
@@ -167,9 +167,9 @@
       hnf in *; subst;
         auto.
 
-  Definition DiscreteFunctor : SpecializedFunctor TypeCat LocallySmallCat.
-    refine (Build_SpecializedFunctor TypeCat LocallySmallCat
-      (fun O => DiscreteCategory O : LocallySmallSpecializedCategory _)
+  Definition DiscreteFunctor : Functor TypeCat LocallyCat.
+    refine (Build_Functor TypeCat LocallyCat
+      (fun O => DiscreteCategory O : LocallyCategory _)
       (fun s d f => InducedDiscreteFunctor _ f)
       _
       _
@@ -177,9 +177,9 @@
     abstract t.
   Defined.
 
-  Definition DiscreteSetFunctor : SpecializedFunctor SetCat SmallCat.
-    refine (Build_SpecializedFunctor SetCat SmallCat
-      (fun O => DiscreteCategory O : SmallSpecializedCategory _)
+  Definition DiscreteSetFunctor : Functor SetCat Cat.
+    refine (Build_Functor SetCat Cat
+      (fun O => DiscreteCategory O : Category _)
       (fun s d f => InducedDiscreteFunctor _ f)
       _
       _
diff -r 4eb6407c6ca3 DualFunctor.v
--- a/DualFunctor.v	Fri May 24 20:09:37 2013 -0400
+++ b/DualFunctor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export Duals SmallCat.
+Require Export Duals Cat.
 
 Set Implicit Arguments.
 
@@ -7,18 +7,18 @@
 Set Universe Polymorphism.
 
 Section OppositeCategory.
-  Definition SmallOppositeFunctor : SpecializedFunctor SmallCat SmallCat.
-    refine (Build_SpecializedFunctor SmallCat SmallCat
-                                     (fun x => OppositeCategory x : SmallSpecializedCategory _)
+  Definition SmallOppositeFunctor : Functor Cat Cat.
+    refine (Build_Functor Cat Cat
+                                     (fun x => OppositeCategory x : Category _)
                                      (fun _ _ m => OppositeFunctor m)
                                      _
                                      _);
     simpl; abstract functor_eq.
   Defined.
 
-  Definition LocallySmallOppositeFunctor : SpecializedFunctor LocallySmallCat LocallySmallCat.
-    refine (Build_SpecializedFunctor LocallySmallCat LocallySmallCat
-                                     (fun x => OppositeCategory x : LocallySmallSpecializedCategory _)
+  Definition LocallySmallOppositeFunctor : Functor LocallyCat LocallyCat.
+    refine (Build_Functor LocallyCat LocallyCat
+                                     (fun x => OppositeCategory x : LocallyCategory _)
                                      (fun _ _ m => OppositeFunctor m)
                                      _
                                      _);
diff -r 4eb6407c6ca3 Duals.v
--- a/Duals.v	Fri May 24 20:09:37 2013 -0400
+++ b/Duals.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq Eqdep.
-Require Export SpecializedCategory CategoryIsomorphisms Functor ProductCategory NaturalTransformation.
+Require Export Category CategoryIsomorphisms Functor ProductCategory NaturalTransformation.
 Require Import Common Notations FEqualDep.
 
 Set Implicit Arguments.
@@ -15,39 +15,38 @@
 Local Open Scope category_scope.
 
 Section OppositeCategory.
-  Definition OppositeCategory `(C : @SpecializedCategory objC) : @SpecializedCategory objC
-    := @Build_SpecializedCategory' objC
-                                (fun s d => Morphism C d s)
-                                (Identity (C := C))
-                                (fun _ _ _ m1 m2 => Compose m2 m1)
-                                (fun _ _ _ _ _ _ _ => @Associativity_sym _ _ _ _ _ _ _ _ _)
-                                (fun _ _ _ _ _ _ _ => @Associativity _ _ _ _ _ _ _ _ _)
-                                (fun _ _ => @RightIdentity _ _ _ _)
-                                (fun _ _ => @LeftIdentity _ _ _ _).
+  Definition OppositeCategory (C : Category) : Category
+    := @Build_Category' C
+                        (fun s d => Morphism C d s)
+                        (Identity (C := C))
+                        (fun _ _ _ m1 m2 => Compose m2 m1)
+                        (fun _ _ _ _ _ _ _ => @Associativity_sym _ _ _ _ _ _ _ _)
+                        (fun _ _ _ _ _ _ _ => @Associativity _ _ _ _ _ _ _ _)
+                        (fun _ _ => @RightIdentity _ _ _)
+                        (fun _ _ => @LeftIdentity _ _ _).
 End OppositeCategory.
 
 (*Notation "C ᵒᵖ" := (OppositeCategory C) : category_scope.*)
 
 Section DualCategories.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
   Lemma op_op_id : OppositeCategory (OppositeCategory C) = C.
-    clear D objD.
+    clear D.
     unfold OppositeCategory; simpl.
     repeat change (fun a => ?f a) with f.
     destruct C; intros; simpl; reflexivity.
   Qed.
 
   Lemma op_distribute_prod : OppositeCategory (C * D) = (OppositeCategory C) * (OppositeCategory D).
-    spcat_eq.
+    cat_eq.
   Qed.
 End DualCategories.
 
 Hint Rewrite @op_op_id @op_distribute_prod : category.
 
 Section DualObjects.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Definition terminal_opposite_initial (o : C) : IsTerminalObject o -> IsInitialObject (C := OppositeCategory C) o
     := fun H o' => H o'.
@@ -65,28 +64,26 @@
 End DualObjects.
 
 Section OppositeFunctor.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
+  Variables C D : Category.
+  Variable F : Functor C D.
   Let COp := OppositeCategory C.
   Let DOp := OppositeCategory D.
 
-  Definition OppositeFunctor : SpecializedFunctor COp DOp.
-    refine (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)
-      (fun d' d s m1 m2 => FCompositionOf F s d d' m2 m1)
-      (FIdentityOf F)
-    ).
+  Definition OppositeFunctor : Functor COp DOp.
+    refine (Build_Functor COp DOp
+                          (fun c : COp => F c : DOp)
+                          (fun (s d : COp) (m : C.(Morphism) d s) => MorphismOf F (s := d) (d := s) m)
+                          (fun d' d s m1 m2 => FCompositionOf F s d d' m2 m1)
+                          (FIdentityOf F)
+           ).
   Defined.
 End OppositeFunctor.
 
 (*Notation "C ᵒᵖ" := (OppositeFunctor C) : functor_scope.*)
 
 Section OppositeFunctor_Id.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
+  Variables C D : Category.
+  Variable F : Functor C D.
 
   Lemma op_op_functor_id : OppositeFunctor (OppositeFunctor F) == F.
     functor_eq; autorewrite with category; trivial.
@@ -97,30 +94,29 @@
 (*Hint Rewrite op_op_functor_id.*)
 
 Section OppositeNaturalTransformation.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variables F G : SpecializedFunctor C D.
-  Variable T : SpecializedNaturalTransformation F G.
+  Variables C D : Category.
+  Variables F G : Functor C D.
+  Variable T : NaturalTransformation F G.
   Let COp := OppositeCategory C.
   Let DOp := OppositeCategory D.
   Let FOp := OppositeFunctor F.
   Let GOp := OppositeFunctor G.
 
-  Definition OppositeNaturalTransformation : SpecializedNaturalTransformation GOp FOp.
-    refine (Build_SpecializedNaturalTransformation GOp FOp
-      (fun c : COp => T.(ComponentsOf) c : DOp.(Morphism) (GOp c) (FOp c))
-      (fun s d m => eq_sym (Commutes T d s m))
-    ).
+  Definition OppositeNaturalTransformation : NaturalTransformation GOp FOp.
+    refine (Build_NaturalTransformation GOp FOp
+                                        (fun c : COp => T.(ComponentsOf) c : DOp.(Morphism) (GOp c) (FOp c))
+                                        (fun s d m => eq_sym (Commutes T d s m))
+           ).
   Defined.
 End OppositeNaturalTransformation.
 
 (*Notation "C ᵒᵖ" := (OppositeNaturalTransformation C) : natural_transformation_scope.*)
 
 Section OppositeNaturalTransformation_Id.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variables F G : SpecializedFunctor C D.
-  Variable T : SpecializedNaturalTransformation F G.
+  Variable C : Category.
+  Variable D : Category.
+  Variables F G : Functor C D.
+  Variable T : NaturalTransformation F G.
 
   Lemma op_op_nt_id : OppositeNaturalTransformation (OppositeNaturalTransformation T) == T.
     nt_eq; intros; try functor_eq; autorewrite with category; subst; trivial.
diff -r 4eb6407c6ca3 EnrichedCategory.v
--- a/EnrichedCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/EnrichedCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory MonoidalCategory.
+Require Export Category MonoidalCategory.
 Require Import Common Notations DefinitionSimplification.
 
 Set Implicit Arguments.
@@ -15,8 +15,8 @@
      *)
   Context `(M : @MonoidalCategory objM).
 
-  Let src `{C : @SpecializedCategory objC} s d (_ : Morphism C s d) := s.
-  Let dst `{C : @SpecializedCategory objC} s d (_ : Morphism C s d) := d.
+  Let src `{C : @Category objC} s d (_ : Morphism C s d) := s.
+  Let dst `{C : @Category objC} s d (_ : Morphism C s d) := d.
 
   Arguments src [objC C s d] _.
   Arguments dst [objC C s d] _.
@@ -273,7 +273,7 @@
             )
           ).
 
-    Let TensorProduct : SpecializedFunctor (m2Cat * m2Cat) m2Cat.
+    Let TensorProduct : Functor (m2Cat * m2Cat) m2Cat.
       eexists (fun _ => tt) (fun _ _ _ => _);
         repeat (unfold Morphism, Object; simpl);
           simpl; intros;
@@ -308,7 +308,7 @@
       define_unit_nt.
     Defined.
 
-    Definition m2MonoidalCat : @MonoidalCategory unit (fun _ _ => SpecializedFunctor TerminalCategory TerminalCategory).
+    Definition m2MonoidalCat : @MonoidalCategory unit (fun _ _ => Functor TerminalCategory TerminalCategory).
       refine (@Build_MonoidalCategory _ _
         m2Cat
         TensorProduct tt Associator LeftUnitor RightUnitor
diff -r 4eb6407c6ca3 Equalizer.v
--- a/Equalizer.v	Fri May 24 20:09:37 2013 -0400
+++ b/Equalizer.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,7 +10,7 @@
 Set Universe Polymorphism.
 
 Section Equalizer.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Variables A B : objC.
   Variables f g : C.(Morphism) A B.
 
@@ -31,8 +31,8 @@
     destruct s, d, d'; simpl in *; trivial.
   Defined.
 
-  Definition EqualizerIndex : @SpecializedCategory EqualizerTwo.
-    refine (@Build_SpecializedCategory _
+  Definition EqualizerIndex : @Category EqualizerTwo.
+    refine (@Build_Category _
                                        EqualizerIndex_Morphism
                                        (fun x => match x with EqualizerA => tt | EqualizerB => tt end)
                                        EqualizerIndex_Compose
@@ -62,10 +62,10 @@
               end.
   Defined.
 
-  Definition EqualizerDiagram : SpecializedFunctor EqualizerIndex C.
+  Definition EqualizerDiagram : Functor EqualizerIndex C.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           EqualizerDiagram_ObjectOf
           EqualizerDiagram_MorphismOf
           _
diff -r 4eb6407c6ca3 EqualizerFunctor.v
--- a/EqualizerFunctor.v	Fri May 24 20:09:37 2013 -0400
+++ b/EqualizerFunctor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -9,10 +9,10 @@
 Set Universe Polymorphism.
 
 Section Equalizer.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Variable HasLimits : forall F : SpecializedFunctor EqualizerIndex C, Limit F.
-  Variable HasColimits : forall F : SpecializedFunctor EqualizerIndex C, Colimit F.
+  Variable HasLimits : forall F : Functor EqualizerIndex C, Limit F.
+  Variable HasColimits : forall F : Functor EqualizerIndex C, Colimit F.
 
   Definition EqualizerFunctor := LimitFunctor HasLimits.
   Definition CoequalizerFunctor := ColimitFunctor HasColimits.
diff -r 4eb6407c6ca3 ExponentialLaws.v
--- a/ExponentialLaws.v	Fri May 24 20:09:37 2013 -0400
+++ b/ExponentialLaws.v	Fri Jun 21 15:07:08 2013 -0400
@@ -19,11 +19,11 @@
       NaturalTransformation_JMeq eq_JMeq.
 
 Section Law0.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition ExponentialLaw0Functor : SpecializedFunctor (C ^ 0) 1
+  Definition ExponentialLaw0Functor : Functor (C ^ 0) 1
     := FunctorTo1 _.
-  Definition ExponentialLaw0Functor_Inverse : SpecializedFunctor 1 (C ^ 0)
+  Definition ExponentialLaw0Functor_Inverse : Functor 1 (C ^ 0)
     := FunctorFrom1 _ (FunctorFrom0 _).
 
   Lemma ExponentialLaw0
@@ -40,17 +40,17 @@
 End Law0.
 
 Section Law0'.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Variable c : objC.
 
-  Definition ExponentialLaw0'Functor : SpecializedFunctor (0 ^ C) 0
-    := Build_SpecializedFunctor (0 ^ C) 0
+  Definition ExponentialLaw0'Functor : Functor (0 ^ C) 0
+    := Build_Functor (0 ^ C) 0
                                 (fun F => F c)
                                 (fun s d m => match (s c) with end)
                                 (fun x _ _ _ _ => match (x c) with end)
                                 (fun x => match (x c) with end).
 
-  Definition ExponentialLaw0'Functor_Inverse : SpecializedFunctor 0 (0 ^ C)
+  Definition ExponentialLaw0'Functor_Inverse : Functor 0 (0 ^ C)
     := FunctorFrom0 _.
 
   Lemma ExponentialLaw0'
@@ -69,10 +69,10 @@
 End Law0'.
 
 Section Law1.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition ExponentialLaw1Functor : SpecializedFunctor (C ^ 1) C
-    := Build_SpecializedFunctor (C ^ 1) C
+  Definition ExponentialLaw1Functor : Functor (C ^ 1) C
+    := Build_Functor (C ^ 1) C
                                 (fun F => F tt)
                                 (fun s d m => m tt)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -86,8 +86,8 @@
   Proof.
     hnf.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
                                                        (fun _ => m)
                                                        _
                )
@@ -101,9 +101,9 @@
 
   Global Arguments ExponentialLaw1Functor_Inverse_MorphismOf / _ _ _.
 
-  Definition ExponentialLaw1Functor_Inverse : SpecializedFunctor C (C ^ 1).
+  Definition ExponentialLaw1Functor_Inverse : Functor C (C ^ 1).
   Proof.
-    refine (Build_SpecializedFunctor
+    refine (Build_Functor
               C (C ^ 1)
               (@FunctorFrom1 _ _)
               ExponentialLaw1Functor_Inverse_MorphismOf
@@ -131,17 +131,17 @@
 End Law1.
 
 Section Law1'.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition ExponentialLaw1'Functor : SpecializedFunctor (1 ^ C) 1
+  Definition ExponentialLaw1'Functor : Functor (1 ^ C) 1
     := FunctorTo1 _.
 
-  Definition ExponentialLaw1'Functor_Inverse : SpecializedFunctor 1 (1 ^ C).
+  Definition ExponentialLaw1'Functor_Inverse : Functor 1 (1 ^ C).
   Proof.
-    refine (Build_SpecializedFunctor
+    refine (Build_Functor
               1 (1 ^ C)
               (fun _ => FunctorTo1 _)
-              (fun s d m => Build_SpecializedNaturalTransformation
+              (fun s d m => Build_NaturalTransformation
                               (FunctorTo1 C) (FunctorTo1 C)
                               (fun _ => Identity (C := 1) tt)
                               (fun _ _ _ => eq_refl))
@@ -164,21 +164,21 @@
 End Law1'.
 
 Section Law2.
-  Context `(D : @SpecializedCategory objD).
-  Context `(C1 : @SpecializedCategory objC1).
-  Context `(C2 : @SpecializedCategory objC2).
+  Variable D : Category.
+  Context `(C1 : @Category objC1).
+  Context `(C2 : @Category objC2).
 
   Definition ExponentialLaw2Functor
-  : SpecializedFunctor (D ^ (C1 + C2)) ((D ^ C1) * (D ^ C2))
+  : Functor (D ^ (C1 + C2)) ((D ^ C1) * (D ^ C2))
     := FunctorProduct (FunctorialComposition C1 (C1 + C2) D ⟨ - , inl_Functor _ _ ⟩)
                       (FunctorialComposition C2 (C1 + C2) D ⟨ - , inr_Functor _ _ ⟩).
 
   Definition ExponentialLaw2Functor_Inverse
-  : SpecializedFunctor ((D ^ C1) * (D ^ C2)) (D ^ (C1 + C2)).
+  : Functor ((D ^ C1) * (D ^ C2)) (D ^ (C1 + C2)).
   Proof.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor
                   C D
                   (fun FG => sum_Functor (fst FG) (snd FG))
                   (fun _ _ m => sum_Functor_Functorial (fst m) (snd m))
@@ -219,15 +219,15 @@
 End Law2.
 
 Section Law3.
-  Context `(C1 : @SpecializedCategory objC1).
-  Context `(C2 : @SpecializedCategory objC2).
-  Context `(D : @SpecializedCategory objD).
+  Context `(C1 : @Category objC1).
+  Context `(C2 : @Category objC2).
+  Variable D : Category.
 
   Definition ExponentialLaw3Functor
-  : SpecializedFunctor ((C1 * C2) ^ D) (C1 ^ D * C2 ^ D).
+  : Functor ((C1 * C2) ^ D) (C1 ^ D * C2 ^ D).
     match goal with
-      | [ |- SpecializedFunctor ?F ?G ] =>
-        refine (Build_SpecializedFunctor
+      | [ |- Functor ?F ?G ] =>
+        refine (Build_Functor
                   F G
                   (fun H => (ComposeFunctors fst_Functor H,
                              ComposeFunctors snd_Functor H))
@@ -256,11 +256,11 @@
   Defined.
 
   Definition ExponentialLaw3Functor_Inverse
-  : SpecializedFunctor (C1 ^ D * C2 ^ D) ((C1 * C2) ^ D).
+  : Functor (C1 ^ D * C2 ^ D) ((C1 * C2) ^ D).
     let FG := (match goal with
-                   [ |- SpecializedFunctor ?F ?G ] => constr:(F, G)
+                   [ |- Functor ?F ?G ] => constr:(F, G)
                end) in
-    refine (Build_SpecializedFunctor
+    refine (Build_Functor
               (fst FG) (snd FG)
               (fun H => FunctorProduct
                           (@fst_Functor _ (C1 ^ D) _ (C2 ^ D) H)
@@ -300,9 +300,9 @@
 End Law3.
 
 Section Law4.
-  Context `(C1 : @SpecializedCategory objC1).
-  Context `(C2 : @SpecializedCategory objC2).
-  Context `(D : @SpecializedCategory objD).
+  Context `(C1 : @Category objC1).
+  Context `(C2 : @Category objC2).
+  Variable D : Category.
 
   Section functor.
     Local Ltac do_exponential4 := intros; simpl;
@@ -322,8 +322,8 @@
     Proof.
       intro F; hnf in F |- *.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
             (fun c1c2 => F (snd c1c2) (fst c1c2))
             (fun s1s2 d1d2 m1m2 => Compose ((F (snd d1d2)).(MorphismOf) (fst m1m2)) ((F.(MorphismOf) (snd m1m2)) (fst s1s2)))
             _
@@ -350,11 +350,11 @@
         ).
     Defined.
 
-    Definition ExponentialLaw4Functor : SpecializedFunctor ((D ^ C1) ^ C2) (D ^ (C1 * C2)).
+    Definition ExponentialLaw4Functor : Functor ((D ^ C1) ^ C2) (D ^ (C1 * C2)).
     Proof.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
             ExponentialLaw4Functor_ObjectOf
             ExponentialLaw4Functor_MorphismOf
             _
@@ -381,15 +381,15 @@
 
 
     Section ObjectOf.
-      Variable F : SpecializedFunctor (C1 * C2) D.
+      Variable F : Functor (C1 * C2) D.
 
       Definition ExponentialLaw4Functor_Inverse_ObjectOf_ObjectOf : C2 -> (D ^ C1)%category.
       Proof.
         intro c2.
         hnf.
         match goal with
-          | [ |- SpecializedFunctor ?C ?D ] =>
-            refine (Build_SpecializedFunctor C D
+          | [ |- Functor ?C ?D ] =>
+            refine (Build_Functor C D
               (fun c1 => F (c1, c2))
               (fun s1 d1 m1 => F.(MorphismOf) (s := (s1, c2)) (d := (d1, c2)) (m1, Identity c2))
               _
@@ -410,8 +410,8 @@
       Proof.
         hnf.
         match goal with
-          | [ |- SpecializedFunctor ?C ?D ] =>
-            refine (Build_SpecializedFunctor C D
+          | [ |- Functor ?C ?D ] =>
+            refine (Build_Functor C D
               ExponentialLaw4Functor_Inverse_ObjectOf_ObjectOf
               ExponentialLaw4Functor_Inverse_ObjectOf_MorphismOf
               _
@@ -441,11 +441,11 @@
 
     Arguments ExponentialLaw4Functor_Inverse_MorphismOf_ComponentsOf / _ _ _ _.
 
-    Definition ExponentialLaw4Functor_Inverse : SpecializedFunctor (D ^ (C1 * C2)) ((D ^ C1) ^ C2).
+    Definition ExponentialLaw4Functor_Inverse : Functor (D ^ (C1 * C2)) ((D ^ C1) ^ C2).
     Proof.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
             ExponentialLaw4Functor_Inverse_ObjectOf
             ExponentialLaw4Functor_Inverse_MorphismOf
             _
diff -r 4eb6407c6ca3 Functor.v
--- a/Functor.v	Fri May 24 20:09:37 2013 -0400
+++ b/Functor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq ProofIrrelevance FunctionalExtensionality.
-Require Export Notations SpecializedCategory Category.
+Require Export Notations Category.
 Require Import Common StructureEquality FEqualDep.
 
 Set Implicit Arguments.
@@ -12,58 +12,47 @@
 
 Local Infix "==" := JMeq.
 
-Section SpecializedFunctor.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+Section Functor.
+  Local Open Scope morphism_scope.
+
+  Variables C D : Category.
 
   (**
      Quoting from the lecture notes for 18.705, Commutative Algebra:
 
      A map of categories is known as a functor. Namely, given
-     categories [C] and [C'], a (covariant) functor [F : C -> C'] is a rule that assigns to
-     each object [A] of [C] an object [F A] of [C'] and to each map [m : A -> B] of [C] a map
-     [F m : F A -> F B] of [C'] preserving composition and identity; that is,
-     (1) [F (m' ○ m) = (F m') ○ (F m)] for maps [m : A -> B] and [m' : B -> C] of [C], and
-     (2) [F (id A) = id (F A)] for any object [A] of [C], where [id A] is the identity morphism of [A].
-     **)
-  Record SpecializedFunctor :=
+     categories [C] and [C'], a (covariant) functor [F : C -> C'] is a
+     rule that assigns to each object [A] of [C] an object [F A] of
+     [C'] and to each map [m : A -> B] of [C] a map [F m : F A -> F B]
+     of [C'] preserving composition and identity; that is,
+
+     (1) [F (m' ∘ m) = (F m') ∘ (F m)] for maps [m : A -> B] and [m' :
+         B -> C] of [C], and
+
+     (2) [F (id A) = id (F A)] for any object [A] of [C], where [id A]
+         is the identity morphism of [A].
+   **)
+
+  Record Functor :=
     {
-      ObjectOf :> objC -> objD;
+      ObjectOf :> C -> D;
       MorphismOf : forall s d, C.(Morphism) s d -> D.(Morphism) (ObjectOf s) (ObjectOf d);
       FCompositionOf : forall s d d' (m1 : C.(Morphism) s d) (m2: C.(Morphism) d d'),
-                          MorphismOf _ _ (Compose m2 m1) = Compose (MorphismOf _ _ m2) (MorphismOf _ _ m1);
+                          MorphismOf _ _ (m2 ∘ m1) = (MorphismOf _ _ m2) ∘ (MorphismOf _ _ m1);
       FIdentityOf : forall x, MorphismOf _ _ (Identity x) = Identity (ObjectOf x)
     }.
-End SpecializedFunctor.
-
-Section Functor.
-  Variable C : Category.
-  Variable D : Category.
-
-  Definition Functor := SpecializedFunctor C D.
 End Functor.
 
-Bind Scope functor_scope with SpecializedFunctor.
 Bind Scope functor_scope with Functor.
 
 Create HintDb functor discriminated.
 
-Identity Coercion Functor_SpecializedFunctor_Id : Functor >-> SpecializedFunctor.
-Definition GeneralizeFunctor objC C objD D (F : @SpecializedFunctor objC C objD D) : Functor C D := F.
-Coercion GeneralizeFunctor : SpecializedFunctor >-> Functor.
+Arguments Functor C D.
+Arguments ObjectOf {C%category D%category} F%functor c%object : rename, simpl nomatch.
+Arguments MorphismOf [C%category D%category] F%functor [s%object d%object] m%morphism : rename, simpl nomatch.
 
-(* try to always unfold [GeneralizeFunctor]; it's in there
-   only for coercions *)
-Arguments GeneralizeFunctor [objC C objD D] F /.
-Hint Extern 0 => unfold GeneralizeFunctor : category 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.
-
-Arguments FCompositionOf [objC C objD D] F _ _ _ _ _ : rename.
-Arguments FIdentityOf [objC C objD D] F _ : rename.
+Arguments FCompositionOf [C D] F _ _ _ _ _ : rename.
+Arguments FIdentityOf [C D] F _ : rename.
 
 Hint Resolve @FCompositionOf @FIdentityOf : category functor.
 Hint Rewrite @FIdentityOf : category.
@@ -72,8 +61,6 @@
 Ltac functor_hideProofs :=
   repeat match goal with
              | [ |- context[{|
-                               ObjectOf := _;
-                               MorphismOf := _;
                                FCompositionOf := ?pf0;
                                FIdentityOf := ?pf1
                              |}] ] =>
@@ -88,24 +75,24 @@
     hnf in H;
     revert H; clear; intro H; clear H;
     match F'' with
-      | @Build_SpecializedFunctor ?objC ?C
-                                  ?objD ?D
-                                  ?OO
-                                  ?MO
-                                  ?FCO
-                                  ?FIO =>
-        refine (@Build_SpecializedFunctor objC C objD D
-                                          OO
-                                          MO
-                                          _
-                                          _);
+      | @Build_Functor ?C
+                       ?D
+                       ?OO
+                       ?MO
+                       ?FCO
+                       ?FIO =>
+        refine (@Build_Functor C D
+                               OO
+                               MO
+                               _
+                               _);
           [ abstract exact FCO | abstract exact FIO ]
     end.
 Ltac functor_abstract_trailing_props F := functor_tac_abstract_trailing_props F ltac:(fun F' => F').
 Ltac functor_simpl_abstract_trailing_props F := functor_tac_abstract_trailing_props F ltac:(fun F' => let F'' := eval simpl in F' in F'').
 
 Section Functors_Equal.
-  Lemma Functor_contr_eq' objC C objD D (F G : @SpecializedFunctor objC C objD D)
+  (*Lemma Functor_contr_eq' objC C objD D (F G : @Functor objC C objD D)
         (D_morphism_proof_irrelevance
          : forall s d (m1 m2 : Morphism D s d) (pf1 pf2 : m1 = m2),
              pf1 = pf2)
@@ -122,7 +109,7 @@
       trivial.
   Qed.
 
-  Lemma Functor_contr_eq objC C objD D (F G : @SpecializedFunctor objC C objD D)
+  Lemma Functor_contr_eq objC C objD D (F G : @Functor objC C objD D)
         (D_object_proof_irrelevance
          : forall (x : D) (pf : x = x),
              pf = eq_refl)
@@ -148,37 +135,49 @@
     subst_eq_refl_dec.
     reflexivity.
   Qed.
-
-  Lemma Functor_eq' objC C objD D : forall (F G : @SpecializedFunctor objC C objD D),
-    ObjectOf F = ObjectOf G
-    -> MorphismOf F == MorphismOf G
-    -> F = G.
-    destruct F, G; simpl; intros; specialize_all_ways; repeat subst;
-      f_equal; apply proof_irrelevance.
+*)
+  Lemma Functor_eq' C D
+  : forall (F G : Functor C D),
+      ObjectOf F = ObjectOf G
+      -> MorphismOf F == MorphismOf G
+      -> F = G.
+  Proof.
+    intros.
+    destruct_head Functor.
+    repeat subst.
+    f_equal; apply proof_irrelevance.
   Qed.
 
-  Lemma Functor_eq objC C objD D :
-    forall (F G : @SpecializedFunctor objC C objD D),
+  Lemma Functor_eq C D
+  : forall (F G : Functor C D),
       (forall x, ObjectOf F x = ObjectOf G x)
       -> (forall s d m, MorphismOf F (s := s) (d := d) m == MorphismOf G (s := s) (d := d) m)
       -> F = G.
-    intros; cut (ObjectOf F = ObjectOf G); intros; try apply Functor_eq'; destruct F, G; simpl in *; repeat subst;
-    try apply eq_JMeq;
-    repeat (apply functional_extensionality_dep; intro); trivial;
-    try apply JMeq_eq; trivial.
+  Proof.
+    intros;
+    destruct_head Functor.
+    repeat match goal with
+             | _ => progress subst
+             | [ H : _ |- _ ] => apply functional_extensionality_dep in H
+           end.
+    repeat match goal with
+             | _ => apply Functor_eq'; simpl; trivial
+             | _ => apply eq_JMeq
+             | _ => apply functional_extensionality_dep; intro
+             | _ => apply JMeq_eq; solve [ trivial ]
+           end.
   Qed.
 
-  Lemma Functor_JMeq objC C objD D objC' C' objD' D' :
-    forall (F : @SpecializedFunctor objC C objD D) (G : @SpecializedFunctor objC' C' objD' D'),
-      objC = objC'
-      -> objD = objD'
-      -> C == C'
-      -> D == D'
+  Lemma Functor_JMeq C D C' D'
+  : forall (F : Functor C D) (G : Functor C' D'),
+      C = C'
+      -> D = D'
       -> ObjectOf F == ObjectOf G
       -> MorphismOf F == MorphismOf G
       -> F == G.
+  Proof.
     simpl; intros; intuition; repeat subst; destruct F, G; simpl in *;
-      repeat subst; JMeq_eq.
+    repeat subst; JMeq_eq.
     f_equal; apply proof_irrelevance.
   Qed.
 End Functors_Equal.
@@ -199,12 +198,12 @@
     hnf in H;
     revert H; clear; intro H; clear H;
     match F'' with
-      | @Build_SpecializedFunctor ?objC ?C
-                                  ?objD ?D
-                                  ?OO
-                                  ?MO
-                                  ?FCO
-                                  ?FIO =>
+      | @Build_Functor ?C
+                       ?D
+                       ?OO
+                       ?MO
+                       ?FCO
+                       ?FIO =>
         let FCO' := fresh in
         let FIO' := fresh in
         let FCOT' := type of FCO in
@@ -213,97 +212,97 @@
         let FIOT := (eval simpl in FIOT') in
         assert (FCO' : FCOT) by abstract exact FCO;
           assert (FIO' : FIOT) by (clear FCO'; abstract exact FIO);
-          exists (@Build_SpecializedFunctor objC C objD D
-                                            OO
-                                            MO
-                                            FCO'
-                                            FIO');
+          exists (@Build_Functor C D
+                                 OO
+                                 MO
+                                 FCO'
+                                 FIO');
           expand; abstract (apply thm; reflexivity) || (apply thm; try reflexivity)
     end.
 
 Ltac functor_tac_abstract_trailing_props_with_equality tac :=
   pre_abstract_trailing_props;
   match goal with
-    | [ |- { F0 : SpecializedFunctor _ _ | F0 = ?F } ] =>
+    | [ |- { F0 : Functor _ _ | F0 = ?F } ] =>
       functor_tac_abstract_trailing_props_with_equality_do tac F @Functor_eq'
-    | [ |- { F0 : SpecializedFunctor _ _ | F0 == ?F } ] =>
+    | [ |- { F0 : Functor _ _ | F0 == ?F } ] =>
       functor_tac_abstract_trailing_props_with_equality_do tac F @Functor_JMeq
   end.
 Ltac functor_abstract_trailing_props_with_equality := functor_tac_abstract_trailing_props_with_equality ltac:(fun F' => F').
 Ltac functor_simpl_abstract_trailing_props_with_equality := functor_tac_abstract_trailing_props_with_equality ltac:(fun F' => let F'' := eval simpl in F' in F'').
 
 Section FunctorComposition.
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
-  Variable G : SpecializedFunctor D E.
-  Variable F : SpecializedFunctor C D.
+  Variables B C D E : Category.
+  Variable G : Functor D E.
+  Variable F : Functor C D.
 
-  Definition ComposeFunctors' : SpecializedFunctor C E
-    := Build_SpecializedFunctor C E
-                                (fun c => G (F c))
-                                (fun _ _ m => G.(MorphismOf) (F.(MorphismOf) m))
-                                (fun _ _ _ m1 m2 =>
-                                   match FCompositionOf G _ _ _ (MorphismOf F m1) (MorphismOf F m2) with
-                                     | eq_refl =>
-                                       match
-                                         FCompositionOf F _ _ _ m1 m2 in (_ = y)
-                                         return
-                                         (MorphismOf G (MorphismOf F (Compose m2 m1)) =
-                                          MorphismOf G y)
-                                       with
-                                         | eq_refl => eq_refl
-                                       end
-                                   end)
-                                (fun x =>
-                                   match FIdentityOf G (F x) with
-                                     | eq_refl =>
-                                       match
-                                         FIdentityOf F x in (_ = y)
-                                         return
-                                         (MorphismOf G (MorphismOf F (Identity x)) =
-                                          MorphismOf G y)
-                                       with
-                                         | eq_refl => eq_refl
-                                       end
-                                   end).
+  Definition ComposeFunctors' : Functor C E
+    := Build_Functor C E
+                     (fun c => G (F c))
+                     (fun _ _ m => G.(MorphismOf) (F.(MorphismOf) m))
+                     (fun _ _ _ m1 m2 =>
+                        match FCompositionOf G _ _ _ (MorphismOf F m1) (MorphismOf F m2) with
+                          | eq_refl =>
+                            match
+                              FCompositionOf F _ _ _ m1 m2 in (_ = y)
+                              return
+                              (MorphismOf G (MorphismOf F (Compose m2 m1)) =
+                               MorphismOf G y)
+                            with
+                              | eq_refl => eq_refl
+                            end
+                        end)
+                     (fun x =>
+                        match FIdentityOf G (F x) with
+                          | eq_refl =>
+                            match
+                              FIdentityOf F x in (_ = y)
+                              return
+                              (MorphismOf G (MorphismOf F (Identity x)) =
+                               MorphismOf G y)
+                            with
+                              | eq_refl => eq_refl
+                            end
+                        end).
 
   Hint Rewrite @FCompositionOf : functor.
 
-  Definition ComposeFunctors : SpecializedFunctor C E.
-    refine (Build_SpecializedFunctor C E
-                                     (fun c => G (F c))
-                                     (fun _ _ m => G.(MorphismOf) (F.(MorphismOf) m))
-                                     _
-                                     _);
+  Definition ComposeFunctors : Functor C E.
+    refine (Build_Functor C E
+                          (fun c => G (F c))
+                          (fun _ _ m => G.(MorphismOf) (F.(MorphismOf) m))
+                          _
+                          _);
     abstract (
         intros; autorewrite with functor; reflexivity
       ).
   Defined.
 End FunctorComposition.
 
+Infix "∘" := ComposeFunctors : functor_scope.
+
+Local Open Scope functor_scope.
+
 Section IdentityFunctor.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   (** There is an identity functor.  It does the obvious thing. *)
-  Definition IdentityFunctor : SpecializedFunctor C C
-    := Build_SpecializedFunctor C C
-                                (fun x => x)
-                                (fun _ _ x => x)
-                                (fun _ _ _ _ _ => eq_refl)
-                                (fun _ => eq_refl).
+  Definition IdentityFunctor : Functor C C
+    := Build_Functor C C
+                     (fun x => x)
+                     (fun _ _ x => x)
+                     (fun _ _ _ _ _ => eq_refl)
+                     (fun _ => eq_refl).
 End IdentityFunctor.
 
 Section IdentityFunctorLemmas.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
-  Lemma LeftIdentityFunctor (F : SpecializedFunctor D C) : ComposeFunctors (IdentityFunctor _) F = F.
+  Lemma LeftIdentityFunctor (F : Functor D C) : IdentityFunctor _ ∘ F = F.
     functor_eq.
   Qed.
 
-  Lemma RightIdentityFunctor (F : SpecializedFunctor C D) : ComposeFunctors F (IdentityFunctor _) = F.
+  Lemma RightIdentityFunctor (F : Functor C D) : F ∘ IdentityFunctor _ = F.
     functor_eq.
   Qed.
 End IdentityFunctorLemmas.
@@ -313,13 +312,10 @@
 Hint Immediate @LeftIdentityFunctor @RightIdentityFunctor : category functor.
 
 Section FunctorCompositionLemmas.
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
+  Variables B C D E : Category.
 
-  Lemma ComposeFunctorsAssociativity (F : SpecializedFunctor B C) (G : SpecializedFunctor C D) (H : SpecializedFunctor D E) :
-    ComposeFunctors (ComposeFunctors H G) F = ComposeFunctors H (ComposeFunctors G F).
+  Lemma ComposeFunctorsAssociativity (F : Functor B C) (G : Functor C D) (H : Functor D E)
+  : (H ∘ G) ∘ F = H ∘ (G ∘ F).
     functor_eq.
   Qed.
 End FunctorCompositionLemmas.
diff -r 4eb6407c6ca3 FunctorAttributes.v
--- a/FunctorAttributes.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorAttributes.v	Fri Jun 21 15:07:08 2013 -0400
@@ -13,12 +13,12 @@
 Local Open Scope category_scope.
 
 Section FullFaithful.
-  Context `(C : @SpecializedCategory objC).
-  Context `(C' : @LocallySmallSpecializedCategory objC').
-  Context `(D : @SpecializedCategory objD).
-  Context `(D' : @LocallySmallSpecializedCategory objD').
-  Variable F : SpecializedFunctor C D.
-  Variable F' : SpecializedFunctor C' D'.
+  Variable C : Category.
+  Context `(C' : @LocallyCategory objC').
+  Variable D : Category.
+  Context `(D' : @LocallyCategory objD').
+  Variable F : Functor C D.
+  Variable F' : Functor C' D'.
   Let COp := OppositeCategory C.
   Let DOp := OppositeCategory D.
   Let FOp := OppositeFunctor F.
@@ -27,8 +27,8 @@
   Let F'Op := OppositeFunctor F'.
 
   Definition InducedHomNaturalTransformation :
-    SpecializedNaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F)).
-    refine (Build_SpecializedNaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F))
+    NaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F)).
+    refine (Build_NaturalTransformation (HomFunctor C) (ComposeFunctors (HomFunctor D) (FOp * F))
       (fun sd : (COp * C) =>
         MorphismOf F (s := _) (d := _))
       _
diff -r 4eb6407c6ca3 FunctorCategory.v
--- a/FunctorCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor NaturalTransformation.
+Require Export Category Functor NaturalTransformation.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -10,20 +10,19 @@
 Set Universe Polymorphism.
 
 Section FunctorCategory.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
   (*
      There is a category Fun(C, D) of functors from [C] to [D].
    *)
-  Definition FunctorCategory : @SpecializedCategory (SpecializedFunctor C D).
-    refine (@Build_SpecializedCategory _
-                                       (SpecializedNaturalTransformation (C := C) (D := D))
-                                       (IdentityNaturalTransformation (C := C) (D := D))
-                                       (NTComposeT (C := C) (D := D))
-                                       _
-                                       _
-                                       _);
+  Definition FunctorCategory : Category.
+    refine (@Build_Category (Functor C D)
+                            (NaturalTransformation (C := C) (D := D))
+                            (IdentityNaturalTransformation (C := C) (D := D))
+                            (NTComposeT (C := C) (D := D))
+                            _
+                            _
+                            _);
     abstract (nt_eq; auto with morphism).
   Defined.
 End FunctorCategory.
diff -r 4eb6407c6ca3 FunctorCategoryFunctorial.v
--- a/FunctorCategoryFunctorial.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorCategoryFunctorial.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Export FunctorCategory NaturalTransformation.
-Require Import Common Notations SmallCat ProductCategory Duals ExponentialLaws CanonicalStructureSimplification FEqualDep.
+Require Import Common Notations Cat ProductCategory Duals ExponentialLaws CanonicalStructureSimplification FEqualDep.
 
 Set Implicit Arguments.
 
@@ -13,13 +13,13 @@
 
 Section FunctorCategoryParts.
   Section MorphismOf.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Context `(C' : @SpecializedCategory objC').
-    Context `(D' : @SpecializedCategory objD').
+    Variable C : Category.
+    Variable D : Category.
+    Context `(C' : @Category objC').
+    Context `(D' : @Category objD').
 
-    Variable F : SpecializedFunctor C C'.
-    Variable G : SpecializedFunctor D' D.
+    Variable F : Functor C C'.
+    Variable G : Functor D' D.
 
     Definition FunctorCategoryFunctor_MorphismOf_ObjectOf : (C ^ D)%functor -> (C' ^ D')%functor
       := fun H => ComposeFunctors F (ComposeFunctors H G).
@@ -32,8 +32,8 @@
 
     Global Arguments FunctorCategoryFunctor_MorphismOf_MorphismOf _ _ _ / .
 
-    Definition FunctorCategoryFunctor_MorphismOf : SpecializedFunctor (C ^ D) (C' ^ D').
-      refine (Build_SpecializedFunctor
+    Definition FunctorCategoryFunctor_MorphismOf : Functor (C ^ D) (C' ^ D').
+      refine (Build_Functor
                 (C ^ D) (C' ^ D')
                 FunctorCategoryFunctor_MorphismOf_ObjectOf
                 FunctorCategoryFunctor_MorphismOf_MorphismOf
@@ -49,8 +49,8 @@
   End MorphismOf.
 
   Section FIdentityOf.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
+    Variable C : Category.
+    Variable D : Category.
 
     Lemma FunctorCategoryFunctor_FIdentityOf : FunctorCategoryFunctor_MorphismOf (IdentityFunctor C) (IdentityFunctor D) = IdentityFunctor _.
       repeat (intro || apply Functor_eq || nt_eq); simpl; subst; JMeq_eq; rsimplify_morphisms; reflexivity.
@@ -58,17 +58,17 @@
   End FIdentityOf.
 
   Section FCompositionOf.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Context `(C' : @SpecializedCategory objC').
-    Context `(D' : @SpecializedCategory objD').
-    Context `(C'' : @SpecializedCategory objC'').
-    Context `(D'' : @SpecializedCategory objD'').
+    Variable C : Category.
+    Variable D : Category.
+    Context `(C' : @Category objC').
+    Context `(D' : @Category objD').
+    Context `(C'' : @Category objC'').
+    Context `(D'' : @Category objD'').
 
-    Variable F' : SpecializedFunctor C' C''.
-    Variable G : SpecializedFunctor D D'.
-    Variable F : SpecializedFunctor C C'.
-    Variable G' : SpecializedFunctor  D' D''.
+    Variable F' : Functor C' C''.
+    Variable G : Functor D D'.
+    Variable F : Functor C C'.
+    Variable G' : Functor  D' D''.
 
     Lemma FunctorCategoryFunctor_FCompositionOf : FunctorCategoryFunctor_MorphismOf (ComposeFunctors F' F) (ComposeFunctors G' G)
                                                   = ComposeFunctors (FunctorCategoryFunctor_MorphismOf F' G) (FunctorCategoryFunctor_MorphismOf F G').
@@ -78,9 +78,9 @@
 End FunctorCategoryParts.
 
 Section FunctorCategoryFunctor.
-  Definition FunctorCategoryFunctor : SpecializedFunctor (LocallySmallCat * (OppositeCategory LocallySmallCat)) LocallySmallCat.
-    refine (Build_SpecializedFunctor (LocallySmallCat * (OppositeCategory LocallySmallCat)) LocallySmallCat
-                                     (fun CD => (fst CD) ^ (snd CD) : LocallySmallSpecializedCategory _)
+  Definition FunctorCategoryFunctor : Functor (LocallyCat * (OppositeCategory LocallyCat)) LocallyCat.
+    refine (Build_Functor (LocallyCat * (OppositeCategory LocallyCat)) LocallyCat
+                                     (fun CD => (fst CD) ^ (snd CD) : LocallyCategory _)
                                      (fun s d m => FunctorCategoryFunctor_MorphismOf (fst m) (snd m))
                                      _
                                      _);
@@ -88,28 +88,28 @@
     abstract (intros; apply FunctorCategoryFunctor_FCompositionOf || apply FunctorCategoryFunctor_FIdentityOf).
   Defined.
 
-  (* Definition FunctorCategoryFunctor : ((LocallySmallCat ^ LocallySmallCat) ^ (OppositeCategory LocallySmallCat))%category
+  (* Definition FunctorCategoryFunctor : ((LocallyCat ^ LocallyCat) ^ (OppositeCategory LocallyCat))%category
     := ExponentialLaw4Functor_Inverse _ _ _ FunctorCategoryUncurriedFunctor. *)
 End FunctorCategoryFunctor.
 
 Notation "F ^ G" := (FunctorCategoryFunctor_MorphismOf F G) : functor_scope.
 
 Section NaturalTransformation.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(C' : @SpecializedCategory objC').
-  Context `(D' : @SpecializedCategory objD').
+  Variable C : Category.
+  Variable D : Category.
+  Context `(C' : @Category objC').
+  Context `(D' : @Category objD').
 
-  Variables F G : SpecializedFunctor C D.
-  Variables F' G' : SpecializedFunctor C' D'.
+  Variables F G : Functor C D.
+  Variables F' G' : Functor C' D'.
 
-  Variable T : SpecializedNaturalTransformation F G.
-  Variable T' : SpecializedNaturalTransformation F' G'.
+  Variable T : NaturalTransformation F G.
+  Variable T' : NaturalTransformation F' G'.
 
-  Definition LiftNaturalTransformationPointwise : SpecializedNaturalTransformation (F ^ F') (G ^ G').
+  Definition LiftNaturalTransformationPointwise : NaturalTransformation (F ^ F') (G ^ G').
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
                                                        (fun _ => NTComposeF T (NTComposeF (IdentityNaturalTransformation _) T'))
                                                        _)
     end.
@@ -134,18 +134,18 @@
 
 Section NaturalTransformation_Properties.
   Section identity.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
+    Variable C : Category.
+    Variable D : Category.
 
     Local Ltac t := intros; simpl; nt_eq; rsimplify_morphisms; try reflexivity.
 
     Section lift.
       Let LiftIdentityPointwise'
-      : SpecializedNaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D).
+      : NaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D).
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            refine (Build_SpecializedNaturalTransformation F G
-                                                           (fun x => (Build_SpecializedNaturalTransformation (F x) (G x)
+          | [ |- NaturalTransformation ?F ?G ] =>
+            refine (Build_NaturalTransformation F G
+                                                           (fun x => (Build_NaturalTransformation (F x) (G x)
                                                                                                              (fun y => Identity (x y))
                                                                                                              _))
                                                            _)
@@ -156,22 +156,22 @@
       Defined.
 
       Let LiftIdentityPointwise''
-      : SpecializedNaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D).
+      : NaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D).
         nt_simpl_abstract_trailing_props LiftIdentityPointwise'.
       Defined.
 
       Definition LiftIdentityPointwise
-      : SpecializedNaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D)
+      : NaturalTransformation (IdentityFunctor (C ^ D)) (IdentityFunctor C ^ IdentityFunctor D)
         := Eval hnf in LiftIdentityPointwise''.
     End lift.
 
     Section inverse.
       Let LiftIdentityPointwise'_Inverse
-      : SpecializedNaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D)).
+      : NaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D)).
         match goal with
-          | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-            refine (Build_SpecializedNaturalTransformation F G
-                                                           (fun x => (Build_SpecializedNaturalTransformation (F x) (G x)
+          | [ |- NaturalTransformation ?F ?G ] =>
+            refine (Build_NaturalTransformation F G
+                                                           (fun x => (Build_NaturalTransformation (F x) (G x)
                                                                                                              (fun y => Identity (x y))
                                                                                                              _))
                                                            _)
@@ -182,12 +182,12 @@
       Defined.
 
       Let LiftIdentityPointwise''_Inverse
-      : SpecializedNaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D)).
+      : NaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D)).
         nt_simpl_abstract_trailing_props LiftIdentityPointwise'_Inverse.
       Defined.
 
       Definition LiftIdentityPointwise_Inverse
-      : SpecializedNaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D))
+      : NaturalTransformation (IdentityFunctor C ^ IdentityFunctor D) (IdentityFunctor (C ^ D))
         := Eval hnf in LiftIdentityPointwise''_Inverse.
     End inverse.
 
@@ -201,21 +201,21 @@
   End identity.
 
   Section compose.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Context `(E : @SpecializedCategory objE).
-    Context `(C' : @SpecializedCategory objC').
-    Context `(D' : @SpecializedCategory objD').
-    Context `(E' : @SpecializedCategory objE').
+    Variable C : Category.
+    Variable D : Category.
+    Variable E : Category.
+    Context `(C' : @Category objC').
+    Context `(D' : @Category objD').
+    Context `(E' : @Category objE').
 
-    Variable G : SpecializedFunctor D E.
-    Variable F : SpecializedFunctor C D.
-    Variable F' : SpecializedFunctor D' E'.
-    Variable G' : SpecializedFunctor C' D'.
+    Variable G : Functor D E.
+    Variable F : Functor C D.
+    Variable F' : Functor D' E'.
+    Variable G' : Functor C' D'.
 
     Section lift.
       Let LiftComposeFunctorsPointwise_ComponentsOf x
-      : SpecializedNaturalTransformation
+      : NaturalTransformation
           (ComposeFunctors (ComposeFunctors G F)
                            (ComposeFunctors x (ComposeFunctors F' G')))
           (ComposeFunctors G
@@ -223,7 +223,7 @@
         nt_solve_associator.
       Defined.
 
-      Definition LiftComposeFunctorsPointwise : SpecializedNaturalTransformation (ComposeFunctors G F ^ ComposeFunctors F' G')
+      Definition LiftComposeFunctorsPointwise : NaturalTransformation (ComposeFunctors G F ^ ComposeFunctors F' G')
                                                                                  (ComposeFunctors (G ^ G') (F ^ F')).
         exists LiftComposeFunctorsPointwise_ComponentsOf;
         subst_body; simpl.
@@ -233,7 +233,7 @@
 
     Section inverse.
       Let LiftComposeFunctorsPointwise_Inverse_ComponentsOf x
-      : SpecializedNaturalTransformation
+      : NaturalTransformation
           (ComposeFunctors G
                            (ComposeFunctors (ComposeFunctors F (ComposeFunctors x F')) G'))
           (ComposeFunctors (ComposeFunctors G F)
@@ -241,7 +241,7 @@
         nt_solve_associator.
       Defined.
 
-      Definition LiftComposeFunctorsPointwise_Inverse : SpecializedNaturalTransformation (ComposeFunctors (G ^ G') (F ^ F'))
+      Definition LiftComposeFunctorsPointwise_Inverse : NaturalTransformation (ComposeFunctors (G ^ G') (F ^ F'))
                                                                                          (ComposeFunctors G F ^ ComposeFunctors F' G').
         exists LiftComposeFunctorsPointwise_Inverse_ComponentsOf;
         subst_body; simpl.
diff -r 4eb6407c6ca3 FunctorIsomorphisms.v
--- a/FunctorIsomorphisms.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorIsomorphisms.v	Fri Jun 21 15:07:08 2013 -0400
@@ -11,61 +11,62 @@
 Set Universe Polymorphism.
 
 Local Open Scope category_scope.
+Local Open Scope functor_scope.
 
 Section FunctorIsomorphism.
   (* Copy definitions from CategoryIsomorphisms.v *)
 
   Section FunctorIsInverseOf.
-    Context `{C : @SpecializedCategory objC}.
-    Context `{D : @SpecializedCategory objD}.
+    Context {C D : Category}.
 
-    Definition FunctorIsInverseOf1 (F : SpecializedFunctor C D) (G : SpecializedFunctor D C) : Prop :=
-      ComposeFunctors G F = IdentityFunctor C.
-    Definition FunctorIsInverseOf2 (F : SpecializedFunctor C D) (G : SpecializedFunctor D C) : Prop :=
-      ComposeFunctors F G = IdentityFunctor D.
+    Definition FunctorIsInverseOf1 (F : Functor C D) (G : Functor D C) : Prop :=
+      G ∘ F = IdentityFunctor C.
+    Definition FunctorIsInverseOf2 (F : Functor C D) (G : Functor D C) : Prop :=
+      F ∘ G = IdentityFunctor D.
 
     Global Arguments FunctorIsInverseOf1 / _ _.
     Global Arguments FunctorIsInverseOf2 / _ _.
 
-    Definition FunctorIsInverseOf (F : SpecializedFunctor C D) (G : SpecializedFunctor D C) : Prop := Eval simpl in
-      FunctorIsInverseOf1 F G /\ FunctorIsInverseOf2 F G.
+    Definition FunctorIsInverseOf (F : Functor C D) (G : Functor D C) : Prop := Eval simpl in
+                                                                                 FunctorIsInverseOf1 F G /\ FunctorIsInverseOf2 F G.
   End FunctorIsInverseOf.
 
-  Lemma FunctorIsInverseOf_sym `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD}
-    (F : SpecializedFunctor C D) (G : SpecializedFunctor D C) :
+  Lemma FunctorIsInverseOf_sym {C D}
+        (F : Functor C D) (G : Functor D C) :
     FunctorIsInverseOf F G -> FunctorIsInverseOf G F.
     intros; hnf in *; split_and; split; trivial.
   Qed.
 
   Section FunctorIsomorphismOf.
-    Record FunctorIsomorphismOf `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} (F : SpecializedFunctor C D) := {
-      FunctorIsomorphismOf_Functor :> _ := F;
-      InverseFunctor : SpecializedFunctor D C;
-      LeftInverseFunctor : ComposeFunctors InverseFunctor F = IdentityFunctor C;
-      RightInverseFunctor : ComposeFunctors F InverseFunctor = IdentityFunctor D
-    }.
+    Record FunctorIsomorphismOf {C D} (F : Functor C D) :=
+      {
+        FunctorIsomorphismOf_Functor :> _ := F;
+        InverseFunctor : Functor D C;
+        LeftInverseFunctor : InverseFunctor ∘ F = IdentityFunctor C;
+        RightInverseFunctor : F ∘ InverseFunctor = IdentityFunctor D
+      }.
 
     Hint Resolve RightInverseFunctor LeftInverseFunctor : category.
     Hint Resolve RightInverseFunctor LeftInverseFunctor : functor.
 
-    Definition FunctorIsomorphismOf_Identity `(C : @SpecializedCategory objC) : FunctorIsomorphismOf (IdentityFunctor C).
+    Definition FunctorIsomorphismOf_Identity (C : Category) : FunctorIsomorphismOf (IdentityFunctor C).
       exists (IdentityFunctor _); eauto with functor.
     Defined.
 
-    Definition InverseOfFunctor `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} (F : SpecializedFunctor C D)
-      (i : FunctorIsomorphismOf F) : FunctorIsomorphismOf (InverseFunctor i).
+    Definition InverseOfFunctor {C D} (F : Functor C D)
+               (i : FunctorIsomorphismOf F) : FunctorIsomorphismOf (InverseFunctor i).
       exists i; auto with functor.
     Defined.
 
-    Definition ComposeFunctorIsmorphismOf `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} `{E : @SpecializedCategory objE}
-      {F : SpecializedFunctor D E} {G : SpecializedFunctor C D} (i1 : FunctorIsomorphismOf F) (i2 : FunctorIsomorphismOf G) :
-      FunctorIsomorphismOf (ComposeFunctors F G).
-      exists (ComposeFunctors (InverseFunctor i2) (InverseFunctor i1));
+    Definition ComposeFunctorIsmorphismOf {C D E}
+               {F : Functor D E} {G : Functor C D} (i1 : FunctorIsomorphismOf F) (i2 : FunctorIsomorphismOf G)
+    : FunctorIsomorphismOf (F ∘ G).
+      exists (InverseFunctor i2 ∘ InverseFunctor i1);
       abstract (
           match goal with
             | [ |- context[ComposeFunctors (ComposeFunctors ?F ?G) (ComposeFunctors ?H ?I)] ] =>
-              transitivity (ComposeFunctors (ComposeFunctors F (ComposeFunctors G H)) I);
-                try solve [ repeat rewrite ComposeFunctorsAssociativity; reflexivity ]; []
+              transitivity ((F ∘ (G ∘ H)) ∘ I);
+            try solve [ repeat rewrite ComposeFunctorsAssociativity; reflexivity ]; []
           end;
           destruct i1, i2; simpl;
           repeat (rewrite_hyp; autorewrite with functor);
@@ -75,19 +76,20 @@
   End FunctorIsomorphismOf.
 
   Section IsomorphismOfCategories.
-    Record IsomorphismOfCategories `(C : @SpecializedCategory objC) `(D : @SpecializedCategory objD) := {
-      IsomorphismOfCategories_Functor : SpecializedFunctor C D;
-      IsomorphismOfCategories_Of :> FunctorIsomorphismOf IsomorphismOfCategories_Functor
-    }.
+    Record IsomorphismOfCategories C D :=
+      {
+        IsomorphismOfCategories_Functor : Functor C D;
+        IsomorphismOfCategories_Of :> FunctorIsomorphismOf IsomorphismOfCategories_Functor
+      }.
 
     Global Coercion Build_IsomorphismOfCategories : FunctorIsomorphismOf >-> IsomorphismOfCategories.
   End IsomorphismOfCategories.
 
   Section FunctorIsIsomorphism.
-    Definition FunctorIsIsomorphism `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} (F : SpecializedFunctor C D) : Prop :=
+    Definition FunctorIsIsomorphism {C D} (F : Functor C D) : Prop :=
       exists G, FunctorIsInverseOf F G.
 
-    Lemma FunctorIsmorphismOf_FunctorIsIsomorphism `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} (F : SpecializedFunctor C D) :
+    Lemma FunctorIsmorphismOf_FunctorIsIsomorphism {C D} (F : Functor C D) :
       FunctorIsomorphismOf F -> FunctorIsIsomorphism F.
       intro i; hnf.
       exists (InverseFunctor i);
@@ -96,7 +98,7 @@
         assumption.
     Qed.
 
-    Lemma FunctorIsIsomorphism_FunctorIsmorphismOf `{C : @SpecializedCategory objC} `{D : @SpecializedCategory objD} (F : SpecializedFunctor C D) :
+    Lemma FunctorIsIsomorphism_FunctorIsmorphismOf {C D} (F : Functor C D) :
       FunctorIsIsomorphism F -> exists _ : FunctorIsomorphismOf F, True.
       intro i; destruct_hypotheses.
       destruct_exists; trivial.
@@ -106,9 +108,9 @@
 
   Section CategoriesIsomorphic.
     Definition CategoriesIsomorphic (C D : Category) : Prop :=
-      exists (F : SpecializedFunctor C D) (G : SpecializedFunctor D C), FunctorIsInverseOf F G.
+      exists (F : Functor C D) (G : Functor D C), FunctorIsInverseOf F G.
 
-    Lemma IsmorphismOfCategories_CategoriesIsomorphic `(C : @SpecializedCategory objC) `(D : @SpecializedCategory objD) :
+    Lemma IsmorphismOfCategories_CategoriesIsomorphic C D :
       IsomorphismOfCategories C D -> CategoriesIsomorphic C D.
       intro i; destruct i as [ m i ].
       exists m.
@@ -123,16 +125,16 @@
     Qed.
 
     Local Ltac t_iso' := intros;
-      repeat match goal with
-               | [ i : CategoriesIsomorphic _ _ |- _ ] => destruct (CategoriesIsomorphic_IsomorphismOfCategories i) as [ ? [] ] ; clear i
-               | [ |- CategoriesIsomorphic _ _ ] => apply IsmorphismOfCategories_CategoriesIsomorphic
-             end.
+                        repeat match goal with
+                                 | [ i : CategoriesIsomorphic _ _ |- _ ] => destruct (CategoriesIsomorphic_IsomorphismOfCategories i) as [ ? [] ] ; clear i
+                                 | [ |- CategoriesIsomorphic _ _ ] => apply IsmorphismOfCategories_CategoriesIsomorphic
+                               end.
 
     Local Ltac t_iso:= t_iso';
-      repeat match goal with
-               | [ i : IsomorphismOfCategories _ _ |- _ ] => unique_pose (IsomorphismOfCategories_Of i); try clear i
-               | [ |- IsomorphismOfCategories _ _ ] => eapply Build_IsomorphismOfCategories
-             end.
+                      repeat match goal with
+                               | [ i : IsomorphismOfCategories _ _ |- _ ] => unique_pose (IsomorphismOfCategories_Of i); try clear i
+                               | [ |- IsomorphismOfCategories _ _ ] => eapply Build_IsomorphismOfCategories
+                             end.
 
     Lemma CategoriesIsomorphic_refl (C : Category) : CategoriesIsomorphic C C.
       t_iso.
@@ -155,17 +157,17 @@
     Qed.
 
     Global Add Parametric Relation : _ @CategoriesIsomorphic
-      reflexivity proved by @CategoriesIsomorphic_refl
-      symmetry proved by @CategoriesIsomorphic_sym
-      transitivity proved by @CategoriesIsomorphic_trans
-        as CategoriesIsomorphic_rel.
+        reflexivity proved by @CategoriesIsomorphic_refl
+        symmetry proved by @CategoriesIsomorphic_sym
+        transitivity proved by @CategoriesIsomorphic_trans
+          as CategoriesIsomorphic_rel.
   End CategoriesIsomorphic.
 End FunctorIsomorphism.
 
 Section Functor_preserves_isomorphism.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
+  Variable C : Category.
+  Variable D : Category.
+  Variable F : Functor C D.
 
   Hint Rewrite <- FCompositionOf : functor.
 
diff -r 4eb6407c6ca3 FunctorProduct.v
--- a/FunctorProduct.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorProduct.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,14 +10,14 @@
 Set Universe Polymorphism.
 
 Section FunctorProduct.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(D' : @SpecializedCategory objD').
+  Variable C : Category.
+  Variable D : Category.
+  Context `(D' : @Category objD').
 
-  Definition FunctorProduct (F : Functor C D) (F' : Functor C D') : SpecializedFunctor C (D * D').
+  Definition FunctorProduct (F : Functor C D) (F' : Functor C D') : Functor C (D * D').
     match goal with
-      | [ |- SpecializedFunctor ?C0 ?D0 ] =>
-        refine (Build_SpecializedFunctor
+      | [ |- Functor ?C0 ?D0 ] =>
+        refine (Build_Functor
                   C0 D0
                   (fun c => (F c, F' c))
                   (fun s d m => (MorphismOf F m, MorphismOf F' m))
@@ -30,12 +30,12 @@
   Definition FunctorProductFunctorial
              (F G : Functor C D)
              (F' G' : Functor C D')
-             (T : SpecializedNaturalTransformation F G)
-             (T' : SpecializedNaturalTransformation F' G')
-  : SpecializedNaturalTransformation (FunctorProduct F F') (FunctorProduct G G').
+             (T : NaturalTransformation F G)
+             (T' : NaturalTransformation F' G')
+  : NaturalTransformation (FunctorProduct F F') (FunctorProduct G G').
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
                                                        (fun x => (T x, T' x))
                                                        _)
     end.
@@ -44,14 +44,14 @@
 End FunctorProduct.
 
 Section FunctorProduct'.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(C' : @SpecializedCategory objC').
-  Context `(D' : @SpecializedCategory objD').
+  Variable C : Category.
+  Variable D : Category.
+  Context `(C' : @Category objC').
+  Context `(D' : @Category objD').
   Variable F : Functor C D.
   Variable F' : Functor C' D'.
 
-  Definition FunctorProduct' : SpecializedFunctor  (C * C') (D * D')
+  Definition FunctorProduct' : Functor  (C * C') (D * D')
     := FunctorProduct (ComposeFunctors F fst_Functor) (ComposeFunctors F' snd_Functor).
 End FunctorProduct'.
 
diff -r 4eb6407c6ca3 FunctorialComposition.v
--- a/FunctorialComposition.v	Fri May 24 20:09:37 2013 -0400
+++ b/FunctorialComposition.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,15 +10,15 @@
 Set Universe Polymorphism.
 
 Section FunctorialComposition.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Context `(E : SpecializedCategory objE).
+  Variable C : Category.
+  Variable D : Category.
+  Variable E : Category.
 
-  Definition FunctorialComposition : SpecializedFunctor ((E ^ D) * (D ^ C)) (E ^ C).
+  Definition FunctorialComposition : Functor ((E ^ D) * (D ^ C)) (E ^ C).
   Proof.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun fg => ComposeFunctors (fst fg) (snd fg))
           (fun _ _ tu => NTComposeF (fst tu) (snd tu))
           _
diff -r 4eb6407c6ca3 Graphs.v
--- a/Graphs.v	Fri May 24 20:09:37 2013 -0400
+++ b/Graphs.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export Functor SetCategory SmallCat FunctorCategory Paths.
+Require Export Functor SetCategory Cat FunctorCategory Paths.
 Require Import Common FEqualDep.
 
 Set Implicit Arguments.
@@ -11,7 +11,7 @@
 Set Universe Polymorphism.
 
 Section GraphObj.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Inductive GraphIndex := GraphIndexSource | GraphIndexTarget.
 
@@ -30,8 +30,8 @@
     destruct s, d, d'; simpl in *; trivial.
   Defined.
 
-  Definition GraphIndexingCategory : @SpecializedCategory GraphIndex.
-    refine (@Build_SpecializedCategory _
+  Definition GraphIndexingCategory : @Category GraphIndex.
+    refine (@Build_Category _
                                        GraphIndex_Morphism
                                        (fun x => match x with GraphIndexSource => tt | GraphIndexTarget => tt end)
                                        GraphIndex_Compose
@@ -67,11 +67,11 @@
       | (GraphIndexTarget, GraphIndexTarget) => fun _ => @id _
     end m.
 
-  Definition UnderlyingGraph : SpecializedFunctor GraphIndexingCategory TypeCat.
+  Definition UnderlyingGraph : Functor GraphIndexingCategory TypeCat.
   Proof.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           UnderlyingGraph_ObjectOf
           UnderlyingGraph_MorphismOf
           _
@@ -88,7 +88,7 @@
 End GraphObj.
 
 Section GraphFunctor.
-  Let UnderlyingGraphFunctor_ObjectOf (C : SmallCat) : SpecializedFunctor GraphIndexingCategory TypeCat :=
+  Let UnderlyingGraphFunctor_ObjectOf (C : Cat) : Functor GraphIndexingCategory TypeCat :=
     UnderlyingGraph C.
 
   Local Ltac t :=
@@ -102,7 +102,7 @@
                | _ => progress destruct_sig; simpl in *
              end.
 
-  Definition UnderlyingGraphFunctor_MorphismOf C D (F : Morphism SmallCat C D) :
+  Definition UnderlyingGraphFunctor_MorphismOf C D (F : Morphism Cat C D) :
     Morphism (TypeCat ^ GraphIndexingCategory) (UnderlyingGraph C) (UnderlyingGraph D).
   Proof.
     exists (fun c => match c as c return (UnderlyingGraph C) c -> (UnderlyingGraph D) c with
@@ -112,11 +112,11 @@
     abstract t.
   Defined.
 
-  Definition UnderlyingGraphFunctor : SpecializedFunctor SmallCat (TypeCat ^ GraphIndexingCategory).
+  Definition UnderlyingGraphFunctor : Functor Cat (TypeCat ^ GraphIndexingCategory).
   Proof.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           UnderlyingGraphFunctor_ObjectOf
           UnderlyingGraphFunctor_MorphismOf
           _
@@ -137,15 +137,15 @@
 End GraphFunctor.
 
 Section FreeCategory.
-  Variable F : SpecializedFunctor GraphIndexingCategory TypeCat.
+  Variable F : Functor GraphIndexingCategory TypeCat.
 
   Let vertices := F GraphIndexTarget.
 
   Hint Rewrite concatenate_p_noedges concatenate_noedges_p concatenate_associative.
 
-  Definition FreeCategory : SpecializedCategory vertices.
+  Definition FreeCategory : Category vertices.
   Proof.
-    refine (@Build_SpecializedCategory
+    refine (@Build_Category
               vertices
               _
               (@NoEdges _ _)
diff -r 4eb6407c6ca3 Grothendieck.v
--- a/Grothendieck.v	Fri May 24 20:09:37 2013 -0400
+++ b/Grothendieck.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common SetCategory.
 
 Set Implicit Arguments.
@@ -24,9 +24,9 @@
      [Hom (Γ F) (c1, x1) (c2, x2)] is the set of morphisms
      [f : c1 -> c2] in [C] such that [F.(MorphismOf) f x1 = x2].
      *)
-  Context `(C : @SpecializedCategory objC).
-  Variable F : SpecializedFunctor C TypeCat.
-  Variable F' : SpecializedFunctor C SetCat.
+  Variable C : Category.
+  Variable F : Functor C TypeCat.
+  Variable F' : Functor C SetCat.
 
   Record GrothendieckPair := {
     GrothendieckC' : objC;
@@ -86,8 +86,8 @@
 
   Hint Extern 1 (@eq (sig _) _ _) => simpl_eq : category.
 
-  Definition CategoryOfElements : @SpecializedCategory GrothendieckPair.
-    refine (@Build_SpecializedCategory _
+  Definition CategoryOfElements : @Category GrothendieckPair.
+    refine (@Build_Category _
                                        (fun s d =>
                                           { f : C.(Morphism) (GrothendieckC s) (GrothendieckC d) | F.(MorphismOf) f (GrothendieckX s) = (GrothendieckX d) })
                                        (fun o => GrothendieckIdentity (GrothendieckC o) (GrothendieckX o))
@@ -101,7 +101,7 @@
       ).
   Defined.
 
-  Definition GrothendieckFunctor : SpecializedFunctor CategoryOfElements C.
+  Definition GrothendieckFunctor : Functor CategoryOfElements C.
     refine {| ObjectOf := (fun o : CategoryOfElements => GrothendieckC o);
       MorphismOf := (fun s d (m : CategoryOfElements.(Morphism) s d) => proj1_sig m)
     |}; abstract (eauto with category; intros; destruct_type CategoryOfElements; simpl; reflexivity).
@@ -109,9 +109,9 @@
 End Grothendieck.
 
 Section SetGrothendieckCoercion.
-  Context `(C : @SpecializedCategory objC).
-  Variable F : SpecializedFunctor C SetCat.
-  Let F' := (F : SpecializedFunctorToSet _) : SpecializedFunctorToType _.
+  Variable C : Category.
+  Variable F : Functor C SetCat.
+  Let F' := (F : FunctorToSet _) : FunctorToType _.
 
   Definition SetGrothendieck2Grothendieck (G : SetGrothendieckPair F) : GrothendieckPair F'
     := {| GrothendieckC' := G.(SetGrothendieckC'); GrothendieckX' := G.(SetGrothendieckX') : F' _ |}.
diff -r 4eb6407c6ca3 GrothendieckCat.v
--- a/GrothendieckCat.v	Fri May 24 20:09:37 2013 -0400
+++ b/GrothendieckCat.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor ComputableCategory.
+Require Export Category Functor ComputableCategory.
 Require Import Common Notations NaturalTransformation.
 
 Set Implicit Arguments.
@@ -10,10 +10,10 @@
 Set Universe Polymorphism.
 
 Section Grothendieck.
-  Context `(Index2Cat : forall i : Index, SpecializedCategory (Index2Object i)).
+  Context `(Index2Cat : forall i : Index, Category (Index2Object i)).
   Let Cat := @ComputableCategory _ Index2Object Index2Cat.
 
-  Local Coercion Index2Cat : Index >-> SpecializedCategory.
+  Local Coercion Index2Cat : Index >-> Category.
 
   (**
      Quoting Wikipedia:
@@ -29,8 +29,8 @@
      [Hom (Γ F) (c1, x1) (c2, x2)] is the set of morphisms
      [f : c1 -> c2] in [C] such that [F.(MorphismOf) f x1 = x2].
      *)
-  Context `(C : @SpecializedCategory objC).
-  Variable F : SpecializedFunctor C Cat.
+  Variable C : Category.
+  Variable F : Functor C Cat.
 
   Record CatGrothendieckPair := {
     CatGrothendieckC' : objC;
@@ -70,7 +70,7 @@
   Local Hint Extern 1 (@eq (sig _) _ _) => simpl_eq : category.
   Local Hint Extern 1 (@eq (sigT _) _ _) => simpl_eq : category.
 
-  Definition CategoryOfCatElements : @SpecializedCategory CatGrothendieckPair.
+  Definition CategoryOfCatElements : @Category CatGrothendieckPair.
     refine {|
         Morphism := (fun s d => _);
         Compose' := (fun _ _ _ m1 m2 => CatGrothendieckCompose m1 m2);
@@ -112,7 +112,7 @@
       ).
   Defined.
 
-  Definition CatGrothendieckProjectionFunctor1 : SpecializedFunctor CategoryOfCatElements C.
+  Definition CatGrothendieckProjectionFunctor1 : Functor CategoryOfCatElements C.
     refine {|
         ObjectOf := (fun o : CategoryOfCatElements => CatGrothendieckC o);
         MorphismOf := (fun s d (m : CategoryOfCatElements.(Morphism) s d) => proj1_sig m)
diff -r 4eb6407c6ca3 GrothendieckFunctorial.v
--- a/GrothendieckFunctorial.v	Fri May 24 20:09:37 2013 -0400
+++ b/GrothendieckFunctorial.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,6 +1,6 @@
 Require Import JMeq.
 Require Export Grothendieck FunctorCategory SetCategory.
-Require Import Common Notations SmallCat SpecializedCommaCategory NatCategory FEqualDep.
+Require Import Common Notations Cat CommaCategory NatCategory FEqualDep.
 
 Set Implicit Arguments.
 
@@ -12,11 +12,11 @@
 
 Section GrothendieckNondependentFunctorial.
   Local Open Scope category_scope.
-  Local Notation "Cat / C" := (SliceSpecializedCategoryOver Cat C).
+  Local Notation "Cat / C" := (SliceCategoryOver Cat C).
 
-  Context `(C : @LocallySmallSpecializedCategory objC).
+  Context `(C : @LocallyCategory objC).
 
-  Let Cat := LocallySmallCat.
+  Let Cat := LocallyCat.
 
   Section functorial.
     Section object_of.
@@ -25,20 +25,20 @@
       : Cat / C.
         hnf.
         match goal with
-          | [ |- CommaSpecializedCategory_Object ?F ?G ] => refine (_ : CommaSpecializedCategory_ObjectT F G)
+          | [ |- CommaCategory_Object ?F ?G ] => refine (_ : CommaCategory_ObjectT F G)
         end.
         hnf; simpl.
         exists (((CategoryOfElements x
-                  : LocallySmallSpecializedCategory _)
-                 : LocallySmallCategory),
+                  : LocallyCategory _)
+                 : LocallyCategory),
                tt).
         exact (GrothendieckFunctor _).
       Defined.
     End object_of.
 
     Section morphism_of.
-      Variables F G : SpecializedFunctor C TypeCat.
-      Variable T : SpecializedNaturalTransformation F G.
+      Variables F G : Functor C TypeCat.
+      Variable T : NaturalTransformation F G.
 
       Definition CategoryOfElementsFunctorial'_MorphismOf_ObjectOf
       : CategoryOfElements F -> CategoryOfElements G
@@ -65,7 +65,7 @@
 
       Definition CategoryOfElementsFunctorial'_MorphismOf
       : Functor (CategoryOfElements F) (CategoryOfElements G).
-        refine (Build_SpecializedFunctor (CategoryOfElements F) (CategoryOfElements G)
+        refine (Build_Functor (CategoryOfElements F) (CategoryOfElements G)
                                          CategoryOfElementsFunctorial'_MorphismOf_ObjectOf
                                          CategoryOfElementsFunctorial'_MorphismOf_MorphismOf
                                          _
@@ -79,8 +79,8 @@
                  (CategoryOfElementsFunctorial_ObjectOf G).
         hnf.
         match goal with
-          | [ |- CommaSpecializedCategory_Morphism ?F ?G ]
-            => refine (_ : CommaSpecializedCategory_MorphismT F G)
+          | [ |- CommaCategory_Morphism ?F ?G ]
+            => refine (_ : CommaCategory_MorphismT F G)
         end.
         hnf; simpl.
         exists (CategoryOfElementsFunctorial'_MorphismOf, eq_refl).
@@ -101,9 +101,9 @@
              end.
 
     Definition CategoryOfElementsFunctorial'
-    : SpecializedFunctor (TypeCat ^ C) Cat.
-      refine (Build_SpecializedFunctor (TypeCat ^ C) Cat
-                                       (fun x => CategoryOfElements x : LocallySmallSpecializedCategory _)
+    : Functor (TypeCat ^ C) Cat.
+      refine (Build_Functor (TypeCat ^ C) Cat
+                                       (fun x => CategoryOfElements x : LocallyCategory _)
                                        CategoryOfElementsFunctorial'_MorphismOf
                                        _
                                        _);
@@ -111,8 +111,8 @@
     Defined.
 
     Definition CategoryOfElementsFunctorial
-    : SpecializedFunctor (TypeCat ^ C) (Cat / C).
-      refine (Build_SpecializedFunctor (TypeCat ^ C) (Cat / C)
+    : Functor (TypeCat ^ C) (Cat / C).
+      refine (Build_Functor (TypeCat ^ C) (Cat / C)
                                        CategoryOfElementsFunctorial_ObjectOf
                                        CategoryOfElementsFunctorial_MorphismOf
                                        _
diff -r 4eb6407c6ca3 GroupCategory.v
--- a/GroupCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/GroupCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Group.
+Require Export Category Group.
 Require Import Notations ComputableCategory SetCategory.
 
 Set Implicit Arguments.
@@ -24,8 +24,8 @@
 Ltac destruct_Trues := destruct_singleton_constructor I.
 
 Section as_category.
-  Definition CategoryOfGroup (G : Group) : SpecializedCategory unit.
-    refine (@Build_SpecializedCategory unit
+  Definition CategoryOfGroup (G : Group) : Category unit.
+    refine (@Build_Category unit
                                        (fun _ _ => G)
                                        (fun _ => @GroupIdentity G)
                                        (fun _ _ _ => @GroupOperation G)
@@ -36,16 +36,16 @@
   Defined.
 End as_category.
 
-Coercion CategoryOfGroup : Group >-> SpecializedCategory.
+Coercion CategoryOfGroup : Group >-> Category.
 
 Section category_of_groups.
-  Definition GroupCat : SpecializedCategory Group
+  Definition GroupCat : Category Group
     := Eval unfold ComputableCategory in ComputableCategory _ CategoryOfGroup.
 End category_of_groups.
 
 Section forgetful_functor.
-  Definition GroupForgetfulFunctor : SpecializedFunctor GroupCat TypeCat.
-    refine (Build_SpecializedFunctor GroupCat TypeCat
+  Definition GroupForgetfulFunctor : Functor GroupCat TypeCat.
+    refine (Build_Functor GroupCat TypeCat
                                      GroupObjects
                                      (fun s d m => MorphismOf m (s := tt) (d := tt))
                                      _
diff -r 4eb6407c6ca3 Groupoid.v
--- a/Groupoid.v	Fri May 24 20:09:37 2013 -0400
+++ b/Groupoid.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import ProofIrrelevance.
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common CategoryIsomorphisms.
 
 Set Implicit Arguments.
@@ -11,7 +11,7 @@
 Set Universe Polymorphism.
 
 Section GroupoidOf.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Inductive GroupoidOf_Morphism (s d : objC) :=
   | hasMorphism : C.(Morphism) s d -> GroupoidOf_Morphism s d
@@ -28,8 +28,8 @@
   (** Quoting Wikipedia:
      A groupoid is a small category in which every morphism is an isomorphism, and hence invertible.
      *)
-  Definition GroupoidOf : @SpecializedCategory objC.
-    refine (@Build_SpecializedCategory _
+  Definition GroupoidOf : @Category objC.
+    refine (@Build_Category _
                                        (fun s d => inhabited (GroupoidOf_Morphism s d))
                                        (fun o : C => inhabits (hasMorphism _ _ (Identity o)))
                                        (@GroupoidOf_Compose)
@@ -46,7 +46,7 @@
 Hint Constructors GroupoidOf_Morphism : category.
 
 Section Groupoid.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Lemma GroupoidOf_Groupoid : CategoryIsGroupoid (GroupoidOf C).
     hnf; intros s d m; hnf; destruct m as [ m ]; induction m;
@@ -64,8 +64,8 @@
         end.
   Qed.
 
-  Definition Groupoid_Functor : SpecializedFunctor C (GroupoidOf C).
-    refine (Build_SpecializedFunctor C (GroupoidOf C)
+  Definition Groupoid_Functor : Functor C (GroupoidOf C).
+    refine (Build_Functor C (GroupoidOf C)
       (fun c => c)
       (fun s d m => inhabits (hasMorphism C _ _ m))
       _
diff -r 4eb6407c6ca3 Hom.v
--- a/Hom.v	Fri May 24 20:09:37 2013 -0400
+++ b/Hom.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export SpecializedCategory Functor Duals SetCategory ProductCategory.
+Require Export Category Functor Duals SetCategory ProductCategory.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -17,12 +17,12 @@
 
 Section covariant_contravariant.
   Local Arguments InducedProductSndFunctor / _ _ _ _ _ _ _ _ _ _ _.
-  Definition CovariantHomFunctor `(C : @SpecializedCategory objC) (A : OppositeCategory C) :=
+  Definition CovariantHomFunctor `(C : @Category objC) (A : OppositeCategory C) :=
     Eval simpl in ((HomFunctor C) [ A, - ])%functor.
-  Definition ContravariantHomFunctor `(C : @SpecializedCategory objC) (A : C) := ((HomFunctor C) [ -, A ])%functor.
+  Definition ContravariantHomFunctor `(C : @Category objC) (A : C) := ((HomFunctor C) [ -, A ])%functor.
 
-  Definition CovariantHomSetFunctor `(C : @LocallySmallSpecializedCategory objC morC) (A : OppositeCategory C) := ((HomSetFunctor C) [ A, - ])%functor.
-  Definition ContravariantHomSetFunctor `(C : @LocallySmallSpecializedCategory objC morC) (A : C) := ((HomSetFunctor C) [ -, A ])%functor.
+  Definition CovariantHomSetFunctor `(C : @LocallyCategory objC morC) (A : OppositeCategory C) := ((HomSetFunctor C) [ A, - ])%functor.
+  Definition ContravariantHomSetFunctor `(C : @LocallyCategory objC morC) (A : C) := ((HomSetFunctor C) [ -, A ])%functor.
 End covariant_contravariant.
 
 but that would introduce an extra identity morphism which some tactics
@@ -30,14 +30,14 @@
 *)
 
 Section HomFunctor.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Let COp := OppositeCategory C.
 
   Section Covariant.
     Variable A : COp.
 
-    Definition CovariantHomFunctor : SpecializedFunctor C TypeCat.
-      refine (Build_SpecializedFunctor C TypeCat
+    Definition CovariantHomFunctor : Functor C TypeCat.
+      refine (Build_Functor C TypeCat
         (fun X : C => C.(Morphism) A X : TypeCat)
         (fun X Y f => (fun g : C.(Morphism) A X => Compose f g))
         _
@@ -50,8 +50,8 @@
   Section Contravariant.
     Variable B : C.
 
-    Definition ContravariantHomFunctor : SpecializedFunctor COp TypeCat.
-      refine (Build_SpecializedFunctor COp TypeCat
+    Definition ContravariantHomFunctor : Functor COp TypeCat.
+      refine (Build_Functor COp TypeCat
         (fun X : COp => COp.(Morphism) B X : TypeCat)
         (fun X Y (h : COp.(Morphism) X Y) => (fun g : COp.(Morphism) B X => Compose h g))
         _
@@ -72,8 +72,8 @@
     exact (Compose f (Compose g h)).
   Defined.
 
-  Definition HomFunctor : SpecializedFunctor (COp * C) TypeCat.
-    refine (Build_SpecializedFunctor (COp * C) TypeCat
+  Definition HomFunctor : Functor (COp * C) TypeCat.
+    refine (Build_Functor (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)
       _
@@ -89,7 +89,7 @@
 End HomFunctor.
 
 Section SplitHomFunctor.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Let COp := OppositeCategory C.
 
   Lemma SplitHom (X Y : COp * C) : forall gh,
diff -r 4eb6407c6ca3 IndiscreteCategory.v
--- a/IndiscreteCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/IndiscreteCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -13,8 +13,8 @@
   (** The indiscrete category has exactly one morphism between any two objects. *)
   Variable O : Type.
 
-  Definition IndiscreteCategory : @SpecializedCategory O
-    := @Build_SpecializedCategory O
+  Definition IndiscreteCategory : @Category O
+    := @Build_Category O
                                   (fun _ _ => unit)
                                   (fun _ => tt)
                                   (fun _ _ _ _ _ => tt)
@@ -25,11 +25,11 @@
 
 Section FunctorToIndiscrete.
   Variable O : Type.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Variable objOf : objC -> O.
 
-  Definition FunctorToIndiscrete : SpecializedFunctor C (IndiscreteCategory O)
-    := Build_SpecializedFunctor C (IndiscreteCategory O)
+  Definition FunctorToIndiscrete : Functor C (IndiscreteCategory O)
+    := Build_Functor C (IndiscreteCategory O)
                                 objOf
                                 (fun _ _ _ => tt)
                                 (fun _ _ _ _ _ => eq_refl)
diff -r 4eb6407c6ca3 InducedLimitFunctors.v
--- a/InducedLimitFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/InducedLimitFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
-Require Export LimitFunctorTheorems SpecializedLaxCommaCategory.
-Require Import Common DefinitionSimplification SpecializedCategory Functor NaturalTransformation Duals CanonicalStructureSimplification.
+Require Export LimitFunctorTheorems LaxCommaCategory.
+Require Import Common DefinitionSimplification Category Functor NaturalTransformation Duals CanonicalStructureSimplification.
 
 Set Implicit Arguments.
 
@@ -14,14 +14,14 @@
      a category that we're coming from.  So prove them separately, so
      we can use them elsewhere, without assuming a full [HasLimits]. *)
   Variable I : Type.
-  Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+  Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
-  Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
-  Local Notation "'CAT' ⇓ D" := (@LaxSliceSpecializedCategory _ _ Index2Cat _ D).
+  Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
+  Local Notation "'CAT' ⇓ D" := (@LaxSliceCategory _ _ Index2Cat _ D).
 
-  Context `(D : @SpecializedCategory objD).
+  Variable D : Category.
   Let DOp := OppositeCategory D.
 
   Section Limit.
@@ -83,10 +83,10 @@
 
     Hint Resolve InducedLimitFunctor_FCompositionOf InducedLimitFunctor_FIdentityOf.
 
-    Definition InducedLimitFunctor : SpecializedFunctor (CAT ⇑ D) D.
+    Definition InducedLimitFunctor : Functor (CAT ⇑ D) D.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
             (fun x => LimitObject (HasLimits x))
             (fun s d => @InducedLimitFunctor_MorphismOf s d (HasLimits s) (HasLimits d))
             _
@@ -155,10 +155,10 @@
 
     Hint Resolve InducedColimitFunctor_FCompositionOf InducedColimitFunctor_FIdentityOf.
 
-    Definition InducedColimitFunctor : SpecializedFunctor (CAT ⇓ D) D.
+    Definition InducedColimitFunctor : Functor (CAT ⇓ D) D.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
             (fun x => ColimitObject (HasColimits x))
             (fun s d => @InducedColimitFunctor_MorphismOf s d (HasColimits s) (HasColimits d))
             _
diff -r 4eb6407c6ca3 InitialTerminalCategory.v
--- a/InitialTerminalCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/InitialTerminalCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common Notations NatCategory.
 
 Set Implicit Arguments.
@@ -10,38 +10,38 @@
 Set Universe Polymorphism.
 
 Section InitialTerminal.
-   Definition InitialCategory : SmallSpecializedCategory _ := 0.
-   Definition TerminalCategory : SmallSpecializedCategory _ := 1.
+   Definition InitialCategory : Category := 0.
+   Definition TerminalCategory : Category := 1.
 End InitialTerminal.
 
 Section Functors.
-  Context `(C : SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition FunctorTo1 : SpecializedFunctor C 1
-    := Build_SpecializedFunctor C 1 (fun _ => tt) (fun _ _ _ => eq_refl) (fun _ _ _ _ _ => eq_refl) (fun _ => eq_refl).
-  Definition FunctorToTerminal : SpecializedFunctor C TerminalCategory := FunctorTo1.
+  Definition FunctorTo1 : Functor C 1
+    := Build_Functor C 1 (fun _ => tt) (fun _ _ _ => eq_refl) (fun _ _ _ _ _ => eq_refl) (fun _ => eq_refl).
+  Definition FunctorToTerminal : Functor C TerminalCategory := FunctorTo1.
 
-  Definition FunctorFrom1 (c : C) : SpecializedFunctor 1 C
-    := Build_SpecializedFunctor 1 C (fun _ => c) (fun _ _ _ => Identity c) (fun _ _ _ _ _ => eq_sym (@RightIdentity _ _ _ _ _)) (fun _ => eq_refl).
-  Definition FunctorFromTerminal (c : C) : SpecializedFunctor TerminalCategory C := FunctorFrom1 c.
+  Definition FunctorFrom1 (c : C) : Functor 1 C
+    := Build_Functor 1 C (fun _ => c) (fun _ _ _ => Identity c) (fun _ _ _ _ _ => eq_sym (@RightIdentity _ _ _ _)) (fun _ => eq_refl).
+  Definition FunctorFromTerminal (c : C) : Functor TerminalCategory C := FunctorFrom1 c.
 
-  Definition FunctorFrom0 : SpecializedFunctor 0 C
-    := Build_SpecializedFunctor 0 C (fun x => match x with end) (fun x _ _ => match x with end) (fun x _ _ _ _ => match x with end) (fun x => match x with end).
-  Definition FunctorFromInitial : SpecializedFunctor InitialCategory C := FunctorFrom0.
+  Definition FunctorFrom0 : Functor 0 C
+    := Build_Functor 0 C (fun x => match x with end) (fun x _ _ => match x with end) (fun x _ _ _ _ => match x with end) (fun x => match x with end).
+  Definition FunctorFromInitial : Functor InitialCategory C := FunctorFrom0.
 End Functors.
 
 Section FunctorsUnique.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Lemma InitialCategoryFunctorUnique
-  : forall F F' : SpecializedFunctor InitialCategory C,
+  : forall F F' : Functor InitialCategory C,
       F = F'.
   Proof.
     functor_eq; destruct_head_hnf @Empty_set.
   Qed.
 
   Lemma InitialCategoryFunctor'Unique
-  : forall F F' : SpecializedFunctor C InitialCategory,
+  : forall F F' : Functor C InitialCategory,
       F = F'.
   Proof.
     intros F F'.
@@ -58,7 +58,7 @@
   Qed.
 
   Lemma TerminalCategoryFunctorUnique
-  : forall F F' : SpecializedFunctor C TerminalCategory,
+  : forall F F' : Functor C TerminalCategory,
       F = F'.
   Proof.
     functor_eq; auto.
diff -r 4eb6407c6ca3 LaxCommaCategory.v
--- a/LaxCommaCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/LaxCommaCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,6 +1,6 @@
 Require Import ProofIrrelevance.
 Require Export Category Functor ProductCategory NaturalTransformation.
-Require Import Common DiscreteCategory ComputableCategory DefinitionSimplification FEqualDep SpecializedLaxCommaCategory.
+Require Import Common DiscreteCategory ComputableCategory DefinitionSimplification FEqualDep LaxCommaCategory.
 
 Set Implicit Arguments.
 
@@ -13,9 +13,9 @@
 Local Open Scope category_scope.
 
 Local Ltac fold_functor := idtac. (*
-  change (@SpecializedFunctor) with (fun objC (C : @SpecializedCategory objC) objD (D : @SpecializedCategory objD) => @Functor C D) in *;
-    change (@SpecializedNaturalTransformation) with (fun objC (C : @SpecializedCategory objC) objD (D : @SpecializedCategory objD)
-      (F G : SpecializedFunctor C D)
+  change (@Functor) with (fun objC (C : @Category objC) objD (D : @Category objD) => @Functor C D) in *;
+    change (@NaturalTransformation) with (fun objC (C : @Category objC) objD (D : @Category objD)
+      (F G : Functor C D)
       => @NaturalTransformation C D F G) in *. *)
 
 Section LaxSliceCategory.
@@ -30,24 +30,24 @@
 
   Variable C : Category.
 
-  (* Pull the definitions from SpecializedLaxCommaCategory.v,
-     removing [Specialized], so that we have smaller definitions.
+  (* Pull the definitions from LaxCommaCategory.v,
+     removing [], so that we have smaller definitions.
      *)
 
-  Let LaxSliceCategory_Object' := Eval hnf in LaxSliceSpecializedCategory_ObjectT _ Index2Cat C.
+  Let LaxSliceCategory_Object' := Eval hnf in LaxSliceCategory_ObjectT _ Index2Cat C.
   Let LaxSliceCategory_Object'' : Type.
     simpl_definition_by_tac_and_exact LaxSliceCategory_Object' ltac:(simpl in *; fold_functor; simpl in *).
   Defined.
   Definition LaxSliceCategory_Object := Eval hnf in LaxSliceCategory_Object''.
 
-  Let LaxSliceCategory_Morphism' (XG X'G' : LaxSliceCategory_Object) := Eval hnf in @LaxSliceSpecializedCategory_MorphismT _ _ Index2Cat _ C XG X'G'.
+  Let LaxSliceCategory_Morphism' (XG X'G' : LaxSliceCategory_Object) := Eval hnf in @LaxSliceCategory_MorphismT _ _ Index2Cat _ C XG X'G'.
   Let LaxSliceCategory_Morphism'' (XG X'G' : LaxSliceCategory_Object) : Type.
     simpl_definition_by_tac_and_exact (LaxSliceCategory_Morphism' XG X'G') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
   Defined.
   Definition LaxSliceCategory_Morphism (XG X'G' : LaxSliceCategory_Object) := Eval hnf in LaxSliceCategory_Morphism'' XG X'G'.
 
   Let LaxSliceCategory_Compose' s d d' Fα F'α'
-    := Eval hnf in @LaxSliceSpecializedCategory_Compose _ _ Index2Cat _ C s d d' Fα F'α'.
+    := Eval hnf in @LaxSliceCategory_Compose _ _ Index2Cat _ C s d d' Fα F'α'.
   Let LaxSliceCategory_Compose'' s d d' (Fα : LaxSliceCategory_Morphism d d') (F'α' : LaxSliceCategory_Morphism s d) :
     LaxSliceCategory_Morphism s d'.
     simpl_definition_by_tac_and_exact (@LaxSliceCategory_Compose' s d d' Fα F'α') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
@@ -56,7 +56,7 @@
     LaxSliceCategory_Morphism s d'
     := Eval hnf in @LaxSliceCategory_Compose'' s d d' Fα F'α'.
 
-  Let LaxSliceCategory_Identity' o := Eval hnf in @LaxSliceSpecializedCategory_Identity _ _ Index2Cat _ C o.
+  Let LaxSliceCategory_Identity' o := Eval hnf in @LaxSliceCategory_Identity _ _ Index2Cat _ C o.
   Let LaxSliceCategory_Identity'' (o : LaxSliceCategory_Object) : LaxSliceCategory_Morphism o o.
     simpl_definition_by_tac_and_exact (@LaxSliceCategory_Identity' o) ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
   Defined.
@@ -68,23 +68,23 @@
 
   Lemma LaxSliceCategory_Associativity o1 o2 o3 o4 (m1 : LaxSliceCategory_Morphism o1 o2) (m2 : LaxSliceCategory_Morphism o2 o3) (m3 : LaxSliceCategory_Morphism o3 o4) :
     LaxSliceCategory_Compose (LaxSliceCategory_Compose m3 m2) m1 = LaxSliceCategory_Compose m3 (LaxSliceCategory_Compose m2 m1).
-    abstract apply (@LaxSliceSpecializedCategory_Associativity _ _ Index2Cat _ C o1 o2 o3 o4 m1 m2 m3).
+    abstract apply (@LaxSliceCategory_Associativity _ _ Index2Cat _ C o1 o2 o3 o4 m1 m2 m3).
   Qed.
 
   Lemma LaxSliceCategory_LeftIdentity (a b : LaxSliceCategory_Object) (f : LaxSliceCategory_Morphism a b) :
     LaxSliceCategory_Compose (LaxSliceCategory_Identity b) f = f.
   Proof.
-    abstract apply (@LaxSliceSpecializedCategory_LeftIdentity _ _ Index2Cat _ C a b f).
+    abstract apply (@LaxSliceCategory_LeftIdentity _ _ Index2Cat _ C a b f).
   Qed.
 
   Lemma LaxSliceCategory_RightIdentity (a b : LaxSliceCategory_Object) (f : LaxSliceCategory_Morphism a b) :
     LaxSliceCategory_Compose f (LaxSliceCategory_Identity a) = f.
   Proof.
-    abstract apply (@LaxSliceSpecializedCategory_RightIdentity _ _ Index2Cat _ C a b f).
+    abstract apply (@LaxSliceCategory_RightIdentity _ _ Index2Cat _ C a b f).
   Qed.
 
   Definition LaxSliceCategory : Category.
-    refine (@Build_SpecializedCategory LaxSliceCategory_Object LaxSliceCategory_Morphism
+    refine (@Build_Category LaxSliceCategory_Object LaxSliceCategory_Morphism
       LaxSliceCategory_Identity
       LaxSliceCategory_Compose
       _ _ _
@@ -108,20 +108,20 @@
 
   Variable C : Category.
 
-  Let LaxCosliceCategory_Object' := Eval hnf in LaxCosliceSpecializedCategory_ObjectT Index2Cat C.
+  Let LaxCosliceCategory_Object' := Eval hnf in LaxCosliceCategory_ObjectT Index2Cat C.
   Let LaxCosliceCategory_Object'' : Type.
     simpl_definition_by_tac_and_exact LaxCosliceCategory_Object' ltac:(simpl in *; fold_functor; simpl in *).
   Defined.
   Definition LaxCosliceCategory_Object := Eval hnf in LaxCosliceCategory_Object''.
 
-  Let LaxCosliceCategory_Morphism' (XG X'G' : LaxCosliceCategory_Object) := Eval hnf in @LaxCosliceSpecializedCategory_MorphismT _ _ Index2Cat _ C XG X'G'.
+  Let LaxCosliceCategory_Morphism' (XG X'G' : LaxCosliceCategory_Object) := Eval hnf in @LaxCosliceCategory_MorphismT _ _ Index2Cat _ C XG X'G'.
   Let LaxCosliceCategory_Morphism'' (XG X'G' : LaxCosliceCategory_Object) : Type.
     simpl_definition_by_tac_and_exact (LaxCosliceCategory_Morphism' XG X'G') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
   Defined.
   Definition LaxCosliceCategory_Morphism (XG X'G' : LaxCosliceCategory_Object) := Eval hnf in LaxCosliceCategory_Morphism'' XG X'G'.
 
   Let LaxCosliceCategory_Compose' s d d' Fα F'α'
-    := Eval hnf in @LaxCosliceSpecializedCategory_Compose _ _ Index2Cat _ C s d d' Fα F'α'.
+    := Eval hnf in @LaxCosliceCategory_Compose _ _ Index2Cat _ C s d d' Fα F'α'.
   Let LaxCosliceCategory_Compose'' s d d' (Fα : LaxCosliceCategory_Morphism d d') (F'α' : LaxCosliceCategory_Morphism s d) :
     LaxCosliceCategory_Morphism s d'.
     simpl_definition_by_tac_and_exact (@LaxCosliceCategory_Compose' s d d' Fα F'α') ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
@@ -130,7 +130,7 @@
     LaxCosliceCategory_Morphism s d'
     := Eval hnf in @LaxCosliceCategory_Compose'' s d d' Fα F'α'.
 
-  Let LaxCosliceCategory_Identity' o := Eval hnf in @LaxCosliceSpecializedCategory_Identity _ _ Index2Cat _ C o.
+  Let LaxCosliceCategory_Identity' o := Eval hnf in @LaxCosliceCategory_Identity _ _ Index2Cat _ C o.
   Let LaxCosliceCategory_Identity'' (o : LaxCosliceCategory_Object) : LaxCosliceCategory_Morphism o o.
     simpl_definition_by_tac_and_exact (@LaxCosliceCategory_Identity' o) ltac:(subst_body; cbv beta in *; fold_functor; cbv beta in *).
   Defined.
@@ -142,23 +142,23 @@
 
   Lemma LaxCosliceCategory_Associativity o1 o2 o3 o4 (m1 : LaxCosliceCategory_Morphism o1 o2) (m2 : LaxCosliceCategory_Morphism o2 o3) (m3 : LaxCosliceCategory_Morphism o3 o4) :
     LaxCosliceCategory_Compose (LaxCosliceCategory_Compose m3 m2) m1 = LaxCosliceCategory_Compose m3 (LaxCosliceCategory_Compose m2 m1).
-    abstract apply (@LaxCosliceSpecializedCategory_Associativity _ _ Index2Cat _ C o1 o2 o3 o4 m1 m2 m3).
+    abstract apply (@LaxCosliceCategory_Associativity _ _ Index2Cat _ C o1 o2 o3 o4 m1 m2 m3).
   Qed.
 
   Lemma LaxCosliceCategory_LeftIdentity (a b : LaxCosliceCategory_Object) (f : LaxCosliceCategory_Morphism a b) :
     LaxCosliceCategory_Compose (LaxCosliceCategory_Identity b) f = f.
   Proof.
-    abstract apply (@LaxCosliceSpecializedCategory_LeftIdentity _ _ Index2Cat _ C a b f).
+    abstract apply (@LaxCosliceCategory_LeftIdentity _ _ Index2Cat _ C a b f).
   Qed.
 
   Lemma LaxCosliceCategory_RightIdentity (a b : LaxCosliceCategory_Object) (f : LaxCosliceCategory_Morphism a b) :
     LaxCosliceCategory_Compose f (LaxCosliceCategory_Identity a) = f.
   Proof.
-    abstract apply (@LaxCosliceSpecializedCategory_RightIdentity _ _ Index2Cat _ C a b f).
+    abstract apply (@LaxCosliceCategory_RightIdentity _ _ Index2Cat _ C a b f).
   Qed.
 
   Definition LaxCosliceCategory : Category.
-    refine (@Build_SpecializedCategory LaxCosliceCategory_Object LaxCosliceCategory_Morphism
+    refine (@Build_Category LaxCosliceCategory_Object LaxCosliceCategory_Morphism
       LaxCosliceCategory_Identity
       LaxCosliceCategory_Compose
       _ _ _
diff -r 4eb6407c6ca3 LimitFunctor2CategoryTheorems.v
--- a/LimitFunctor2CategoryTheorems.v	Fri May 24 20:09:37 2013 -0400
+++ b/LimitFunctor2CategoryTheorems.v	Fri Jun 21 15:07:08 2013 -0400
@@ -15,24 +15,24 @@
 DataMigrationFunctorsAdjoint.v *)
 
 Section LimReady.
-  (* Variables C D : LocallySmallCategory. *)
+  (* Variables C D : LocallyCategory. *)
   Variable S : Category.
 
-  Local Notation "A ↓ F" := (CosliceSpecializedCategory A F).
-  (*Local Notation "C / c" := (@SliceSpecializedCategoryOver _ _ C c).*)
+  Local Notation "A ↓ F" := (CosliceCategory A F).
+  (*Local Notation "C / c" := (@SliceCategoryOver _ _ C c).*)
 
   Variable I : Type.
   Variable Index2Object : I -> Type.
-  Variable Index2Cat : forall i : I, SpecializedCategory (Index2Object i).
+  Variable Index2Cat : forall i : I, Category (Index2Object i).
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
-  Local Notation "'CAT' ⇑ D" := (@LaxCosliceSpecializedCategory _ _ Index2Cat _ D).
+  Local Notation "'CAT' ⇑ D" := (@LaxCosliceCategory _ _ Index2Cat _ D).
 
   Variable HasLimits : forall C0 : CAT ⇑ S, Limit (projT2 C0).
 
   Definition CatUpSet2Morphism A B (m1 m2 : Morphism (CAT ⇑ S) A B) : Type
-    := { T : SpecializedNaturalTransformation (snd (projT1 m1)) (snd (projT1 m2)) |
+    := { T : NaturalTransformation (snd (projT1 m1)) (snd (projT1 m2)) |
          NTComposeT (projT2 m2) (NTComposeF (IdentityNaturalTransformation (projT2 A)) T) = projT2 m1 }.
 
   Lemma LimReady A B m1 m2 (LimitF := @InducedLimitFunctor _ _ Index2Cat _ S HasLimits) :
@@ -87,7 +87,7 @@
           hnf in T
     end.
     match goal with
-      | [ T : SpecializedNaturalTransformation _ _ |- _ ] =>
+      | [ T : NaturalTransformation _ _ |- _ ] =>
         let H := fresh in
         pose proof (Commutes T) as H;
           simpl in H;
diff -r 4eb6407c6ca3 LimitFunctorTheorems.v
--- a/LimitFunctorTheorems.v	Fri May 24 20:09:37 2013 -0400
+++ b/LimitFunctorTheorems.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Export LimitFunctors.
-Require Import Common DefinitionSimplification SpecializedCategory Functor NaturalTransformation.
+Require Import Common DefinitionSimplification Category Functor NaturalTransformation.
 
 Set Implicit Arguments.
 
@@ -34,12 +34,12 @@
      injects one set into its union with another and [lim G] projects a
      product of two sets onto one factor.
      *)
-  Context `(C1 : @SpecializedCategory objC1).
-  Context `(C2 : @SpecializedCategory objC2).
-  Context `(D : @SpecializedCategory objD).
-  Variable F1 : SpecializedFunctor C1 D.
-  Variable F2 : SpecializedFunctor C2 D.
-  Variable G : SpecializedFunctor C1 C2.
+  Context `(C1 : @Category objC1).
+  Context `(C2 : @Category objC2).
+  Variable D : Category.
+  Variable F1 : Functor C1 D.
+  Variable F2 : Functor C2 D.
+  Variable G : Functor C1 C2.
 
   Section Limit.
     Variable T : NaturalTransformation (ComposeFunctors F2 G) F1.
@@ -50,7 +50,7 @@
     Let limF1 := LimitObject F1_Limit.
     Let limF2 := LimitObject F2_Limit.
 
-    Definition InducedLimitMapNT' : SpecializedNaturalTransformation ((DiagonalFunctor D C1) limF2) F1.
+    Definition InducedLimitMapNT' : NaturalTransformation ((DiagonalFunctor D C1) limF2) F1.
       unfold LimitObject, Limit in *;
         intro_universal_morphisms.
       subst limF1 limF2.
@@ -61,8 +61,8 @@
       unfold ComposeFunctors at 1.
       simpl.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun x => Identity _)
             _
           )
@@ -70,12 +70,12 @@
       simpl; reflexivity.
     Defined.
 
-    Definition InducedLimitMapNT'' : SpecializedNaturalTransformation ((DiagonalFunctor D C1) limF2) F1.
+    Definition InducedLimitMapNT'' : NaturalTransformation ((DiagonalFunctor D C1) limF2) F1.
       simpl_definition_by_exact InducedLimitMapNT'.
     Defined.
 
     (* Then we clean up a bit with reduction. *)
-    Definition InducedLimitMapNT : SpecializedNaturalTransformation ((DiagonalFunctor D C1) limF2) F1
+    Definition InducedLimitMapNT : NaturalTransformation ((DiagonalFunctor D C1) limF2) F1
       := Eval cbv beta iota zeta delta [InducedLimitMapNT''] in InducedLimitMapNT''.
 
     Definition InducedLimitMap : D.(Morphism) limF2 limF1
@@ -91,7 +91,7 @@
     Let colimF1 := ColimitObject F1_Colimit.
     Let colimF2 := ColimitObject F2_Colimit.
 
-    Definition InducedColimitMapNT' : SpecializedNaturalTransformation F1 ((DiagonalFunctor D C1) colimF2).
+    Definition InducedColimitMapNT' : NaturalTransformation F1 ((DiagonalFunctor D C1) colimF2).
       unfold ColimitObject, Colimit in *;
         intro_universal_morphisms.
       subst colimF1 colimF2.
@@ -102,8 +102,8 @@
       unfold ComposeFunctors at 1.
       simpl.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun x => Identity _)
             _
           )
@@ -111,12 +111,12 @@
       simpl; reflexivity.
     Defined.
 
-    Definition InducedColimitMapNT'' : SpecializedNaturalTransformation F1 ((DiagonalFunctor D C1) colimF2).
+    Definition InducedColimitMapNT'' : NaturalTransformation F1 ((DiagonalFunctor D C1) colimF2).
       simpl_definition_by_exact InducedColimitMapNT'.
     Defined.
 
     (* Then we clean up a bit with reduction. *)
-    Definition InducedColimitMapNT : SpecializedNaturalTransformation F1 ((DiagonalFunctor D C1) colimF2)
+    Definition InducedColimitMapNT : NaturalTransformation F1 ((DiagonalFunctor D C1) colimF2)
       := Eval cbv beta iota zeta delta [InducedColimitMapNT''] in InducedColimitMapNT''.
 
     Definition InducedColimitMap : Morphism D colimF1 colimF2
diff -r 4eb6407c6ca3 LimitFunctors.v
--- a/LimitFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/LimitFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -12,21 +12,21 @@
 Local Open Scope category_scope.
 
 Section LimitFunctors.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variable C : Category.
+  Variable D : Category.
 
-  Definition HasLimits' := forall F : SpecializedFunctor D C, exists _ : Limit F, True.
-  Definition HasLimits := forall F : SpecializedFunctor D C, Limit F.
+  Definition HasLimits' := forall F : Functor D C, exists _ : Limit F, True.
+  Definition HasLimits := forall F : Functor D C, Limit F.
 
-  Definition HasColimits' := forall F : SpecializedFunctor D C, exists _ : Colimit F, True.
-  Definition HasColimits := forall F : SpecializedFunctor D C, Colimit F.
+  Definition HasColimits' := forall F : Functor D C, exists _ : Colimit F, True.
+  Definition HasColimits := forall F : Functor D C, Colimit F.
 
   Hypothesis HL : HasLimits.
   Hypothesis HC : HasColimits.
 
-  Definition LimitFunctor : SpecializedFunctor (C ^ D) C
+  Definition LimitFunctor : Functor (C ^ D) C
     := Eval unfold AdjointFunctorOfTerminalMorphism in AdjointFunctorOfTerminalMorphism HL.
-  Definition ColimitFunctor : SpecializedFunctor (C ^ D) C
+  Definition ColimitFunctor : Functor (C ^ D) C
     := Eval unfold AdjointFunctorOfInitialMorphism in AdjointFunctorOfInitialMorphism HC.
 
   Definition LimitAdjunction : Adjunction (DiagonalFunctor C D) LimitFunctor
diff -r 4eb6407c6ca3 Limits.v
--- a/Limits.v	Fri May 24 20:09:37 2013 -0400
+++ b/Limits.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor UniversalProperties.
+Require Export Category Functor UniversalProperties.
 Require Import Common FunctorCategory NaturalTransformation.
 
 Set Implicit Arguments.
@@ -12,8 +12,8 @@
 Local Open Scope category_scope.
 
 Section DiagonalFunctor.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variable C : Category.
+  Variable D : Category.
 
   (**
      Quoting Dwyer and Spalinski:
@@ -37,8 +37,8 @@
     simpl; unfold diagonal_functor_object_of; intro m.
     hnf.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun d => m : C.(Morphism) ((diagonal_functor_object_of o1) d) ((diagonal_functor_object_of o2) d))
           _
         )
@@ -46,10 +46,10 @@
       simpl; abstract (intros; autorewrite with morphism; trivial).
   Defined.
 
-  Definition DiagonalFunctor' : SpecializedFunctor C (C ^ D).
+  Definition DiagonalFunctor' : Functor C (C ^ D).
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           diagonal_functor_object_of
           diagonal_functor_morphism_of
           _
@@ -63,18 +63,18 @@
 End DiagonalFunctor.
 
 Section DiagonalFunctorLemmas.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(D' : @SpecializedCategory objD').
+  Variable C : Category.
+  Variable D : Category.
+  Context `(D' : @Category objD').
 
-  Lemma Compose_DiagonalFunctor x (F : SpecializedFunctor D' D) :
+  Lemma Compose_DiagonalFunctor x (F : Functor D' D) :
     ComposeFunctors (DiagonalFunctor C D x) F = DiagonalFunctor _ _ x.
     functor_eq.
   Qed.
 
-  Definition Compose_DiagonalFunctor_NT x (F : SpecializedFunctor D' D) :
-    SpecializedNaturalTransformation (ComposeFunctors (DiagonalFunctor C D x) F) (DiagonalFunctor _ _ x)
-    := Build_SpecializedNaturalTransformation (ComposeFunctors (DiagonalFunctor C D x) F)
+  Definition Compose_DiagonalFunctor_NT x (F : Functor D' D) :
+    NaturalTransformation (ComposeFunctors (DiagonalFunctor C D x) F) (DiagonalFunctor _ _ x)
+    := Build_NaturalTransformation (ComposeFunctors (DiagonalFunctor C D x) F)
                                               (DiagonalFunctor _ _ x)
                                               (fun z => Identity x)
                                               (fun _ _ _ => eq_refl).
@@ -83,9 +83,9 @@
 Hint Rewrite @Compose_DiagonalFunctor.
 
 Section Limit.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor D C.
+  Variable C : Category.
+  Variable D : Category.
+  Variable F : Functor D C.
 
   (**
      Quoting Dwyer and Spalinski:
@@ -98,8 +98,8 @@
   Definition Limit := TerminalMorphism (DiagonalFunctor C D) F.
   Definition IsLimit := IsTerminalMorphism (U := DiagonalFunctor C D) (X := F).
   (*  Definition Limit (L : C) :=
-    { t : SmallSpecializedNaturalTransformation ((DiagonalFunctor C D) L) F &
-      forall X : C, forall s : SmallSpecializedNaturalTransformation ((DiagonalFunctor C D) X) F,
+    { t : SmallNaturalTransformation ((DiagonalFunctor C D) L) F &
+      forall X : C, forall s : SmallNaturalTransformation ((DiagonalFunctor C D) X) F,
         { s' : C.(Morphism) X L |
           unique
           (fun s' => SNTComposeT t ((DiagonalFunctor C D).(MorphismOf) s') = s)
@@ -118,8 +118,8 @@
   Definition Colimit := @InitialMorphism (C ^ D) _ F (DiagonalFunctor C D).
   Definition IsColimit := @IsInitialMorphism (C ^ D) _ F (DiagonalFunctor C D).
   (*  Definition Colimit (c : C) :=
-    { t : SmallSpecializedNaturalTransformation F ((DiagonalFunctor C D) c) &
-      forall X : C, forall s : SmallSpecializedNaturalTransformation F ((DiagonalFunctor C D) X),
+    { t : SmallNaturalTransformation F ((DiagonalFunctor C D) c) &
+      forall X : C, forall s : SmallNaturalTransformation F ((DiagonalFunctor C D) X),
         { s' : C.(Morphism) c X | is_unique s' /\
           SNTComposeT ((DiagonalFunctor C D).(MorphismOf) s') t = s
         }
@@ -160,9 +160,9 @@
 End Limit.
 
 Section LimitMorphisms.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor D C.
+  Variable C : Category.
+  Variable D : Category.
+  Variable F : Functor D C.
 
   Definition MorphismBetweenLimits (L L' : Limit F) : C.(Morphism) (LimitObject L) (LimitObject L').
     unfold Limit, LimitObject in *.
diff -r 4eb6407c6ca3 LtacReifiedSimplification.v
--- a/LtacReifiedSimplification.v	Fri May 24 20:09:37 2013 -0400
+++ b/LtacReifiedSimplification.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import Bool.
-Require Import SpecializedCategory Functor NaturalTransformation.
+Require Import Category Functor NaturalTransformation.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -17,71 +17,70 @@
     repeat rewrite ?Associativity;
     repeat rewrite ?FCompositionOf.
 
+Local Open Scope morphism_scope.
+
 Section ReifiedMorphism.
-  Inductive ReifiedMorphism : forall objC (C : SpecializedCategory objC), C -> C -> Type :=
-  | ReifiedIdentityMorphism : forall objC C x, @ReifiedMorphism objC C x x
-  | ReifiedComposedMorphism : forall objC C s d d', ReifiedMorphism C d d' -> ReifiedMorphism C s d -> @ReifiedMorphism objC C s d'
-  | ReifiedNaturalTransformationMorphism : forall objB (B : SpecializedCategory objB)
-                                                  objC (C : SpecializedCategory objC)
-                                                  (F G : SpecializedFunctor B C)
-                                                  (T : SpecializedNaturalTransformation F G)
+  Inductive ReifiedMorphism : forall (C : Category), C -> C -> Type :=
+  | ReifiedIdentityMorphism : forall C x, @ReifiedMorphism C x x
+  | ReifiedComposedMorphism : forall C s d d', ReifiedMorphism C d d' -> ReifiedMorphism C s d -> ReifiedMorphism C s d'
+  | ReifiedNaturalTransformationMorphism : forall B C
+                                                  (F G : Functor B C)
+                                                  (T : NaturalTransformation F G)
                                                   x,
                                              ReifiedMorphism C (F x) (G x)
-  | ReifiedFunctorMorphism : forall objB (B : SpecializedCategory objB)
-                                    objC (C : SpecializedCategory objC)
-                                    (F : SpecializedFunctor B C)
+  | ReifiedFunctorMorphism : forall B C
+                                    (F : Functor B C)
                                     s d,
-                               @ReifiedMorphism objB B s d -> @ReifiedMorphism objC C (F s) (F d)
-  | ReifiedGenericMorphism : forall objC (C : SpecializedCategory objC) s d, Morphism C s d -> @ReifiedMorphism objC C s d.
+                               ReifiedMorphism B s d -> ReifiedMorphism C (F s) (F d)
+  | ReifiedGenericMorphism : forall C s d, Morphism C s d -> ReifiedMorphism C s d.
 
-  Fixpoint ReifiedMorphismDenote objC C s d (m : @ReifiedMorphism objC C s d) : Morphism C s d :=
-    match m in @ReifiedMorphism objC C s d return Morphism C s d with
-      | ReifiedIdentityMorphism _ _ x => Identity x
-      | ReifiedComposedMorphism _ _ _ _ _ m1 m2 => Compose (@ReifiedMorphismDenote _ _ _ _ m1)
-                                                           (@ReifiedMorphismDenote _ _ _ _ m2)
-      | ReifiedNaturalTransformationMorphism _ _ _ _ _ _ T x => T x
-      | ReifiedFunctorMorphism _ _ _ _ F _ _ m => MorphismOf F (@ReifiedMorphismDenote _ _ _ _ m)
-      | ReifiedGenericMorphism _ _ _ _ m => m
+  Fixpoint ReifiedMorphismDenote C s d (m : @ReifiedMorphism C s d) : Morphism C s d :=
+    match m in @ReifiedMorphism C s d return Morphism C s d with
+      | ReifiedIdentityMorphism _ x => Identity x
+      | ReifiedComposedMorphism _ _ _ _ m1 m2 => (@ReifiedMorphismDenote _ _ _ m1)
+                                                   ∘ (@ReifiedMorphismDenote _ _ _ m2)
+      | ReifiedNaturalTransformationMorphism _ _ _ _ T x => T x
+      | ReifiedFunctorMorphism _ _ F _ _ m => MorphismOf F (@ReifiedMorphismDenote _ _ _ m)
+      | ReifiedGenericMorphism _ _ _ m => m
     end.
 End ReifiedMorphism.
 
 Ltac Ltac_reify_morphism m :=
-  let objC := match type of m with @Morphism ?objC _ ?s ?d => constr:(objC) end in
-  let C := match type of m with @Morphism ?objC ?C ?s ?d => constr:(C) end in
-  let s := match type of m with @Morphism ?objC _ ?s ?d => constr:(s) end in
-  let d := match type of m with @Morphism ?objC _ ?s ?d => constr:(d) end in
+  let C := match type of m with @Morphism ?C ?s ?d => constr:(C) end in
+  let s := match type of m with @Morphism _ ?s ?d => constr:(s) end in
+  let d := match type of m with @Morphism _ ?s ?d => constr:(d) end in
   match m with
-    | @Identity _ _ ?x => constr:(@ReifiedIdentityMorphism _ C x)
+    | @Identity _ ?x => constr:(@ReifiedIdentityMorphism C x)
     | Compose ?m1 ?m2 => let m1' := Ltac_reify_morphism m1 in
                          let m2' := Ltac_reify_morphism m2 in
-                         constr:(@ReifiedComposedMorphism _ _ _ _ _ m1' m2')
-    | ComponentsOf ?T ?x => constr:(@ReifiedNaturalTransformationMorphism _ _ _ _ _ _ T x)
+                         constr:(@ReifiedComposedMorphism _ _ _ _ m1' m2')
+    | ComponentsOf ?T ?x => constr:(@ReifiedNaturalTransformationMorphism _ _ _ _ T x)
     | MorphismOf ?F ?m => let m' := Ltac_reify_morphism m in
-                          constr:(@ReifiedFunctorMorphism _ _ _ _ F _ _ m')
+                          constr:(@ReifiedFunctorMorphism _ _ F _ _ m')
     | ReifiedMorphismDenote ?m' => constr:(m')
-    | _ => constr:(@ReifiedGenericMorphism objC C s d m)
+    | _ => constr:(@ReifiedGenericMorphism C s d m)
   end.
 
 Section SimplifiedMorphism.
-  Fixpoint ReifiedHasIdentities `(m : @ReifiedMorphism objC C s d) : bool
+  Fixpoint ReifiedHasIdentities `(m : @ReifiedMorphism C s d) : bool
     := match m with
-         | ReifiedIdentityMorphism _ _ _ => true
-         | ReifiedComposedMorphism _ _ _ _ _ m1 m2 => orb (@ReifiedHasIdentities _ _ _ _ m1) (@ReifiedHasIdentities _ _ _ _ m2)
-         | ReifiedNaturalTransformationMorphism _ _ _ _ _ _ _ _ => false
-         | ReifiedFunctorMorphism _ _ _ _ _ _ _ m0 => (@ReifiedHasIdentities _ _ _ _ m0)
-         | ReifiedGenericMorphism _ _ _ _ _ => false
+         | ReifiedIdentityMorphism _ _ => true
+         | ReifiedComposedMorphism _ _ _ _ m1 m2 => orb (@ReifiedHasIdentities _ _ _ m1) (@ReifiedHasIdentities _ _ _ m2)
+         | ReifiedNaturalTransformationMorphism _ _ _ _ _ _ => false
+         | ReifiedFunctorMorphism _ _ _ _ _ m0 => (@ReifiedHasIdentities _ _ _ m0)
+         | ReifiedGenericMorphism _ _ _ _ => false
        end.
 
-  Fixpoint ReifiedMorphismSimplifyWithProof objC C s d (m : @ReifiedMorphism objC C s d) {struct m}
+  Fixpoint ReifiedMorphismSimplifyWithProof C s d (m : @ReifiedMorphism C s d) {struct m}
   : { m' : ReifiedMorphism C s d | ReifiedMorphismDenote m = ReifiedMorphismDenote m' }.
   refine match m with
-           | ReifiedComposedMorphism _ _ s0 d0 d0' m1 m2 => _
-           | ReifiedFunctorMorphism _ _ _ _ F _ _ m' => _
-           | ReifiedIdentityMorphism _ _ x => exist _ _ eq_refl
-           | ReifiedNaturalTransformationMorphism _ _ _ _ _ _ T x => exist _ _ eq_refl
-           | ReifiedGenericMorphism _ _ _ _ m => exist _ _ eq_refl
+           | ReifiedComposedMorphism _ s0 d0 d0' m1 m2 => _
+           | ReifiedFunctorMorphism _ _ F _ _ m' => _
+           | ReifiedIdentityMorphism _ x => exist _ _ eq_refl
+           | ReifiedNaturalTransformationMorphism _ _ _ _ T x => exist _ _ eq_refl
+           | ReifiedGenericMorphism _ _ _ m => exist _ _ eq_refl
          end; clear m;
-  [ destruct (@ReifiedMorphismSimplifyWithProof _ _ _ _ m1) as [ m1' H1 ], (@ReifiedMorphismSimplifyWithProof _ _ _ _ m2) as [ m2' H2 ];
+  [ destruct (@ReifiedMorphismSimplifyWithProof _ _ _ m1) as [ m1' H1 ], (@ReifiedMorphismSimplifyWithProof _ _ _ m2) as [ m2' H2 ];
     clear ReifiedMorphismSimplifyWithProof;
     destruct m1';
     ((destruct m2')
@@ -97,7 +96,7 @@
          match type of m2' with
            | ReifiedMorphism _ _ (?f ?x) => generalize dependent (f x); intros; destruct m2'
          end))
-    | destruct (@ReifiedMorphismSimplifyWithProof _ _ _ _ m') as [ m'' ? ];
+    | destruct (@ReifiedMorphismSimplifyWithProof _ _ _ m') as [ m'' ? ];
       clear ReifiedMorphismSimplifyWithProof;
       destruct m''];
   simpl in *;
@@ -179,16 +178,16 @@
       destruct H';
       simpl in H.
 
-  Fixpoint ReifiedMorphismSimplifyWithProofNoIdentity `(C : @SpecializedCategory objC) s d
+  Fixpoint ReifiedMorphismSimplifyWithProofNoIdentity C s d
            (m : ReifiedMorphism C s d)
   : {ReifiedHasIdentities (proj1_sig (ReifiedMorphismSimplifyWithProof m)) = false}
     + { exists H : s = d, proj1_sig (ReifiedMorphismSimplifyWithProof m) = match H with eq_refl => ReifiedIdentityMorphism C s end }.
   destruct m;
   try solve [ right; clear; abstract (exists eq_refl; subst; reflexivity)
             | left; clear; reflexivity ];
-  [ destruct (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ _ m1), (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ _ m2);
+  [ destruct (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ m1), (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ m2);
     [ left | left | left | right ]
-  | destruct (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ _ m);
+  | destruct (@ReifiedMorphismSimplifyWithProofNoIdentity _ _ _ m);
     [ left | right ] ];
   clear ReifiedMorphismSimplifyWithProofNoIdentity;
   try abstract (
@@ -202,8 +201,8 @@
         repeat match goal with
                  | _ => progress subst
                  | _ => progress solve_t'
-                 | [ H : ReifiedHasIdentities (?f ?x) = false |- _ ] => generalize dependent (f x); intros
-                                                                                                  | _ => progress simpl in *
+                 | [ H : ReifiedHasIdentities (?f ?x) = false |- _ ] => (generalize dependent (f x); intros)
+                 | _ => progress simpl in *
                end
       ).
   Defined.
@@ -211,14 +210,14 @@
   Section ReifiedMorphismSimplify.
     Local Arguments ReifiedMorphismSimplifyWithProof / .
 
-    Definition ReifiedMorphismSimplify objC C s d (m : @ReifiedMorphism objC C s d)
+    Definition ReifiedMorphismSimplify C s d (m : @ReifiedMorphism C s d)
     : ReifiedMorphism C s d
       := Eval simpl in proj1_sig (ReifiedMorphismSimplifyWithProof m).
 
     (*Local Arguments ReifiedMorphismSimplify / .*)
   End ReifiedMorphismSimplify.
 
-  Lemma ReifiedMorphismSimplifyOk objC C s d (m : @ReifiedMorphism objC C s d)
+  Lemma ReifiedMorphismSimplifyOk C s d (m : @ReifiedMorphism C s d)
   : ReifiedMorphismDenote m =
     ReifiedMorphismDenote (ReifiedMorphismSimplify m).
   Proof.
@@ -227,7 +226,7 @@
 
   Local Arguments ReifiedMorphismSimplifyWithProofNoIdentity / .
 
-  Definition ReifiedMorphismSimplifyNoIdentity `(C : @SpecializedCategory objC) s d
+  Definition ReifiedMorphismSimplifyNoIdentity C s d
              (m : ReifiedMorphism C s d)
   : {ReifiedHasIdentities (ReifiedMorphismSimplify m) = false}
     + { exists H : s = d, ReifiedMorphismSimplify m = match H with eq_refl => ReifiedIdentityMorphism C s end }
@@ -252,66 +251,56 @@
 (*******************************************************************************)
 Section good_examples.
   Section id.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objC).
-    Variables F G : SpecializedFunctor C D.
-    Variable T : SpecializedNaturalTransformation F G.
+    Variables C D : Category.
+    Variables F G : Functor C D.
+    Variable T : NaturalTransformation F G.
 
-    Lemma good_example_00001 (x : C) :Compose (Identity x) (Identity x) = Identity (C := C) x.
+    Lemma good_example_00001 (x : C) : Identity x ∘ Identity x = Identity (C := C) x.
       Ltac_rsimplify_morphisms.
       reflexivity.
     Qed.
 
     Lemma good_example_00002 s d (m : Morphism C s d)
-    : MorphismOf F (Compose m (Identity s)) = MorphismOf F m.
+    : MorphismOf F (m ∘ Identity s) = MorphismOf F m.
       Ltac_rsimplify_morphisms.
       reflexivity.
     Qed.
 
     Lemma good_example_00003 s d (m : Morphism C s d)
-    : MorphismOf F (Compose (Identity d) m) = MorphismOf F m.
+    : MorphismOf F (Identity d ∘ m) = MorphismOf F m.
       Ltac_rsimplify_morphisms.
       reflexivity.
     Qed.
   End id.
 
   Lemma good_example_00004
-  : forall (objC : Type) (C : SpecializedCategory objC)
-           (objD : Type) (D : SpecializedCategory objD) (objC' : Type)
-           (C' : SpecializedCategory objC') (objD' : Type)
-           (D' : SpecializedCategory objD') (F : SpecializedFunctor C C')
-           (G : SpecializedFunctor D' D) (s d d' : SpecializedFunctor D C)
-           (m1 : SpecializedNaturalTransformation s d)
-           (m2 : SpecializedNaturalTransformation d d') (x : objD'),
-      Compose (MorphismOf F (m2 (G x))) (MorphismOf F (m1 (G x))) =
-      Compose
-        (Compose
-           (MorphismOf F (Compose (MorphismOf d' (Identity (G x))) (m2 (G x))))
-           (Identity (F (d (G x))))) (MorphismOf F (m1 (G x))).
+  : forall C D C' D'
+           (F : Functor C C')
+           (G : Functor D' D) (s d d' : Functor D C)
+           (m1 : NaturalTransformation s d)
+           (m2 : NaturalTransformation d d') (x : D'),
+      MorphismOf F (m2 (G x)) ∘ MorphismOf F (m1 (G x)) =
+      (MorphismOf F (MorphismOf d' (Identity (G x)) ∘ m2 (G x))
+                  ∘ Identity (F (d (G x)))) ∘ MorphismOf F (m1 (G x)).
+  Proof.
     intros.
     Ltac_rsimplify_morphisms.
     reflexivity.
   Qed.
 
   Lemma good_example_00005
-  : forall (objC : Type) (C : SpecializedCategory objC)
-           (objD : Type) (D : SpecializedCategory objD) (objC' : Type)
-           (C' : SpecializedCategory objC') (objD' : Type)
-           (D' : SpecializedCategory objD') (F : SpecializedFunctor C C')
-           (G : SpecializedFunctor D' D) (s d d' : SpecializedFunctor D C)
-           (m1 : SpecializedNaturalTransformation s d)
-           (m2 : SpecializedNaturalTransformation d d') (x : objD'),
-      Compose
-        (MorphismOf F
-                    (Compose (MorphismOf d' (Identity (G x)))
-                             (Compose (m2 (G x)) (m1 (G x))))) (Identity (F (s (G x)))) =
-      Compose
-        (Compose
-           (MorphismOf F (Compose (MorphismOf d' (Identity (G x))) (m2 (G x))))
-           (Identity (F (d (G x)))))
-        (Compose
-           (MorphismOf F (Compose (MorphismOf d (Identity (G x))) (m1 (G x))))
-           (Identity (F (s (G x))))).
+  : forall C D C' D'
+           (F : Functor C C')
+           (G : Functor D' D) (s d d' : Functor D C)
+           (m1 : NaturalTransformation s d)
+           (m2 : NaturalTransformation d d') (x : D'),
+      MorphismOf F (MorphismOf d' (Identity (G x)) ∘ m2 (G x) ∘ m1 (G x))
+                 ∘ Identity (F (s (G x))) =
+      (MorphismOf F (MorphismOf d' (Identity (G x)) ∘ m2 (G x))
+                  ∘ Identity (F (d (G x))))
+        ∘ MorphismOf F (MorphismOf d (Identity (G x)) ∘ m1 (G x))
+        ∘ Identity (F (s (G x))).
+  Proof.
     intros.
     Ltac_rsimplify_morphisms.
     reflexivity.
@@ -324,16 +313,14 @@
 Section bad_examples.
   Require Import SumCategory.
   Section bad_example_0001.
-    Context `(C0 : SpecializedCategory objC0).
-    Context `(C1 : SpecializedCategory objC1).
-    Context `(D : SpecializedCategory objD).
+    Variables C0 C1 D : Category.
 
     Variables s d d' : C0.
     Variable m1 : Morphism C0 s d.
     Variable m2 : Morphism C0 d d'.
-    Variable F : SpecializedFunctor (C0 + C1) D.
+    Variable F : Functor (C0 + C1) D.
 
-    Goal MorphismOf F (s := inl _) (d := inl _) (Compose m2 m1) = Compose (MorphismOf F (s := inl _) (d := inl _) m2) (MorphismOf F (s := inl _) (d := inl _) m1).
+    Goal MorphismOf F (s := inl _) (d := inl _) (m2 ∘ m1) = (MorphismOf F (s := inl _) (d := inl _) m2) ∘ (MorphismOf F (s := inl _) (d := inl _) m1).
     Ltac_rsimplify_morphisms.
     Fail reflexivity.
     Abort.
diff -r 4eb6407c6ca3 Makefile
--- a/Makefile	Fri May 24 20:09:37 2013 -0400
+++ b/Makefile	Fri Jun 21 15:07:08 2013 -0400
@@ -12,7 +12,7 @@
 	EqdepFacts_one_variable \
 	Eqdep_dec_one_variable \
 	\
-	SpecializedCategory \
+	Category \
 	Functor \
 	NaturalTransformation \
 	ProductCategory \
@@ -23,8 +23,6 @@
 	TypeclassUnreifiedSimplification \
 	CanonicalStructureSimplification \
 	\
-	Category \
-	CategoryEquality \
 	CategoryIsomorphisms \
 	FunctorIsomorphisms \
 	NaturalEquivalence \
@@ -54,9 +52,8 @@
 	SetCategory \
 	SetCategoryProductFunctor \
 	InitialTerminalCategory \
-	SmallCat \
+	Cat \
 	CommaCategory \
-	SpecializedCommaCategory \
 	LaxCommaCategory \
 	SpecializedLaxCommaCategory \
 	CommaCategoryProjection \
diff -r 4eb6407c6ca3 MonoidalCategory.v
--- a/MonoidalCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/MonoidalCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import ProofIrrelevance.
-Require Export SpecializedCategory Functor ProductCategory FunctorProduct NaturalTransformation NaturalEquivalence.
+Require Export Category Functor ProductCategory FunctorProduct NaturalTransformation NaturalEquivalence.
 Require Import Common Notations StructureEquality DefinitionSimplification.
 
 Set Implicit Arguments.
@@ -17,12 +17,12 @@
   (** Quoting Wikipedia:
      A  monoidal category is a category [C] equipped with
      *)
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   (**
      - a bifunctor [ ⊗ : C × C -> C] called the tensor product or
          monoidal product,
     *)
-  Variable TensorProduct : SpecializedFunctor (C * C) C.
+  Variable TensorProduct : Functor (C * C) C.
 
   Let src {s d} (_ : Morphism C s d) := s.
   Let dst {s d} (_ : Morphism C s d) := d.
@@ -132,8 +132,8 @@
     Morphism C (TriMonoidalProductR_ObjectOf s) (TriMonoidalProductR_ObjectOf d)
     := Eval cbv beta iota zeta delta [TriMonoidalProductR_MorphismOf''] in @TriMonoidalProductR_MorphismOf'' s d m.
 
-  Definition TriMonoidalProductL' : SpecializedFunctor (C * C * C) C.
-    refine (Build_SpecializedFunctor (C * C * C) C
+  Definition TriMonoidalProductL' : Functor (C * C * C) C.
+    refine (Build_Functor (C * C * C) C
       TriMonoidalProductL_ObjectOf
       TriMonoidalProductL_MorphismOf
       _
@@ -148,8 +148,8 @@
     ).
   Defined.
 
-  Definition TriMonoidalProductR' : SpecializedFunctor (C * C * C) C.
-    refine (Build_SpecializedFunctor (C * C * C) C
+  Definition TriMonoidalProductR' : Functor (C * C * C) C.
+    refine (Build_Functor (C * C * C) C
       TriMonoidalProductR_ObjectOf
       TriMonoidalProductR_MorphismOf
       _
@@ -164,11 +164,11 @@
     ).
   Defined.
 
-  Definition TriMonoidalProductL'' : SpecializedFunctor (C * C * C) C.
+  Definition TriMonoidalProductL'' : Functor (C * C * C) C.
     simpl_definition_by_tac_and_exact TriMonoidalProductL' ltac:(idtac; subst_body).
   Defined.
 
-  Definition TriMonoidalProductR'' : SpecializedFunctor (C * C * C) C.
+  Definition TriMonoidalProductR'' : Functor (C * C * C) C.
     simpl_definition_by_tac_and_exact TriMonoidalProductR' ltac:(idtac; subst_body).
   Defined.
 
@@ -195,7 +195,7 @@
            right unitor, with components
            [λ A : I ⊗ A ~= A] and [ρ A : A ⊗ I ~= A].
      *)
-  Definition LeftUnitorFunctor : SpecializedFunctor C C.
+  Definition LeftUnitorFunctor : Functor C C.
     clear TriMonoidalProductL_MorphismOf TriMonoidalProductR_MorphismOf
       TriMonoidalProductL_MorphismOf' TriMonoidalProductR_MorphismOf'
       TriMonoidalProductL_MorphismOf'' TriMonoidalProductR_MorphismOf''
@@ -213,7 +213,7 @@
     ).
   Defined.
 
-  Definition RightUnitorFunctor : SpecializedFunctor C C.
+  Definition RightUnitorFunctor : Functor C C.
     clear TriMonoidalProductL_MorphismOf TriMonoidalProductR_MorphismOf
       TriMonoidalProductL_MorphismOf' TriMonoidalProductR_MorphismOf'
       TriMonoidalProductL_MorphismOf'' TriMonoidalProductR_MorphismOf''
@@ -357,15 +357,15 @@
   Local Reserved Notation "'λ'".
   Local Reserved Notation "'ρ'".
 
-  Let src (C : @SpecializedCategory objC) {s d} (_ : Morphism C s d) := s.
-  Let dst (C : @SpecializedCategory objC) s d (_ : Morphism C s d) := d.
+  Let src (C : @Category objC) {s d} (_ : Morphism C s d) := s.
+  Let dst (C : @Category objC) s d (_ : Morphism C s d) := d.
 
   Let AssociatorCoherenceCondition' := Eval unfold AssociatorCoherenceCondition in @AssociatorCoherenceCondition.
   Let UnitorCoherenceCondition' := Eval unfold UnitorCoherenceCondition in @UnitorCoherenceCondition.
 
   Record MonoidalCategory := {
-    MonoidalUnderlyingCategory :> @SpecializedCategory objC;
-    TensorProduct : SpecializedFunctor (MonoidalUnderlyingCategory * MonoidalUnderlyingCategory) MonoidalUnderlyingCategory
+    MonoidalUnderlyingCategory :> @Category objC;
+    TensorProduct : Functor (MonoidalUnderlyingCategory * MonoidalUnderlyingCategory) MonoidalUnderlyingCategory
       where "A ⊗ B" := (TensorProduct (A, B)) and "A ⊗m B" := (TensorProduct.(MorphismOf) (s := (@src _ _ _ A, @src _ _ _ B)) (d := (@dst _ _ _ A, @dst _ _ _ B)) (A, B)%morphism);
 
     IdentityObject : objC where "'I'" := IdentityObject;
diff -r 4eb6407c6ca3 NatCategory.v
--- a/NatCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/NatCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory DiscreteCategory.
+Require Export Category DiscreteCategory.
 Require Import Common.
 
 Fixpoint CardinalityRepresentative (n : nat) : Set :=
@@ -12,4 +12,4 @@
 
 Definition NatCategory (n : nat) := Eval unfold DiscreteCategory, DiscreteCategory_Identity in DiscreteCategory n.
 
-Coercion NatCategory : nat >-> SpecializedCategory.
+Coercion NatCategory : nat >-> Category.
diff -r 4eb6407c6ca3 NaturalEquivalence.v
--- a/NaturalEquivalence.v	Fri May 24 20:09:37 2013 -0400
+++ b/NaturalEquivalence.v	Fri Jun 21 15:07:08 2013 -0400
@@ -17,21 +17,19 @@
 
 Section NaturalIsomorphism.
   Section NaturalIsomorphism.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Variables F G : SpecializedFunctor C D.
+    Variables C D : Category.
+    Variables F G : Functor C D.
 
     Record NaturalIsomorphism :=
       {
-        NaturalIsomorphism_Transformation :> SpecializedNaturalTransformation F G;
-        NaturalIsomorphism_Isomorphism : forall x : objC, IsomorphismOf (NaturalIsomorphism_Transformation x)
+        NaturalIsomorphism_Transformation :> NaturalTransformation F G;
+        NaturalIsomorphism_Isomorphism : forall x : C, IsomorphismOf (NaturalIsomorphism_Transformation x)
       }.
   End NaturalIsomorphism.
 
   Section Inverse.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Variables F G : SpecializedFunctor C D.
+    Variables C D : Category.
+    Variables F G : Functor C D.
 
     Definition InverseNaturalIsomorphism_NT (T : NaturalIsomorphism F G) : NaturalTransformation G F.
       exists (fun x => Inverse (NaturalIsomorphism_Isomorphism T x)).
@@ -43,7 +41,7 @@
                  end;
           destruct T as [ [ ] ];
           simpl in *;
-          pre_compose_to_identity; post_compose_to_identity;
+            pre_compose_to_identity; post_compose_to_identity;
           auto with functor
         ).
     Defined.
@@ -56,11 +54,11 @@
   End Inverse.
 
   Section Composition.
-    Context `(C : @SpecializedCategory objC).
-    Context `(D : @SpecializedCategory objD).
-    Context `(E : @SpecializedCategory objE).
-    Variables F F' F'' : SpecializedFunctor C D.
-    Variables G G' : SpecializedFunctor D E.
+    Variable C : Category.
+    Variable D : Category.
+    Variable E : Category.
+    Variables F F' F'' : Functor C D.
+    Variables G G' : Functor D E.
 
     Local Ltac t :=
       simpl;
@@ -94,25 +92,25 @@
 End NaturalIsomorphism.
 
 Section NaturalIsomorphismOfCategories.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variable C : Category.
+  Variable D : Category.
 
   Local Reserved Notation "'F'".
   Local Reserved Notation "'G'".
 
   Record NaturalIsomorphismOfCategories := {
-    NaturalIsomorphismOfCategories_F : SpecializedFunctor C D where "'F'" := NaturalIsomorphismOfCategories_F;
-    NaturalIsomorphismOfCategories_G : SpecializedFunctor D C where "'G'" := NaturalIsomorphismOfCategories_G;
-    NaturalIsomorphismOfCategories_Isomorphism_C :> NaturalIsomorphism (IdentityFunctor C) (ComposeFunctors G F);
-    NaturalIsomorphismOfCategories_Isomorphism_D : NaturalIsomorphism (IdentityFunctor D) (ComposeFunctors F G)
-  }.
+                                            NaturalIsomorphismOfCategories_F : Functor C D where "'F'" := NaturalIsomorphismOfCategories_F;
+                                            NaturalIsomorphismOfCategories_G : Functor D C where "'G'" := NaturalIsomorphismOfCategories_G;
+                                            NaturalIsomorphismOfCategories_Isomorphism_C :> NaturalIsomorphism (IdentityFunctor C) (ComposeFunctors G F);
+                                            NaturalIsomorphismOfCategories_Isomorphism_D : NaturalIsomorphism (IdentityFunctor D) (ComposeFunctors F G)
+                                          }.
 
   Record IsomorphismOfCategories := {
-    IsomorphismOfCategories_F : SpecializedFunctor C D where "'F'" := IsomorphismOfCategories_F;
-    IsomorphismOfCategories_G : SpecializedFunctor D C where "'G'" := IsomorphismOfCategories_G;
-    IsomorphismOfCategories_Isomorphism_C : ComposeFunctors G F = IdentityFunctor C;
-    IsomorphismOfCategories_Isomorphism_D : ComposeFunctors F G = IdentityFunctor D
-  }.
+                                     IsomorphismOfCategories_F : Functor C D where "'F'" := IsomorphismOfCategories_F;
+                                     IsomorphismOfCategories_G : Functor D C where "'G'" := IsomorphismOfCategories_G;
+                                     IsomorphismOfCategories_Isomorphism_C : ComposeFunctors G F = IdentityFunctor C;
+                                     IsomorphismOfCategories_Isomorphism_D : ComposeFunctors F G = IdentityFunctor D
+                                   }.
 End NaturalIsomorphismOfCategories.
 
 Section NaturalEquivalence.
@@ -127,17 +125,17 @@
     exists T : NaturalTransformation F G, exists NE : NaturalEquivalenceOf T, True.
 End NaturalEquivalence.
 
-(* grumble, grumble, grumble, [Functors] take [SpecializedCategories] and so we don't get type inference.
-   Perhaps I should fix this?  But I don't like having to prefix so many things with [Specialized] and
+(* grumble, grumble, grumble, [Functors] take [Categories] and so we don't get type inference.
+   Perhaps I should fix this?  But I don't like having to prefix so many things with [] and
    having duplicated code just so I get type inference.
-   *)
+ *)
 Arguments NaturalEquivalenceOf [C D F G] T.
 Arguments FunctorsNaturallyEquivalent [C D] F G.
 
 Section Coercions.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variables F G : SpecializedFunctor C D.
+  Variable C : Category.
+  Variable D : Category.
+  Variables F G : Functor C D.
   Variable C' : Category.
   Variable D' : Category.
   Variables F' G' : Functor C' D'.
@@ -181,13 +179,13 @@
 
   Definition NaturalEquivalenceInverse : NaturalEquivalenceOf T -> NaturalTransformation G F.
     refine (fun X => {| ComponentsOf := (fun c => proj1_sig (X c)) |});
-      abstract (
+    abstract (
         intros; destruct (X s); destruct (X d);
-          simpl; unfold InverseOf in *; destruct_hypotheses;
-            pre_compose_to_identity; post_compose_to_identity;
-            auto
+        simpl; unfold InverseOf in *; destruct_hypotheses;
+        pre_compose_to_identity; post_compose_to_identity;
+        auto
       ).
-    (*morphisms 2*)
+  (*morphisms 2*)
   Defined.
 
   Lemma NaturalEquivalenceInverse_NaturalEquivalence (TE : NaturalEquivalenceOf T) : NaturalEquivalenceOf (NaturalEquivalenceInverse TE).
@@ -228,19 +226,19 @@
   Lemma functors_naturally_equivalent_trans (F G H : Functor C D) :
     FunctorsNaturallyEquivalent F G -> FunctorsNaturallyEquivalent G H -> FunctorsNaturallyEquivalent F H.
     destruct 1 as [ T [ ] ]; destruct 1 as [ U [ ] ];
-      exists (NTComposeT U T); eexists; hnf; simpl; eauto with category.
+    exists (NTComposeT U T); eexists; hnf; simpl; eauto with category.
   Qed.
 End FunctorNaturalEquivalenceRelation.
 
 Add Parametric Relation (C D : Category) : _ (@FunctorsNaturallyEquivalent C D)
-  reflexivity proved by (@functors_naturally_equivalent_refl _ _)
-  symmetry proved by (@functors_naturally_equivalent_sym _ _)
-  transitivity proved by (@functors_naturally_equivalent_trans _ _)
-    as functors_naturally_equivalent.
+    reflexivity proved by (@functors_naturally_equivalent_refl _ _)
+    symmetry proved by (@functors_naturally_equivalent_sym _ _)
+    transitivity proved by (@functors_naturally_equivalent_trans _ _)
+      as functors_naturally_equivalent.
 
 Add Parametric Morphism (C D E : Category) :
   (ComposeFunctors (C := C) (D := D) (E := E))
-  with signature (@FunctorsNaturallyEquivalent _ _) ==> (@FunctorsNaturallyEquivalent _ _) ==> (@FunctorsNaturallyEquivalent _ _) as functor_n_eq_mor.
+    with signature (@FunctorsNaturallyEquivalent _ _) ==> (@FunctorsNaturallyEquivalent _ _) ==> (@FunctorsNaturallyEquivalent _ _) as functor_n_eq_mor.
   intros;
   destruct_hypotheses;
   eexists (NTComposeF _ _);
@@ -305,9 +303,9 @@
 
   Lemma categories_naturally_equivalent_refl C : CategoriesNaturallyEquivalent C C.
     repeat (exists (IdentityFunctor C)); split;
-      match goal with
-        | [ |- FunctorsNaturallyEquivalent ?a ?b ] => cut (a = b); try let H' := fresh in solve [ intro H'; rewrite <- H'; reflexivity || trivial ]
-      end; functor_eq.
+    match goal with
+      | [ |- FunctorsNaturallyEquivalent ?a ?b ] => cut (a = b); try let H' := fresh in solve [ intro H'; rewrite <- H'; reflexivity || trivial ]
+    end; functor_eq.
   Qed.
 
   Lemma categories_naturally_equivalent_sym C D :
@@ -335,7 +333,7 @@
 End CategoryNaturalEquivalenceRelation.
 
 Add Parametric Relation : _ CategoriesNaturallyEquivalent
-  reflexivity proved by categories_naturally_equivalent_refl
-  symmetry proved by categories_naturally_equivalent_sym
-  transitivity proved by categories_naturally_equivalent_trans
-    as categories_naturally_equivalent.
+    reflexivity proved by categories_naturally_equivalent_refl
+    symmetry proved by categories_naturally_equivalent_sym
+    transitivity proved by categories_naturally_equivalent_trans
+      as categories_naturally_equivalent.
diff -r 4eb6407c6ca3 NaturalNumbersObject.v
--- a/NaturalNumbersObject.v	Fri May 24 20:09:37 2013 -0400
+++ b/NaturalNumbersObject.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory.
+Require Export Category.
 Require Import CategoryIsomorphisms.
 
 Set Implicit Arguments.
@@ -90,7 +90,7 @@
           [NaturalNumbersPreObject], to make the distinction slightly
           more obvious. *)
 
-  Context `(E : SpecializedCategory objE).
+  Variable E : Category.
 
   Local Reserved Notation "'ℕ'".
   Local Reserved Notation "'S'".
diff -r 4eb6407c6ca3 NaturalTransformation.v
--- a/NaturalTransformation.v	Fri May 24 20:09:37 2013 -0400
+++ b/NaturalTransformation.v	Fri Jun 21 15:07:08 2013 -0400
@@ -12,19 +12,23 @@
 
 Local Infix "==" := JMeq.
 
-Section SpecializedNaturalTransformation.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variables F G : SpecializedFunctor C D.
+Section NaturalTransformation.
+  Local Open Scope morphism_scope.
+
+  Variables C D : Category.
+  Variables F G : Functor C D.
 
   (**
      Quoting from the lecture notes for 18.705, Commutative Algebra:
 
-     A map of functors is known as a natural transformation. Namely, given two functors
-     [F : C -> D], [G : C -> D], a natural transformation [T: F -> G] is a collection of maps
-     [T A : F A -> G A], one for each object [A] of [C], such that [(T B) ○ (F m) = (G m) ○ (T A)]
-     for every map [m : A -> B] of [C]; that is, the following diagram is commutative:
+     A map of functors is known as a natural transformation. Namely,
+     given two functors [F : C -> D], [G : C -> D], a natural
+     transformation [T: F -> G] is a collection of maps [T A : F A ->
+     G A], one for each object [A] of [C], such that [(T B) ∘ (F m) =
+     (G m) ∘ (T A)] for every map [m : A -> B] of [C]; that is, the
+     following diagram is commutative:
 
+<<
            F m
      F A -------> F B
       |            |
@@ -33,48 +37,28 @@
       |            |
       V    G m     V
      G A --------> G B
+>>
      **)
-  Record SpecializedNaturalTransformation :=
+  Record NaturalTransformation :=
     {
       ComponentsOf :> forall c, D.(Morphism) (F c) (G c);
       Commutes : forall s d (m : C.(Morphism) s d),
-                   Compose (ComponentsOf d) (F.(MorphismOf) m) = Compose (G.(MorphismOf) m) (ComponentsOf s)
+                   ComponentsOf d ∘ F.(MorphismOf) m = G.(MorphismOf) m ∘ ComponentsOf s
     }.
-End SpecializedNaturalTransformation.
-
-Section NaturalTransformation.
-  Variable C : Category.
-  Variable D : Category.
-  Variable F G : Functor C D.
-
-  Definition NaturalTransformation := SpecializedNaturalTransformation F G.
 End NaturalTransformation.
 
-Bind Scope natural_transformation_scope with SpecializedNaturalTransformation.
 Bind Scope natural_transformation_scope with NaturalTransformation.
 
 Create HintDb natural_transformation discriminated.
 
-Arguments ComponentsOf {objC%type C%category objD%type D%category F%functor G%functor} T%natural_transformation c%object : rename, simpl nomatch.
-Arguments Commutes [objC C objD D F G] T _ _ _ : rename.
-
-Identity Coercion NaturalTransformation_SpecializedNaturalTransformation_Id : NaturalTransformation >-> SpecializedNaturalTransformation.
-Definition GeneralizeNaturalTransformation `(T : @SpecializedNaturalTransformation objC C objD D F G) :
-  NaturalTransformation F G := T.
-Global Coercion GeneralizeNaturalTransformation : SpecializedNaturalTransformation >-> NaturalTransformation.
-
-Arguments GeneralizeNaturalTransformation [objC C objD D F G] T /.
-Hint Extern 0 => unfold GeneralizeNaturalTransformation : category natural_transformation.
-Ltac fold_NT :=
-  change @SpecializedNaturalTransformation with
-    (fun objC (C : SpecializedCategory objC) objD (D : SpecializedCategory objD) => @NaturalTransformation C D) in *; simpl in *.
+Arguments ComponentsOf {C%category D%category F%functor G%functor} T%natural_transformation c%object : rename, simpl nomatch.
+Arguments Commutes [C D F G] T _ _ _ : rename.
 
 Hint Resolve @Commutes : category natural_transformation.
 
 Ltac nt_hideProofs :=
   repeat match goal with
              | [ |- context[{|
-                               ComponentsOf := _;
                                Commutes := ?pf
                              |}] ] =>
                hideProofs pf
@@ -88,87 +72,89 @@
     hnf in H;
     revert H; clear; intro H; clear H;
     match T'' with
-      | @Build_SpecializedNaturalTransformation ?objC ?C
-                                                ?objD ?D
-                                                ?F
-                                                ?G
-                                                ?CO
-                                                ?COM =>
-        refine (@Build_SpecializedNaturalTransformation objC C objD D F G
-                                                        CO
-                                                        _);
+      | @Build_NaturalTransformation ?C
+                                     ?D
+                                     ?F
+                                     ?G
+                                     ?CO
+                                     ?COM =>
+        refine (@Build_NaturalTransformation C D F G
+                                             CO
+                                             _);
           abstract exact COM
     end.
 Ltac nt_abstract_trailing_props T := nt_tac_abstract_trailing_props T ltac:(fun T' => T').
 Ltac nt_simpl_abstract_trailing_props T := nt_tac_abstract_trailing_props T ltac:(fun T' => let T'' := eval simpl in T' in T'').
 
 Section NaturalTransformations_Equal.
-  Lemma NaturalTransformation_contr_eq' objC C objD D F G
-        (T U : @SpecializedNaturalTransformation objC C objD D F G)
+  Lemma NaturalTransformation_contr_eq' C D F G
+        (T U : @NaturalTransformation C D F G)
         (D_morphism_proof_irrelevance
          : forall s d (m1 m2 : Morphism D s d) (pf1 pf2 : m1 = m2),
              pf1 = pf2)
   : ComponentsOf T = ComponentsOf U
     -> T = U.
-    destruct T, U; simpl; intros; repeat subst;
+  Proof.
+    intros.
+    destruct_head NaturalTransformation; simpl; intros; repeat subst;
     f_equal;
     repeat (apply functional_extensionality_dep; intro).
     trivial.
   Qed.
 
-  Lemma NaturalTransformation_contr_eq objC C objD D F G
-        (T U : @SpecializedNaturalTransformation objC C objD D F G)
+  Lemma NaturalTransformation_contr_eq C D F G
+        (T U : @NaturalTransformation C D F G)
         (D_morphism_proof_irrelevance
          : forall s d (m1 m2 : Morphism D s d) (pf1 pf2 : m1 = m2),
              pf1 = pf2)
   : (forall x, ComponentsOf T x = ComponentsOf U x)
     -> T = U.
+  Proof.
     intros; apply NaturalTransformation_contr_eq'; try assumption.
     destruct T, U; simpl in *;
     repeat (apply functional_extensionality_dep; intro);
     trivial.
   Qed.
 
-  Lemma NaturalTransformation_eq' objC C objD D F G :
-    forall (T U : @SpecializedNaturalTransformation objC C objD D F G),
-    ComponentsOf T = ComponentsOf U
-    -> T = U.
+  Lemma NaturalTransformation_eq' C D F G
+  : forall (T U : @NaturalTransformation C D F G),
+      ComponentsOf T = ComponentsOf U
+      -> T = U.
+  Proof.
     intros T U.
     apply NaturalTransformation_contr_eq'.
     intros; apply proof_irrelevance.
   Qed.
 
-  Lemma NaturalTransformation_eq objC C objD D F G :
-    forall (T U : @SpecializedNaturalTransformation objC C objD D F G),
-    (forall x, ComponentsOf T x = ComponentsOf U x)
-    -> T = U.
+  Lemma NaturalTransformation_eq C D F G
+  : forall (T U : @NaturalTransformation C D F G),
+      (forall x, ComponentsOf T x = ComponentsOf U x)
+      -> T = U.
+  Proof.
     intros T U.
     apply NaturalTransformation_contr_eq.
     intros; apply proof_irrelevance.
   Qed.
 
-  Lemma NaturalTransformation_JMeq' objC C objD D objC' C' objD' D' :
-    forall F G F' G'
-      (T : @SpecializedNaturalTransformation objC C objD D F G) (U : @SpecializedNaturalTransformation objC' C' objD' D' F' G'),
-      objC = objC'
-      -> objD = objD'
-      -> C == C'
-      -> D == D'
+  Lemma NaturalTransformation_JMeq' C D C' D'
+  : forall F G F' G'
+           (T : @NaturalTransformation C D F G) (U : @NaturalTransformation C' D' F' G'),
+      C = C'
+      -> D = D'
       -> F == F'
       -> G == G'
       -> ComponentsOf T == ComponentsOf U
       -> T == U.
     simpl; intros; intuition; destruct T, U; simpl in *; repeat subst;
-      JMeq_eq.
+    JMeq_eq.
     f_equal; apply proof_irrelevance.
   Qed.
-  Lemma NaturalTransformation_JMeq objC C objD D objC' C' objD' D' :
-    forall F G F' G'
-      (T : @SpecializedNaturalTransformation objC C objD D F G) (U : @SpecializedNaturalTransformation objC' C' objD' D' F' G'),
-      objC = objC'
-      -> objD = objD'
-      -> C == C'
-      -> D == D'
+
+  Lemma NaturalTransformation_JMeq C D C' D'
+  : forall F G F' G'
+           (T : @NaturalTransformation C D F G) (U : @NaturalTransformation C' D' F' G'),
+      C = C'
+      -> D = D'
       -> F == F'
       -> G == G'
       -> (forall x x', x == x' -> ComponentsOf T x == ComponentsOf U x')
@@ -197,38 +183,38 @@
     hnf in H;
     revert H; clear; intro H; clear H;
     match T'' with
-      | @Build_SpecializedNaturalTransformation ?objC ?C
-                                                ?objD ?D
-                                                ?F
-                                                ?G
-                                                ?CO
-                                                ?COM =>
+      | @Build_NaturalTransformation ?C
+                                     ?D
+                                     ?F
+                                     ?G
+                                     ?CO
+                                     ?COM =>
         let COM' := fresh in
         let COMT' := type of COM in
         let COMT := (eval simpl in COMT') in
         assert (COM' : COMT) by abstract exact COM;
-          exists (@Build_SpecializedNaturalTransformation objC C objD D F G
-                                                          CO
-                                                          COM');
+          exists (@Build_NaturalTransformation C D F G
+                                               CO
+                                               COM');
           expand; abstract (apply thm; reflexivity) || (apply thm; try reflexivity)
     end.
 Ltac nt_tac_abstract_trailing_props_with_equality tac :=
   pre_abstract_trailing_props;
   match goal with
-    | [ |- { T0 : SpecializedNaturalTransformation _ _ | T0 = ?T } ] =>
+    | [ |- { T0 : NaturalTransformation _ _ | T0 = ?T } ] =>
       nt_tac_abstract_trailing_props_with_equality_do tac T @NaturalTransformation_eq'
-    | [ |- { T0 : SpecializedNaturalTransformation _ _ | T0 == ?T } ] =>
+    | [ |- { T0 : NaturalTransformation _ _ | T0 == ?T } ] =>
       nt_tac_abstract_trailing_props_with_equality_do tac T @NaturalTransformation_JMeq
   end.
 Ltac nt_abstract_trailing_props_with_equality := nt_tac_abstract_trailing_props_with_equality ltac:(fun T' => T').
 Ltac nt_simpl_abstract_trailing_props_with_equality := nt_tac_abstract_trailing_props_with_equality ltac:(fun T' => let T'' := eval simpl in T' in T'').
 
 Section NaturalTransformationComposition.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
-  Variables F F' F'' : SpecializedFunctor C D.
-  Variables G G' : SpecializedFunctor D E.
+  Local Open Scope morphism_scope.
+
+  Variables C D E : Category.
+  Variables F F' F'' : Functor C D.
+  Variables G G' : Functor D E.
 
   Hint Resolve @Commutes f_equal f_equal2 : natural_transformation.
   Hint Rewrite @Associativity : natural_transformation.
@@ -269,13 +255,13 @@
 
   *)
 
-  Definition NTComposeT (T' : SpecializedNaturalTransformation F' F'') (T : SpecializedNaturalTransformation F F') :
-    SpecializedNaturalTransformation F F''.
-    exists (fun c => Compose (T' c) (T c));
+  Definition NTComposeT (T' : NaturalTransformation F' F'') (T : NaturalTransformation F F') :
+    NaturalTransformation F F''.
+    exists (fun c => T' c ∘ T c);
     (* XXX TODO: Find a way to get rid of [m] in the transitivity call *)
     abstract (
         intros;
-        transitivity (Compose (T' _) (Compose (MorphismOf F' m) (T _)));
+        transitivity (T' _ ∘ (MorphismOf F' m ∘ T _));
         try_associativity ltac:(eauto with natural_transformation)
       ).
   Defined.
@@ -305,21 +291,9 @@
   *)
   (* XXX TODO: Automate this better *)
 
-  Hint Rewrite @Commutes : natural_transformation.
-  Hint Resolve f_equal2 : natural_transformation.
-  Hint Extern 1 (_ = _) => apply @FCompositionOf : natural_transformation.
-
-  Lemma FCompositionOf2 : forall `(C : @SpecializedCategory objC) `(D : @SpecializedCategory objD)
-    (F : SpecializedFunctor C D) x y z u (m1 : C.(Morphism) x z) (m2 : C.(Morphism) y x) (m3 : D.(Morphism) u _),
-    Compose (MorphismOf F m1) (Compose (MorphismOf F m2) m3) = Compose (MorphismOf F (Compose m1 m2)) m3.
-    intros; symmetry; try_associativity ltac:(eauto with natural_transformation).
-  Qed.
-
-  Hint Rewrite @FCompositionOf2 : natural_transformation.
-
-  Definition NTComposeF (U : SpecializedNaturalTransformation G G') (T : SpecializedNaturalTransformation F F'):
-    SpecializedNaturalTransformation (ComposeFunctors G F) (ComposeFunctors G' F').
-    exists (fun c => Compose (G'.(MorphismOf) (T c)) (U (F c)));
+  Definition NTComposeF (U : NaturalTransformation G G') (T : NaturalTransformation F F')
+  : NaturalTransformation (G ∘ F) (G' ∘ F').
+    exists (fun c => G'.(MorphismOf) (T c) ∘ U (F c));
     abstract (
         simpl; intros; autorewrite with category;
         repeat try_associativity ltac:(repeat rewrite <- @Commutes; repeat rewrite <- @FCompositionOf);
@@ -329,22 +303,35 @@
 End NaturalTransformationComposition.
 
 Section IdentityNaturalTransformation.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
+  Variables C D : Category.
+  Print Commutes.
+  Definition IdentityNaturalTransformation'
+             (oo : C -> D)
+             (mo : forall s d, Morphism C s d -> Morphism D (oo s) (oo d))
+             ido ido'
+             fco fco'
+  : NaturalTransformation
+      (@Build_Functor C D oo mo ido fco)
+      (@Build_Functor C D oo mo ido' fco').
+  Proof.
+    exists (fun c => Identity (oo c)); simpl; clear.
+    abstract (intros; autorewrite with category; reflexivity).
+  Defined.
+
+  Variable F : Functor C D.
 
   (* There is an identity natrual transformation. *)
-  Definition IdentityNaturalTransformation : SpecializedNaturalTransformation F F.
+  Definition IdentityNaturalTransformation : NaturalTransformation F F.
     exists (fun c => Identity (F c));
     abstract (intros; autorewrite with morphism; reflexivity).
   Defined.
 
-  Lemma LeftIdentityNaturalTransformation (F' : SpecializedFunctor C D) (T : SpecializedNaturalTransformation F' F) :
+  Lemma LeftIdentityNaturalTransformation (F' : Functor C D) (T : NaturalTransformation F' F) :
     NTComposeT IdentityNaturalTransformation T = T.
     nt_eq; auto with morphism.
   Qed.
 
-  Lemma RightIdentityNaturalTransformation (F' : SpecializedFunctor C D) (T : SpecializedNaturalTransformation F F') :
+  Lemma RightIdentityNaturalTransformation (F' : Functor C D) (T : NaturalTransformation F F') :
     NTComposeT T IdentityNaturalTransformation = T.
     nt_eq; auto with morphism.
   Qed.
@@ -354,14 +341,12 @@
 Hint Rewrite @LeftIdentityNaturalTransformation @RightIdentityNaturalTransformation : natural_transformation.
 
 Section IdentityNaturalTransformationF.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
-  Variable G : SpecializedFunctor D E.
-  Variable F : SpecializedFunctor C D.
+  Variables C D E : Category.
+  Variable G : Functor D E.
+  Variable F : Functor C D.
 
   Lemma NTComposeFIdentityNaturalTransformation :
-    NTComposeF (IdentityNaturalTransformation G) (IdentityNaturalTransformation F) = IdentityNaturalTransformation (ComposeFunctors G F).
+    NTComposeF (IdentityNaturalTransformation G) (IdentityNaturalTransformation F) = IdentityNaturalTransformation (G ∘ F).
   Proof.
     nt_eq; repeat rewrite FIdentityOf; auto with morphism.
   Qed.
@@ -371,21 +356,20 @@
 Hint Rewrite @NTComposeFIdentityNaturalTransformation : natural_transformation.
 
 Section Associativity.
-  Context `(B : @SpecializedCategory objB).
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
-  Variable F : SpecializedFunctor D E.
-  Variable G : SpecializedFunctor C D.
-  Variable H : SpecializedFunctor B C.
+  Local Open Scope functor_scope.
 
-  Let F0 := ComposeFunctors (ComposeFunctors F G) H.
-  Let F1 := ComposeFunctors F (ComposeFunctors G H).
+  Variables B C D E : Category.
+  Variable F : Functor D E.
+  Variable G : Functor C D.
+  Variable H : Functor B C.
 
-  Definition ComposeFunctorsAssociator1 : SpecializedNaturalTransformation F0 F1.
-    refine (Build_SpecializedNaturalTransformation F0 F1
-                                                   (fun _ => Identity (C := E) _)
-                                                   _
+  Let F0 := (F ∘ G) ∘ H.
+  Let F1 := F ∘ (G ∘ H).
+
+  Definition ComposeFunctorsAssociator1 : NaturalTransformation F0 F1.
+    refine (Build_NaturalTransformation F0 F1
+                                        (fun _ => Identity (C := E) _)
+                                        _
            );
     abstract (
         simpl; intros;
@@ -393,10 +377,10 @@
       ).
   Defined.
 
-  Definition ComposeFunctorsAssociator2 : SpecializedNaturalTransformation F1 F0.
-    refine (Build_SpecializedNaturalTransformation F1 F0
-                                                   (fun _ => Identity (C := E) _)
-                                                   _
+  Definition ComposeFunctorsAssociator2 : NaturalTransformation F1 F0.
+    refine (Build_NaturalTransformation F1 F0
+                                        (fun _ => Identity (C := E) _)
+                                        _
            );
     abstract (
         simpl; intros;
@@ -406,24 +390,23 @@
 End Associativity.
 
 Section IdentityFunctor.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
   Local Ltac t :=
     repeat match goal with
-             | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-               refine (Build_SpecializedNaturalTransformation F G
-                                                              (fun _ => Identity _)
-                                                              _)
+             | [ |- NaturalTransformation ?F ?G ] =>
+               refine (Build_NaturalTransformation F G
+                                                   (fun _ => Identity _)
+                                                   _)
              | _ => abstract (simpl; intros; autorewrite with morphism; reflexivity)
              | _ => split; nt_eq
            end.
 
   Section left.
-    Variable F : SpecializedFunctor D C.
+    Variable F : Functor D C.
 
-    Definition LeftIdentityFunctorNaturalTransformation1 : SpecializedNaturalTransformation (ComposeFunctors (IdentityFunctor _) F) F. t. Defined.
-    Definition LeftIdentityFunctorNaturalTransformation2 : SpecializedNaturalTransformation F (ComposeFunctors (IdentityFunctor _) F). t. Defined.
+    Definition LeftIdentityFunctorNaturalTransformation1 : NaturalTransformation (IdentityFunctor _ ∘ F) F. t. Defined.
+    Definition LeftIdentityFunctorNaturalTransformation2 : NaturalTransformation F (IdentityFunctor _ ∘ F). t. Defined.
 
     Theorem LeftIdentityFunctorNT_Isomorphism
     : NTComposeT LeftIdentityFunctorNaturalTransformation1 LeftIdentityFunctorNaturalTransformation2 = IdentityNaturalTransformation _
@@ -433,10 +416,10 @@
   End left.
 
   Section right.
-    Variable F : SpecializedFunctor C D.
+    Variable F : Functor C D.
 
-    Definition RightIdentityFunctorNaturalTransformation1 : SpecializedNaturalTransformation (ComposeFunctors F (IdentityFunctor _)) F. t. Defined.
-    Definition RightIdentityFunctorNaturalTransformation2 : SpecializedNaturalTransformation F (ComposeFunctors F (IdentityFunctor _)). t. Defined.
+    Definition RightIdentityFunctorNaturalTransformation1 : NaturalTransformation (F ∘ (IdentityFunctor _)) F. t. Defined.
+    Definition RightIdentityFunctorNaturalTransformation2 : NaturalTransformation F (F ∘ (IdentityFunctor _)). t. Defined.
 
     Theorem RightIdentityFunctorNT_Isomorphism
     : NTComposeT RightIdentityFunctorNaturalTransformation1 RightIdentityFunctorNaturalTransformation2 = IdentityNaturalTransformation _
@@ -447,18 +430,16 @@
 End IdentityFunctor.
 
 Section NaturalTransformationExchangeLaw.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Context `(E : SpecializedCategory objE).
+  Variables C D E : Category.
 
-  Variables F G H : SpecializedFunctor C D.
-  Variables F' G' H' : SpecializedFunctor D E.
+  Variables F G H : Functor C D.
+  Variables F' G' H' : Functor D E.
 
-  Variable T : SpecializedNaturalTransformation F G.
-  Variable U : SpecializedNaturalTransformation G H.
+  Variable T : NaturalTransformation F G.
+  Variable U : NaturalTransformation G H.
 
-  Variable T' : SpecializedNaturalTransformation F' G'.
-  Variable U' : SpecializedNaturalTransformation G' H'.
+  Variable T' : NaturalTransformation F' G'.
+  Variable U' : NaturalTransformation G' H'.
 
   Local Ltac t_progress := progress repeat
     match goal with
@@ -492,9 +473,9 @@
   repeat match goal with
            | _ => exact (ComposeFunctorsAssociator1 _ _ _)
            | _ => exact (ComposeFunctorsAssociator2 _ _ _)
-           | [ |- SpecializedNaturalTransformation (ComposeFunctors ?F _) (ComposeFunctors ?F _) ] =>
+           | [ |- NaturalTransformation (ComposeFunctors ?F _) (ComposeFunctors ?F _) ] =>
              refine (NTComposeF (IdentityNaturalTransformation F) _)
-           | [ |- SpecializedNaturalTransformation (ComposeFunctors _ ?F) (ComposeFunctors _ ?F) ] =>
+           | [ |- NaturalTransformation (ComposeFunctors _ ?F) (ComposeFunctors _ ?F) ] =>
              refine (NTComposeF _ (IdentityNaturalTransformation F))
          end.
 Ltac nt_solve_associator :=
diff -r 4eb6407c6ca3 Notations.v
--- a/Notations.v	Fri May 24 20:09:37 2013 -0400
+++ b/Notations.v	Fri Jun 21 15:07:08 2013 -0400
@@ -20,12 +20,17 @@
 Reserved Notation "x ⊗ y" (at level 40, left associativity).
 Reserved Notation "x ⊗m y" (at level 40, left associativity).
 
-Reserved Notation "f ○ g" (at level 70, right associativity).
+Reserved Infix "○" (at level 40, left associativity).
+Reserved Infix "∘" (at level 40, left associativity).
+Reserved Infix "∘₀" (at level 40, left associativity).
+Reserved Infix "∘₁" (at level 40, left associativity).
 
 Reserved Notation "x ~> y" (at level 99, right associativity, y at level 200).
 
 Reserved Notation "x ∏ y" (at level 40, left associativity).
 Reserved Notation "x ∐ y" (at level 50, left associativity).
+Reserved Infix "Π" (at level 40, left associativity).
+Reserved Infix "⊔" (at level 50, left associativity).
 
 Reserved Notation "∏_{ x } f" (at level 0, x at level 99).
 Reserved Notation "∏_{ x : A } f" (at level 0, x at level 99).
@@ -43,6 +48,8 @@
 
 Reserved Notation "∫ F" (at level 0).
 
+Reserved Infix "\" (at level 40, left associativity).
+
 Delimit Scope object_scope with object.
 Delimit Scope morphism_scope with morphism.
 Delimit Scope category_scope with category.
diff -r 4eb6407c6ca3 PathsCategory.v
--- a/PathsCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/PathsCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Paths.
+Require Export Category Paths.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -14,8 +14,8 @@
   Hint Immediate concatenate_associative.
   Hint Rewrite concatenate_associative.
 
-  Definition PathsCategory : @SpecializedCategory V.
-    refine (@Build_SpecializedCategory _
+  Definition PathsCategory : @Category V.
+    refine (@Build_Category _
                                        (@path V E)
                                        (@NoEdges _ _)
                                        (fun _ _ _ p p' => concatenate p' p)
diff -r 4eb6407c6ca3 PathsCategoryFunctors.v
--- a/PathsCategoryFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/PathsCategoryFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export PathsCategory Functor SetCategory ComputableCategory SmallCat NaturalTransformation.
+Require Export PathsCategory Functor SetCategory ComputableCategory Cat NaturalTransformation.
 Require Import Common Adjoint.
 
 Set Implicit Arguments.
@@ -13,7 +13,7 @@
 Section FunctorFromPaths.
   Variable V : Type.
   Variable E : V -> V -> Type.
-  Context `(D : @SpecializedCategory objD).
+  Variable D : Category.
   Variable objOf : V -> objD.
   Variable morOf : forall s d, E s d -> Morphism D (objOf s) (objOf d).
 
@@ -29,7 +29,7 @@
     induction p2; t_with t'; autorewrite with morphism; reflexivity.
   Qed.
 
-  Definition FunctorFromPaths : SpecializedFunctor (PathsCategory E) D.
+  Definition FunctorFromPaths : Functor (PathsCategory E) D.
   Proof.
     refine {|
       ObjectOf := objOf;
@@ -41,5 +41,5 @@
 End FunctorFromPaths.
 
 Section Underlying.
-  Definition UnderlyingGraph `(C : @SpecializedCategory objC) := @PathsCategory objC (Morphism C).
+  Definition UnderlyingGraph `(C : @Category objC) := @PathsCategory objC (Morphism C).
 End Underlying.
diff -r 4eb6407c6ca3 ProductCategory.v
--- a/ProductCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/ProductCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -9,39 +9,39 @@
 
 Set Universe Polymorphism.
 
+Local Open Scope morphism_scope.
+
 Section ProductCategory.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
-  Definition ProductCategory : @SpecializedCategory (objC * objD)%type.
-    refine (@Build_SpecializedCategory _
-                                       (fun s d => (C.(Morphism) (fst s) (fst d) * D.(Morphism) (snd s) (snd d))%type)
-                                       (fun o => (Identity (fst o), Identity (snd o)))
-                                       (fun s d d' m2 m1 => (Compose (fst m2) (fst m1), Compose (snd m2) (snd m1)))
-                                       _
-                                       _
-                                       _);
-    abstract (intros; simpl_eq; auto with morphism).
+  Definition ProductCategory : Category.
+    refine (@Build_Category (C * D)%type
+                            (fun s d => (C.(Morphism) (fst s) (fst d) * D.(Morphism) (snd s) (snd d))%type)
+                            (fun o => (Identity (fst o), Identity (snd o)))
+                            (fun s d d' m2 m1 => (fst m2 ∘ fst m1, snd m2 ∘ snd m1))
+                            _
+                            _
+                            _);
+    abstract (intros; apply injective_projections; simpl; auto with morphism).
   Defined.
 End ProductCategory.
 
 Infix "*" := ProductCategory : category_scope.
 
 Section ProductCategoryFunctors.
-  Context `{C : @SpecializedCategory objC}.
-  Context `{D : @SpecializedCategory objD}.
+  Context {C D : Category}.
 
-  Definition fst_Functor : SpecializedFunctor (C * D) C
-    := Build_SpecializedFunctor (C * D) C
-                                (@fst _ _)
-                                (fun _ _ => @fst _ _)
-                                (fun _ _ _ _ _ => eq_refl)
-                                (fun _ => eq_refl).
+  Definition fst_Functor : Functor (C * D) C
+    := Build_Functor (C * D) C
+                     (@fst _ _)
+                     (fun _ _ => @fst _ _)
+                     (fun _ _ _ _ _ => eq_refl)
+                     (fun _ => eq_refl).
 
-  Definition snd_Functor : SpecializedFunctor (C * D) D
-    := Build_SpecializedFunctor (C * D) D
-                                (@snd _ _)
-                                (fun _ _ => @snd _ _)
-                                (fun _ _ _ _ _ => eq_refl)
-                                (fun _ => eq_refl).
+  Definition snd_Functor : Functor (C * D) D
+    := Build_Functor (C * D) D
+                     (@snd _ _)
+                     (fun _ _ => @snd _ _)
+                     (fun _ _ _ _ _ => eq_refl)
+                     (fun _ => eq_refl).
 End ProductCategoryFunctors.
diff -r 4eb6407c6ca3 ProductFunctors.v
--- a/ProductFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/ProductFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,11 +10,11 @@
 Set Universe Polymorphism.
 
 Section Products.
-  Context `{C : @SpecializedCategory objC}.
+  Context {C : Category}.
   Variable I : Type.
 
-  Variable HasLimits : forall F : SpecializedFunctor (DiscreteCategory I) C, Limit F.
-  Variable HasColimits : forall F : SpecializedFunctor (DiscreteCategory I) C, Colimit F.
+  Variable HasLimits : forall F : Functor (DiscreteCategory I) C, Limit F.
+  Variable HasColimits : forall F : Functor (DiscreteCategory I) C, Colimit F.
 
   Definition ProductFunctor := LimitFunctor HasLimits.
   Definition CoproductFunctor := ColimitFunctor HasColimits.
diff -r 4eb6407c6ca3 ProductInducedFunctors.v
--- a/ProductInducedFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/ProductInducedFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -15,20 +15,20 @@
   reflexivity.
 
 Section ProductInducedFunctors.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
+  Variable C : Category.
+  Variable D : Category.
+  Variable E : Category.
 
-  Variable F : SpecializedFunctor (C * D) E.
+  Variable F : Functor (C * D) E.
 
-  Definition InducedProductFstFunctor (d : D) : SpecializedFunctor C E.
+  Definition InducedProductFstFunctor (d : D) : Functor C E.
     refine {| ObjectOf := (fun c => F (c, d));
       MorphismOf := (fun _ _ m => MorphismOf F (s := (_, d)) (d := (_, d)) (m, Identity d))
     |};
     abstract t.
   Defined.
 
-  Definition InducedProductSndFunctor (c : C) : SpecializedFunctor D E.
+  Definition InducedProductSndFunctor (c : C) : Functor D E.
     refine {| ObjectOf := (fun d => F (c, d));
       MorphismOf := (fun _ _ m => MorphismOf F (s := (c, _)) (d := (c, _)) (Identity c, m))
     |};
@@ -40,16 +40,16 @@
 Notation "F ⟨ - , d ⟩" := (InducedProductFstFunctor F d) : functor_scope.
 
 Section ProductInducedNaturalTransformations.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
-  Context `(E : @SpecializedCategory objE).
+  Variable C : Category.
+  Variable D : Category.
+  Variable E : Category.
 
-  Variable F : SpecializedFunctor (C * D) E.
+  Variable F : Functor (C * D) E.
 
-  Definition InducedProductFstNaturalTransformation {s d} (m : Morphism C s d) : SpecializedNaturalTransformation (F ⟨ s, - ⟩) (F ⟨ d, - ⟩).
+  Definition InducedProductFstNaturalTransformation {s d} (m : Morphism C s d) : NaturalTransformation (F ⟨ s, - ⟩) (F ⟨ d, - ⟩).
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F0 ?G0 ] =>
-        refine (Build_SpecializedNaturalTransformation F0 G0
+      | [ |- NaturalTransformation ?F0 ?G0 ] =>
+        refine (Build_NaturalTransformation F0 G0
           (fun d => MorphismOf F (s := (_, d)) (d := (_, d)) (m, Identity (C := D) d))
           _
         )
@@ -57,10 +57,10 @@
     abstract t.
   Defined.
 
-  Definition InducedProductSndNaturalTransformation {s d} (m : Morphism D s d) : SpecializedNaturalTransformation (F ⟨ -, s ⟩) (F ⟨ - , d ⟩).
+  Definition InducedProductSndNaturalTransformation {s d} (m : Morphism D s d) : NaturalTransformation (F ⟨ -, s ⟩) (F ⟨ - , d ⟩).
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F0 ?G0 ] =>
-        refine (Build_SpecializedNaturalTransformation F0 G0
+      | [ |- NaturalTransformation ?F0 ?G0 ] =>
+        refine (Build_NaturalTransformation F0 G0
           (fun c => MorphismOf F (s := (c, _)) (d := (c, _)) (Identity (C := C) c, m))
           _
         )
diff -r 4eb6407c6ca3 ProductLaws.v
--- a/ProductLaws.v	Fri May 24 20:09:37 2013 -0400
+++ b/ProductLaws.v	Fri Jun 21 15:07:08 2013 -0400
@@ -12,25 +12,25 @@
 Local Open Scope category_scope.
 
 Section swap.
-  Definition SwapFunctor `(C : @SpecializedCategory objC) `(D : @SpecializedCategory objD)
-  : SpecializedFunctor (C * D) (D * C)
-    := Build_SpecializedFunctor (C * D) (D * C)
+  Definition SwapFunctor `(C : @Category objC) `(D : @Category objD)
+  : Functor (C * D) (D * C)
+    := Build_Functor (C * D) (D * C)
                                 (fun cd => (snd cd, fst cd))
                                 (fun _ _ m => (snd m, fst m))
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Lemma ProductLawSwap `(C : @SpecializedCategory objC) `(D : @SpecializedCategory objD)
+  Lemma ProductLawSwap `(C : @Category objC) `(D : @Category objD)
   : ComposeFunctors (SwapFunctor C D) (SwapFunctor D C) = IdentityFunctor _.
     functor_eq; intuition.
   Qed.
 End swap.
 
 Section Law0.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition ProductLaw0Functor : SpecializedFunctor (C * 0) 0.
-    refine (Build_SpecializedFunctor (C * 0) 0
+  Definition ProductLaw0Functor : Functor (C * 0) 0.
+    refine (Build_Functor (C * 0) 0
                                      (@snd _ _)
                                      (fun _ _ => @snd _ _)
                                      _
@@ -42,7 +42,7 @@
       ).
   Defined.
 
-  Definition ProductLaw0Functor_Inverse : SpecializedFunctor 0 (C * 0).
+  Definition ProductLaw0Functor_Inverse : Functor 0 (C * 0).
     repeat esplit;
     intros; destruct_head_hnf Empty_set.
     Grab Existential Variables.
@@ -60,17 +60,17 @@
 End Law0.
 
 Section Law0'.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Let ProductLaw0'Functor' : SpecializedFunctor (0 * C) 0.
+  Let ProductLaw0'Functor' : Functor (0 * C) 0.
     functor_simpl_abstract_trailing_props (ComposeFunctors (ProductLaw0Functor C) (SwapFunctor _ _)).
   Defined.
-  Definition ProductLaw0'Functor : SpecializedFunctor (0 * C) 0 := Eval hnf in ProductLaw0'Functor'.
+  Definition ProductLaw0'Functor : Functor (0 * C) 0 := Eval hnf in ProductLaw0'Functor'.
 
-  Let ProductLaw0'Functor_Inverse' : SpecializedFunctor 0 (0 * C).
+  Let ProductLaw0'Functor_Inverse' : Functor 0 (0 * C).
     functor_simpl_abstract_trailing_props (ComposeFunctors (SwapFunctor _ _) (ProductLaw0Functor_Inverse C)).
   Defined.
-  Definition ProductLaw0'Functor_Inverse : SpecializedFunctor 0 (0 * C) := Eval hnf in ProductLaw0'Functor_Inverse'.
+  Definition ProductLaw0'Functor_Inverse : Functor 0 (0 * C) := Eval hnf in ProductLaw0'Functor_Inverse'.
 
   Lemma ProductLaw0' : ComposeFunctors ProductLaw0'Functor ProductLaw0'Functor_Inverse = IdentityFunctor _ /\
     ComposeFunctors ProductLaw0'Functor_Inverse ProductLaw0'Functor = IdentityFunctor _.
@@ -82,16 +82,16 @@
 End Law0'.
 
 Section Law1.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Let ProductLaw1Functor' : SpecializedFunctor (C * 1) C.
+  Let ProductLaw1Functor' : Functor (C * 1) C.
     functor_simpl_abstract_trailing_props (fst_Functor (C := C) (D := 1)).
   Defined.
-  Definition ProductLaw1Functor : SpecializedFunctor (C * 1) C
+  Definition ProductLaw1Functor : Functor (C * 1) C
     := Eval hnf in ProductLaw1Functor'.
 
-  Definition ProductLaw1Functor_Inverse : SpecializedFunctor C (C * 1).
-    refine (Build_SpecializedFunctor C (C * 1)
+  Definition ProductLaw1Functor_Inverse : Functor C (C * 1).
+    refine (Build_Functor C (C * 1)
                                      (fun c => (c, tt))
                                      (fun s d m => (m, eq_refl))
                                      _
@@ -113,17 +113,17 @@
 End Law1.
 
 Section Law1'.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition ProductLaw1'Functor' : SpecializedFunctor (1 * C) C.
+  Definition ProductLaw1'Functor' : Functor (1 * C) C.
     functor_simpl_abstract_trailing_props (ComposeFunctors (ProductLaw1Functor C) (SwapFunctor _ _)).
   Defined.
-  Definition ProductLaw1'Functor : SpecializedFunctor (1 * C) C := Eval hnf in ProductLaw1'Functor'.
+  Definition ProductLaw1'Functor : Functor (1 * C) C := Eval hnf in ProductLaw1'Functor'.
 
-  Let ProductLaw1'Functor_Inverse' : SpecializedFunctor C (1 * C).
+  Let ProductLaw1'Functor_Inverse' : Functor C (1 * C).
     functor_simpl_abstract_trailing_props (ComposeFunctors (SwapFunctor _ _) (ProductLaw1Functor_Inverse C)).
   Defined.
-  Definition ProductLaw1'Functor_Inverse : SpecializedFunctor C (1 * C) := Eval hnf in ProductLaw1'Functor_Inverse'.
+  Definition ProductLaw1'Functor_Inverse : Functor C (1 * C) := Eval hnf in ProductLaw1'Functor_Inverse'.
 
   Lemma ProductLaw1' : ComposeFunctors ProductLaw1'Functor ProductLaw1'Functor_Inverse = IdentityFunctor _ /\
     ComposeFunctors ProductLaw1'Functor_Inverse ProductLaw1'Functor = IdentityFunctor _.
diff -r 4eb6407c6ca3 ProductNaturalTransformation.v
--- a/ProductNaturalTransformation.v	Fri May 24 20:09:37 2013 -0400
+++ b/ProductNaturalTransformation.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,17 +10,17 @@
 Set Universe Polymorphism.
 
 Section ProductNaturalTransformation.
-  Context `{A : @SpecializedCategory objA}.
-  Context `{B : @SpecializedCategory objB}.
-  Context `{C : @SpecializedCategory objC}.
-  Context `{D : @SpecializedCategory objD}.
-  Variables F F' : SpecializedFunctor A B.
-  Variables G G' : SpecializedFunctor C D.
-  Variable T : SpecializedNaturalTransformation F F'.
-  Variable U : SpecializedNaturalTransformation G G'.
+  Context {A : Category}.
+  Context {B : Category}.
+  Context {C : Category}.
+  Context {D : Category}.
+  Variables F F' : Functor A B.
+  Variables G G' : Functor C D.
+  Variable T : NaturalTransformation F F'.
+  Variable U : NaturalTransformation G G'.
 
-  Definition ProductNaturalTransformation : SpecializedNaturalTransformation (F * G) (F' * G').
-    refine (Build_SpecializedNaturalTransformation (F * G) (F' * G')
+  Definition ProductNaturalTransformation : NaturalTransformation (F * G) (F' * G').
+    refine (Build_NaturalTransformation (F * G) (F' * G')
       (fun ac : A * C => (T (fst ac), U (snd ac)))
       _
     );
diff -r 4eb6407c6ca3 Products.v
--- a/Products.v	Fri May 24 20:09:37 2013 -0400
+++ b/Products.v	Fri Jun 21 15:07:08 2013 -0400
@@ -10,7 +10,7 @@
 Set Universe Polymorphism.
 
 Section Products.
-  Context `{C : @SpecializedCategory objC}.
+  Context {C : Category}.
   Variable I : Type.
   Variable f : I -> C.
 
diff -r 4eb6407c6ca3 Pullback.v
--- a/Pullback.v	Fri May 24 20:09:37 2013 -0400
+++ b/Pullback.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Export Limits.
-Require Import Common Duals SpecializedCommaCategory FunctorCategory.
+Require Import Common Duals CommaCategory FunctorCategory.
 
 Set Implicit Arguments.
 
@@ -24,7 +24,7 @@
      ]]
    *)
 
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Section pullback.
     Variables a b c : objC.
     Variable f : C.(Morphism) a c.
@@ -49,8 +49,8 @@
              destruct s, d, d'; simpl in *; trivial.
            Defined.
 
-    Definition PullbackIndex : @SpecializedCategory PullbackThree.
-      refine (@Build_SpecializedCategory _
+    Definition PullbackIndex : @Category PullbackThree.
+      refine (@Build_Category _
                                          PullbackIndex_Morphism
                                          (fun x => match x as p return (PullbackIndex_Morphism p p) with
                                                        PullbackA => tt | PullbackB => tt | PullbackC => tt
@@ -81,10 +81,10 @@
       try solve [ destruct m ].
     Defined.
 
-    Definition PullbackDiagram : SpecializedFunctor PullbackIndex C.
+    Definition PullbackDiagram : Functor PullbackIndex C.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            PullbackDiagram_ObjectOf
                                            PullbackDiagram_MorphismOf
                                            _
@@ -104,7 +104,7 @@
   End pullback.
 
   Section pullback_functorial.
-    Local Infix "/" := SliceSpecializedCategoryOver.
+    Local Infix "/" := SliceCategoryOver.
 
     Variable c : C.
 
@@ -132,10 +132,10 @@
         ).
     Defined.
 
-    Definition PullbackDiagramFunctor : SpecializedFunctor ((C / c) * (C / c)) (C ^ PullbackIndex).
+    Definition PullbackDiagramFunctor : Functor ((C / c) * (C / c)) (C ^ PullbackIndex).
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            PullbackDiagramFunctor_ObjectOf
                                            PullbackDiagramFunctor_MorphismOf
                                            _
@@ -171,10 +171,10 @@
       try solve [ destruct m ].
     Defined.
 
-    Definition PushoutDiagram : SpecializedFunctor PushoutIndex C.
+    Definition PushoutDiagram : Functor PushoutIndex C.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            PushoutDiagram_ObjectOf
                                            PushoutDiagram_MorphismOf
                                            _
@@ -220,10 +220,10 @@
         ).
     Defined.
 
-    Definition PushoutDiagramFunctor : SpecializedFunctor ((c \ C) * (c \ C)) (C ^ PushoutIndex).
+    Definition PushoutDiagramFunctor : Functor ((c \ C) * (c \ C)) (C ^ PushoutIndex).
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            PushoutDiagramFunctor_ObjectOf
                                            PushoutDiagramFunctor_MorphismOf
                                            _
@@ -242,7 +242,7 @@
 End Pullback.
 
 Section PullbackObjects.
-  Context `{C : @SpecializedCategory objC}.
+  Context {C : Category}.
   Variables a b c : objC.
 
   (** Does an object [d] together with the functions [i] and [j]
diff -r 4eb6407c6ca3 PullbackFunctor.v
--- a/PullbackFunctor.v	Fri May 24 20:09:37 2013 -0400
+++ b/PullbackFunctor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -9,10 +9,10 @@
 Set Universe Polymorphism.
 
 Section Equalizer.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Variable HasLimits : forall F : SpecializedFunctor PullbackIndex C, Limit F.
-  Variable HasColimits : forall F : SpecializedFunctor PushoutIndex C, Colimit F.
+  Variable HasLimits : forall F : Functor PullbackIndex C, Limit F.
+  Variable HasColimits : forall F : Functor PushoutIndex C, Colimit F.
 
   Definition PullbackFunctor := LimitFunctor HasLimits.
   Definition PushoutFunctor := ColimitFunctor HasColimits.
diff -r 4eb6407c6ca3 SemiSimplicialSets.v
--- a/SemiSimplicialSets.v	Fri May 24 20:09:37 2013 -0400
+++ b/SemiSimplicialSets.v	Fri Jun 21 15:07:08 2013 -0400
@@ -21,17 +21,17 @@
      which we will use to make simplicial sets without having to worry
      about "degneracies". *)
 
-  Definition SemiSimplexCategory : SpecializedCategory nat.
+  Definition SemiSimplexCategory : Category nat.
     eapply (WideSubcategory Δ (@IsMonomorphism _ _));
     abstract eauto with morphism.
   Defined.
 
   Local Notation Γ := SemiSimplexCategory.
 
-  Definition SemiSimplexCategoryInclusionFunctor : SpecializedFunctor Γ Δ
+  Definition SemiSimplexCategoryInclusionFunctor : Functor Γ Δ
     := WideSubcategoryInclusionFunctor _ _ _ _.
 
-  Definition SemiSimplicialCategory `(C : SpecializedCategory objC) := (C ^ (OppositeCategory Γ))%category.
+  Definition SemiSimplicialCategory `(C : Category objC) := (C ^ (OppositeCategory Γ))%category.
 
   Definition SemiSimplicialSet := SemiSimplicialCategory SetCat.
   Definition SemiSimplicialType := SemiSimplicialCategory TypeCat.
diff -r 4eb6407c6ca3 SetCategory.v
--- a/SetCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SetCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -9,13 +9,13 @@
 
 Set Universe Polymorphism.
 
-Notation IndexedCatOf obj coerce := (@Build_SpecializedCategory obj
-                                                                (fun s d => coerce s -> coerce d)
-                                                                (fun _ => (fun x => x))
-                                                                (fun _ _ _ f g => (fun x => f (g x)))
-                                                                (fun _ _ _ _ _ _ f => eq_refl)
-                                                                (fun _ _ f => eq_refl : (fun x => f x) = f)
-                                                                (fun _ _ f => eq_refl : (fun x => f x) = f)
+Notation IndexedCatOf obj coerce := (@Build_Category obj
+                                                     (fun s d => coerce s -> coerce d)
+                                                     (fun _ => (fun x => x))
+                                                     (fun _ _ _ f g => (fun x => f (g x)))
+                                                     (fun _ _ _ _ _ _ f => eq_refl)
+                                                     (fun _ _ f => eq_refl : (fun x => f x) = f)
+                                                     (fun _ _ f => eq_refl : (fun x => f x) = f)
                                     ).
 
 Notation CatOf obj := (IndexedCatOf obj (fun x => x)).
@@ -23,81 +23,81 @@
 
 (* There is a category [Set], where the objects are sets and the morphisms are set morphisms *)
 Section CSet.
-  Definition TypeCat : @SpecializedCategory Type := Eval cbv beta in CatOf Type.
-  Definition SetCat : @SpecializedCategory Set := Eval cbv beta in CatOf Set.
-  Definition PropCat : @SpecializedCategory Prop := Eval cbv beta in CatOf Prop.
+  Definition TypeCat : Category := Eval cbv beta in CatOf Type.
+  Definition SetCat : Category := Eval cbv beta in CatOf Set.
+  Definition PropCat : Category := Eval cbv beta in CatOf Prop.
 
   Definition IndexedTypeCat (Index : Type) (Index2Object : Index -> Type) := Eval cbv beta in IndexedCatOf Index Index2Object.
 End CSet.
 
 Section SetCoercionsDefinitions.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Definition SpecializedFunctorToProp := SpecializedFunctor C PropCat.
-  Definition SpecializedFunctorToSet := SpecializedFunctor C SetCat.
-  Definition SpecializedFunctorToType := SpecializedFunctor C TypeCat.
+  Definition FunctorToProp := Functor C PropCat.
+  Definition FunctorToSet := Functor C SetCat.
+  Definition FunctorToType := Functor C TypeCat.
 
-  Definition SpecializedFunctorFromProp := SpecializedFunctor PropCat C.
-  Definition SpecializedFunctorFromSet := SpecializedFunctor SetCat C.
-  Definition SpecializedFunctorFromType := SpecializedFunctor TypeCat C.
+  Definition FunctorFromProp := Functor PropCat C.
+  Definition FunctorFromSet := Functor SetCat C.
+  Definition FunctorFromType := Functor TypeCat C.
 End SetCoercionsDefinitions.
 
-Identity Coercion SpecializedFunctorToProp_Id : SpecializedFunctorToProp >-> SpecializedFunctor.
-Identity Coercion SpecializedFunctorToSet_Id : SpecializedFunctorToSet >-> SpecializedFunctor.
-Identity Coercion SpecializedFunctorToType_Id : SpecializedFunctorToType >-> SpecializedFunctor.
-Identity Coercion SpecializedFunctorFromProp_Id : SpecializedFunctorFromProp >-> SpecializedFunctor.
-Identity Coercion SpecializedFunctorFromSet_Id : SpecializedFunctorFromSet >-> SpecializedFunctor.
-Identity Coercion SpecializedFunctorFromType_Id : SpecializedFunctorFromType >-> SpecializedFunctor.
+Identity Coercion FunctorToProp_Id : FunctorToProp >-> Functor.
+Identity Coercion FunctorToSet_Id : FunctorToSet >-> Functor.
+Identity Coercion FunctorToType_Id : FunctorToType >-> Functor.
+Identity Coercion FunctorFromProp_Id : FunctorFromProp >-> Functor.
+Identity Coercion FunctorFromSet_Id : FunctorFromSet >-> Functor.
+Identity Coercion FunctorFromType_Id : FunctorFromType >-> Functor.
 
 Section SetCoercions.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Local Notation BuildFunctor C D F :=
-    (Build_SpecializedFunctor C D
-                              (fun x => ObjectOf F%functor x)
-                              (fun s d m => MorphismOf F%functor m)
-                              (fun s d d' m m' => FCompositionOf F%functor s d d' m m')
-                              (fun x => FIdentityOf F%functor x)).
+    (Build_Functor C D
+                   (fun x => ObjectOf F%functor x)
+                   (fun s d m => MorphismOf F%functor m)
+                   (fun s d d' m m' => FCompositionOf F%functor s d d' m m')
+                   (fun x => FIdentityOf F%functor x)).
 
-  Definition SpecializedFunctorTo_Prop2Set (F : SpecializedFunctorToProp C) : SpecializedFunctorToSet C := BuildFunctor C SetCat F.
-  Definition SpecializedFunctorTo_Prop2Type (F : SpecializedFunctorToProp C) : SpecializedFunctorToType C := BuildFunctor C TypeCat F.
-  Definition SpecializedFunctorTo_Set2Type (F : SpecializedFunctorToSet C) : SpecializedFunctorToType C := BuildFunctor C TypeCat F.
+  Definition FunctorTo_Prop2Set (F : FunctorToProp C) : FunctorToSet C := BuildFunctor C SetCat F.
+  Definition FunctorTo_Prop2Type (F : FunctorToProp C) : FunctorToType C := BuildFunctor C TypeCat F.
+  Definition FunctorTo_Set2Type (F : FunctorToSet C) : FunctorToType C := BuildFunctor C TypeCat F.
 
-  Definition SpecializedFunctorFrom_Set2Prop (F : SpecializedFunctorFromSet C) : SpecializedFunctorFromProp C := BuildFunctor PropCat C F.
-  Definition SpecializedFunctorFrom_Type2Prop (F : SpecializedFunctorFromType C) : SpecializedFunctorFromProp C := BuildFunctor PropCat C F.
-  Definition SpecializedFunctorFrom_Type2Set (F : SpecializedFunctorFromType C) : SpecializedFunctorFromSet C := BuildFunctor SetCat C F.
+  Definition FunctorFrom_Set2Prop (F : FunctorFromSet C) : FunctorFromProp C := BuildFunctor PropCat C F.
+  Definition FunctorFrom_Type2Prop (F : FunctorFromType C) : FunctorFromProp C := BuildFunctor PropCat C F.
+  Definition FunctorFrom_Type2Set (F : FunctorFromType C) : FunctorFromSet C := BuildFunctor SetCat C F.
 End SetCoercions.
 
-Coercion SpecializedFunctorTo_Prop2Set : SpecializedFunctorToProp >-> SpecializedFunctorToSet.
-Coercion SpecializedFunctorTo_Prop2Type : SpecializedFunctorToProp >-> SpecializedFunctorToType.
-Coercion SpecializedFunctorTo_Set2Type : SpecializedFunctorToSet >-> SpecializedFunctorToType.
-Coercion SpecializedFunctorFrom_Set2Prop : SpecializedFunctorFromSet >-> SpecializedFunctorFromProp.
-Coercion SpecializedFunctorFrom_Type2Prop : SpecializedFunctorFromType >-> SpecializedFunctorFromProp.
-Coercion SpecializedFunctorFrom_Type2Set : SpecializedFunctorFromType >-> SpecializedFunctorFromSet.
+Coercion FunctorTo_Prop2Set : FunctorToProp >-> FunctorToSet.
+Coercion FunctorTo_Prop2Type : FunctorToProp >-> FunctorToType.
+Coercion FunctorTo_Set2Type : FunctorToSet >-> FunctorToType.
+Coercion FunctorFrom_Set2Prop : FunctorFromSet >-> FunctorFromProp.
+Coercion FunctorFrom_Type2Prop : FunctorFromType >-> FunctorFromProp.
+Coercion FunctorFrom_Type2Set : FunctorFromType >-> FunctorFromSet.
 
 (* There is a category [Set*], where the objects are sets with a distinguished elements and the morphisms are set morphisms *)
 Section PointedSet.
   Local Notation PointedCatOf obj := (let pointed_set_fun_eta := ((fun _ _ _ _ _ => eq_refl) :
                                                                     forall {T : Type} {proj : T -> Type} (a b : T) (f : proj a -> proj b),
                                                                       (fun x => f x) = f) in
-                                      @Build_SpecializedCategory { A : obj & A }
-                                                                 (fun s d => (projT1 s) -> (projT1 d))
-                                                                 (fun _ => (fun x => x))
-                                                                 (fun _ _ _ f g => (fun x => f (g x)))
-                                                                 (fun _ _ _ _ _ _ _ => eq_refl)
-                                                                 (@pointed_set_fun_eta { A : obj & A } (@projT1 obj _))
-                                                                 (@pointed_set_fun_eta { A : obj & A } (@projT1 obj _))).
+                                      @Build_Category { A : obj & A }
+                                                      (fun s d => (projT1 s) -> (projT1 d))
+                                                      (fun _ => (fun x => x))
+                                                      (fun _ _ _ f g => (fun x => f (g x)))
+                                                      (fun _ _ _ _ _ _ _ => eq_refl)
+                                                      (@pointed_set_fun_eta { A : obj & A } (@projT1 obj _))
+                                                      (@pointed_set_fun_eta { A : obj & A } (@projT1 obj _))).
 
-  Local Notation PointedCatProjectionOf obj := (Build_SpecializedFunctor (PointedCatOf obj) (CatOf obj)
-                                                                         (fun c => projT1 c)
-                                                                         (fun s d (m : (projT1 s) -> (projT1 d)) => m)
-                                                                         (fun _ _ _ _ _ => eq_refl)
-                                                                         (fun _ => eq_refl)).
+  Local Notation PointedCatProjectionOf obj := (Build_Functor (PointedCatOf obj) (CatOf obj)
+                                                              (fun c => projT1 c)
+                                                              (fun s d (m : (projT1 s) -> (projT1 d)) => m)
+                                                              (fun _ _ _ _ _ => eq_refl)
+                                                              (fun _ => eq_refl)).
 
-  Definition PointedTypeCat : @SpecializedCategory { A : Type & A } := Eval cbv beta zeta in PointedCatOf Type.
-  Definition PointedTypeProjection : SpecializedFunctor PointedTypeCat TypeCat := PointedCatProjectionOf Type.
-  Definition PointedSetCat : @SpecializedCategory { A : Set & A } := Eval cbv beta zeta in PointedCatOf Set.
-  Definition PointedSetProjection : SpecializedFunctor PointedSetCat SetCat := PointedCatProjectionOf Set.
-  Definition PointedPropCat : @SpecializedCategory { A : Prop & A } := Eval cbv beta zeta in PointedCatOf Prop.
-  Definition PointedPropProjection : SpecializedFunctor PointedPropCat PropCat := PointedCatProjectionOf Prop.
+  Definition PointedTypeCat : Category := Eval cbv beta zeta in PointedCatOf Type.
+  Definition PointedTypeProjection : Functor PointedTypeCat TypeCat := PointedCatProjectionOf Type.
+  Definition PointedSetCat : Category := Eval cbv beta zeta in PointedCatOf Set.
+  Definition PointedSetProjection : Functor PointedSetCat SetCat := PointedCatProjectionOf Set.
+  Definition PointedPropCat : Category := Eval cbv beta zeta in PointedCatOf Prop.
+  Definition PointedPropProjection : Functor PointedPropCat PropCat := PointedCatProjectionOf Prop.
 End PointedSet.
diff -r 4eb6407c6ca3 SetCategoryProductFunctor.v
--- a/SetCategoryProductFunctor.v	Fri May 24 20:09:37 2013 -0400
+++ b/SetCategoryProductFunctor.v	Fri Jun 21 15:07:08 2013 -0400
@@ -17,8 +17,8 @@
   Local Ltac build_functor :=
     hnf;
     match goal with
-      | [ |- @SpecializedFunctor ?objC ?C ?objD ?D ] =>
-        refine (@Build_SpecializedFunctor objC C objD D
+      | [ |- @Functor ?objC ?C ?objD ?D ] =>
+        refine (@Build_Functor objC C objD D
                                           (fun ab => (fst ab) * (snd ab))
                                           (fun _ _ fg => (fun xy => ((fst fg) (fst xy), (snd fg) (snd xy))))
                                           _
@@ -26,7 +26,7 @@
           abstract eauto
     end.
 
-  Definition TypeProductFunctor : SpecializedFunctor (TypeCat * TypeCat) TypeCat. build_functor. Defined.
-  Definition SetProductFunctor  : SpecializedFunctor (SetCat * SetCat) SetCat. build_functor. Defined.
-  Definition PropProductFunctor : SpecializedFunctor (PropCat * PropCat) PropCat. build_functor. Defined.
+  Definition TypeProductFunctor : Functor (TypeCat * TypeCat) TypeCat. build_functor. Defined.
+  Definition SetProductFunctor  : Functor (SetCat * SetCat) SetCat. build_functor. Defined.
+  Definition PropProductFunctor : Functor (PropCat * PropCat) PropCat. build_functor. Defined.
 End ProductFunctor.
diff -r 4eb6407c6ca3 SetColimits.v
--- a/SetColimits.v	Fri May 24 20:09:37 2013 -0400
+++ b/SetColimits.v	Fri Jun 21 15:07:08 2013 -0400
@@ -46,9 +46,9 @@
   (* An element of the colimit is an pair [c : C] and [x : F c] with
      (c, x) ~ (c', x') an equivalence relation generated by the
      existance of a morphism [m : c -> c'] with m c x = m c x'. *)
-  Context `(C : @SmallSpecializedCategory objC).
-  Variable F : SpecializedFunctor C SetCat.
-  Let F' : SpecializedFunctorToType _ := F : SpecializedFunctorToSet _.
+  Context `(C : @Category objC).
+  Variable F : Functor C SetCat.
+  Let F' : FunctorToType _ := F : FunctorToSet _.
 
   Definition SetColimit_Object_pre := SetGrothendieckPair F. (* { c : objC & F.(ObjectOf) c }.*)
   Global Arguments SetColimit_Object_pre /.
@@ -97,8 +97,8 @@
     Definition SetColimit_Morphism : Morphism (SetCat ^ C) F ((DiagonalFunctor SetCat C) SetColimit_Object).
       hnf; simpl.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun c : objC =>
               (fun S : F c =>
                 chooser' (Build_SetGrothendieckPair F c S)
@@ -172,8 +172,8 @@
   (* An element of the colimit is an pair [c : C] and [x : F c] with
      (c, x) ~ (c', x') an equivalence relation generated by the
      existance of a morphism [m : c -> c'] with m c x = m c x'. *)
-  Context `(C : @SpecializedCategory objC).
-  Variable F : SpecializedFunctor C TypeCat.
+  Variable C : Category.
+  Variable F : Functor C TypeCat.
 
   Definition TypeColimit_Object_pre := GrothendieckPair F. (* { c : objC & F.(ObjectOf) c }. *)
   Global Arguments TypeColimit_Object_pre /.
@@ -222,8 +222,8 @@
     Definition TypeColimit_Morphism : Morphism (TypeCat ^ C) F ((DiagonalFunctor TypeCat C) TypeColimit_Object).
       hnf; simpl.
       match goal with
-        | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-          refine (Build_SpecializedNaturalTransformation F G
+        | [ |- NaturalTransformation ?F ?G ] =>
+          refine (Build_NaturalTransformation F G
             (fun c : objC =>
               (fun S : F c =>
                 chooser' (Build_GrothendieckPair F c S)
diff -r 4eb6407c6ca3 SetLimits.v
--- a/SetLimits.v	Fri May 24 20:09:37 2013 -0400
+++ b/SetLimits.v	Fri Jun 21 15:07:08 2013 -0400
@@ -19,7 +19,7 @@
 
   repeat (apply functional_extensionality_dep; intro; try simpl_eq);
 
-  destruct_head @SpecializedNaturalTransformation;
+  destruct_head @NaturalTransformation;
   fg_equal;
 
   destruct_sig;
@@ -27,8 +27,8 @@
   trivial; t_with t'; intuition.
 
 Section SetLimits.
-  Context `(C : @SmallSpecializedCategory objC).
-  Variable F : SpecializedFunctor C SetCat.
+  Context `(C : @Category objC).
+  Variable F : Functor C SetCat.
 
   (* Quoting David:
      let F:C-->Set be a functor. An element of the limit is a collection of elements x_c,
@@ -37,12 +37,12 @@
   Definition SetLimit_Object : SetCat :=
     { S : forall c : objC, F c | forall c c' (g : C.(Morphism) c c'), F.(MorphismOf) g (S c) = (S c') }.
 
-  Definition SetLimit_Morphism : SpecializedNaturalTransformation
+  Definition SetLimit_Morphism : NaturalTransformation
                                    ((DiagonalFunctor SetCat C) SetLimit_Object)
                                    F.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun c : objC => (fun S => (proj1_sig S) c))
           _
         )
@@ -51,7 +51,7 @@
   Defined.
 
   Definition SetLimit_Property_Morphism A'
-             (φ' : SpecializedNaturalTransformation ((DiagonalFunctor SetCat C) A') F) :
+             (φ' : NaturalTransformation ((DiagonalFunctor SetCat C) A') F) :
     A' -> SetLimit_Object.
     intro x; hnf.
     exists (fun c => ComponentsOf φ' c x).
@@ -67,8 +67,8 @@
 End SetLimits.
 
 Section TypeLimits.
-  Context `(C : @SmallSpecializedCategory objC).
-  Variable F : SpecializedFunctor C TypeCat.
+  Context `(C : @Category objC).
+  Variable F : Functor C TypeCat.
 
   (* Quoting David:
      let F:C-->Type be a functor. An element of the limit is a collection of elements x_c,
@@ -77,12 +77,12 @@
   Definition TypeLimit_Object : TypeCat :=
     { S : forall c : objC, F c | forall c c' (g : C.(Morphism) c c'), F.(MorphismOf) g (S c) = (S c') }.
 
-  Definition TypeLimit_Morphism : SpecializedNaturalTransformation
+  Definition TypeLimit_Morphism : NaturalTransformation
                                    ((DiagonalFunctor TypeCat C) TypeLimit_Object)
                                    F.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun c : objC => (fun S => (proj1_sig S) c))
           _
         )
@@ -91,7 +91,7 @@
   Defined.
 
   Definition TypeLimit_Property_Morphism A'
-             (φ' : SpecializedNaturalTransformation ((DiagonalFunctor TypeCat C) A') F) :
+             (φ' : NaturalTransformation ((DiagonalFunctor TypeCat C) A') F) :
     A' -> TypeLimit_Object.
     intro x; hnf.
     exists (fun c => ComponentsOf φ' c x).
diff -r 4eb6407c6ca3 SigCategory.v
--- a/SigCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq FunctionalExtensionality.
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common Notations FunctorAttributes.
 
 Set Implicit Arguments.
@@ -23,16 +23,16 @@
           end).
 
 Section sig_obj_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Prop.
   Variable Pmor : forall s d : sig Pobj, A.(Morphism) (proj1_sig s) (proj1_sig d) -> Prop.
   Variable Pidentity : forall x, @Pmor x x (Identity (C := A) _).
   Variable Pcompose : forall s d d', forall m1 m2, @Pmor d d' m1 -> @Pmor s d m2 -> @Pmor s d' (Compose (C := A) m1 m2).
 
-  Definition SpecializedCategory_sig : @SpecializedCategory (sig Pobj).
+  Definition Category_sig : @Category (sig Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sig (@Pmor s d))
           (fun x => exist _ _ (Pidentity x))
           (fun s d d' m1 m2 => exist _ _ (Pcompose (proj2_sig m1) (proj2_sig m2)))
@@ -44,8 +44,8 @@
     abstract (intros; simpl_eq; auto with category).
   Defined.
 
-  Definition proj1_sig_functor : SpecializedFunctor SpecializedCategory_sig A
-    := Build_SpecializedFunctor SpecializedCategory_sig A
+  Definition proj1_sig_functor : Functor Category_sig A
+    := Build_Functor Category_sig A
                                 (@proj1_sig _ _)
                                 (fun s d => @proj1_sig _ _)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -59,13 +59,13 @@
 Arguments proj1_sig_functor {objA A Pobj Pmor Pidentity Pcompose}.
 
 Section sig_obj.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Prop.
 
-  Definition SpecializedCategory_sig_obj : @SpecializedCategory (sig Pobj).
+  Definition Category_sig_obj : @Category (sig Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => A.(Morphism) (proj1_sig s) (proj1_sig d))
           (fun x => Identity (C := A) (proj1_sig x))
           (fun s d d' m1 m2 => Compose (C := A) m1 m2)
@@ -77,25 +77,25 @@
     abstract (intros; destruct_sig; simpl; auto with category).
   Defined.
 
-  Definition proj1_sig_obj_functor : SpecializedFunctor SpecializedCategory_sig_obj A
-    := Build_SpecializedFunctor SpecializedCategory_sig_obj A
+  Definition proj1_sig_obj_functor : Functor Category_sig_obj A
+    := Build_Functor Category_sig_obj A
                                 (@proj1_sig _ _)
                                 (fun s d m => m)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition SpecializedCategory_sig_obj_as_sig : @SpecializedCategory (sig Pobj)
-    := @SpecializedCategory_sig _ A Pobj (fun _ _ _ => True) (fun _ => I) (fun _ _ _ _ _ _ _ => I).
+  Definition Category_sig_obj_as_sig : @Category (sig Pobj)
+    := @Category_sig _ A Pobj (fun _ _ _ => True) (fun _ => I) (fun _ _ _ _ _ _ _ => I).
 
-  Definition sig_functor_obj : SpecializedFunctor SpecializedCategory_sig_obj_as_sig SpecializedCategory_sig_obj
-    := Build_SpecializedFunctor SpecializedCategory_sig_obj_as_sig SpecializedCategory_sig_obj
+  Definition sig_functor_obj : Functor Category_sig_obj_as_sig Category_sig_obj
+    := Build_Functor Category_sig_obj_as_sig Category_sig_obj
                                 (fun x => x)
                                 (fun _ _ => @proj1_sig _ _)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition sig_functor_obj_inv : SpecializedFunctor SpecializedCategory_sig_obj SpecializedCategory_sig_obj_as_sig
-    := Build_SpecializedFunctor SpecializedCategory_sig_obj SpecializedCategory_sig_obj_as_sig
+  Definition sig_functor_obj_inv : Functor Category_sig_obj Category_sig_obj_as_sig
+    := Build_Functor Category_sig_obj Category_sig_obj_as_sig
                                 (fun x => x)
                                 (fun _ _ m => exist _ m I)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -117,16 +117,16 @@
 Arguments proj1_sig_obj_functor {objA A Pobj}.
 
 Section sig_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pmor : forall s d, A.(Morphism) s d -> Prop.
 
   Variable Pidentity : forall x, @Pmor x x (Identity (C := A) _).
   Variable Pcompose : forall s d d', forall m1 m2, @Pmor d d' m1 -> @Pmor s d m2 -> @Pmor s d' (Compose (C := A) m1 m2).
 
-  Definition SpecializedCategory_sig_mor : @SpecializedCategory objA.
+  Definition Category_sig_mor : @Category objA.
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sig (@Pmor s d))
           (fun x => exist _ (Identity (C := A) x) (Pidentity x))
           (fun s d d' m1 m2 => exist _ (Compose (proj1_sig m1) (proj1_sig m2)) (Pcompose (proj2_sig m1) (proj2_sig m2)))
@@ -138,25 +138,25 @@
     abstract (intros; simpl_eq; auto with category).
   Defined.
 
-  Definition proj1_sig_mor_functor : SpecializedFunctor SpecializedCategory_sig_mor A
-    := Build_SpecializedFunctor SpecializedCategory_sig_mor A
+  Definition proj1_sig_mor_functor : Functor Category_sig_mor A
+    := Build_Functor Category_sig_mor A
                                 (fun x => x)
                                 (fun s d => @proj1_sig _ _)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition SpecializedCategory_sig_mor_as_sig : @SpecializedCategory (sig (fun _ : objA => True))
-    := @SpecializedCategory_sig _ A _ (fun s d => @Pmor (proj1_sig s) (proj1_sig d)) (fun _ => Pidentity _) (fun _ _ _ _ _ m1 m2 => Pcompose m1 m2).
+  Definition Category_sig_mor_as_sig : @Category (sig (fun _ : objA => True))
+    := @Category_sig _ A _ (fun s d => @Pmor (proj1_sig s) (proj1_sig d)) (fun _ => Pidentity _) (fun _ _ _ _ _ m1 m2 => Pcompose m1 m2).
 
-  Definition sig_functor_mor : SpecializedFunctor SpecializedCategory_sig_mor_as_sig SpecializedCategory_sig_mor
-    := Build_SpecializedFunctor SpecializedCategory_sig_mor_as_sig SpecializedCategory_sig_mor
+  Definition sig_functor_mor : Functor Category_sig_mor_as_sig Category_sig_mor
+    := Build_Functor Category_sig_mor_as_sig Category_sig_mor
                                 (@proj1_sig _ _)
                                 (fun _ _ m => m)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition sig_functor_mor_inv : SpecializedFunctor SpecializedCategory_sig_mor SpecializedCategory_sig_mor_as_sig
-    := Build_SpecializedFunctor SpecializedCategory_sig_mor SpecializedCategory_sig_mor_as_sig
+  Definition sig_functor_mor_inv : Functor Category_sig_mor Category_sig_mor_as_sig
+    := Build_Functor Category_sig_mor Category_sig_mor_as_sig
                                 (fun x => exist _ x I)
                                 (fun _ _ m => m)
                                 (fun _ _ _ _ _ => eq_refl)
diff -r 4eb6407c6ca3 SigSigTCategory.v
--- a/SigSigTCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigSigTCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq.
-Require Export SpecializedCategory Functor SigTCategory.
+Require Export Category Functor SigTCategory.
 Require Import Common Notations.
 
 Set Implicit Arguments.
@@ -13,7 +13,7 @@
 Local Infix "==" := JMeq.
 
 Section sig_sigT_obj_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Prop.
   Variable Pmor : forall s d : sig Pobj, A.(Morphism) (proj1_sig s) (proj1_sig d) -> Type.
 
@@ -32,10 +32,10 @@
     @Pcompose a a b f _ f' (@Pidentity a) ==
     f'.
 
-  Definition SpecializedCategory_sig_sigT : @SpecializedCategory (sig Pobj).
+  Definition Category_sig_sigT : @Category (sig Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sigT (@Pmor s d))
           (fun x => existT _ (Identity (C := A) (proj1_sig x)) (Pidentity x))
           (fun s d d' m1 m2 => existT _ (Compose (C := A) (projT1 m1) (projT1 m2)) (Pcompose (projT2 m1) (projT2 m2)))
@@ -67,8 +67,8 @@
     f'
     := P_RightIdentity f'.
 
-  Let SpecializedCategory_sig_sigT_as_sigT : @SpecializedCategory (sigT Pobj).
-    eapply (@SpecializedCategory_sigT _ A
+  Let Category_sig_sigT_as_sigT : @Category (sigT Pobj).
+    eapply (@Category_sigT _ A
       Pobj
       Pmor'
       Pidentity'
@@ -81,8 +81,8 @@
     subst_body; destruct s, d; simpl in *; eta_red; exact m.
   Defined.
 
-  Definition sig_sigT_functor_sigT : SpecializedFunctor SpecializedCategory_sig_sigT SpecializedCategory_sig_sigT_as_sigT.
-    refine (Build_SpecializedFunctor SpecializedCategory_sig_sigT SpecializedCategory_sig_sigT_as_sigT
+  Definition sig_sigT_functor_sigT : Functor Category_sig_sigT Category_sig_sigT_as_sigT.
+    refine (Build_Functor Category_sig_sigT Category_sig_sigT_as_sigT
       (fun x => x)
       (@sig_sigT_functor_sigT_MorphismOf)
       _
@@ -95,8 +95,8 @@
     subst_body; destruct s, d; simpl in *; eta_red; exact m.
   Defined.
 
-  Definition sigT_functor_sig_sigT : SpecializedFunctor SpecializedCategory_sig_sigT_as_sigT SpecializedCategory_sig_sigT.
-    refine (Build_SpecializedFunctor SpecializedCategory_sig_sigT_as_sigT SpecializedCategory_sig_sigT
+  Definition sigT_functor_sig_sigT : Functor Category_sig_sigT_as_sigT Category_sig_sigT.
+    refine (Build_Functor Category_sig_sigT_as_sigT Category_sig_sigT
       (fun x => x)
       (@sigT_functor_sig_sigT_MorphismOf)
       _
@@ -111,6 +111,6 @@
     split; functor_eq; destruct_sig; reflexivity.
   Qed.
 
-  Definition proj1_functor_sig_sigT : SpecializedFunctor SpecializedCategory_sig_sigT A
+  Definition proj1_functor_sig_sigT : Functor Category_sig_sigT A
     := ComposeFunctors projT1_functor sig_sigT_functor_sigT.
 End sig_sigT_obj_mor.
diff -r 4eb6407c6ca3 SigTCategory.v
--- a/SigTCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigTCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality JMeq.
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common Notations FunctorAttributes FEqualDep.
 
 Set Implicit Arguments.
@@ -27,7 +27,7 @@
           end).
 
 Section sigT_obj_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Type.
   Variable Pmor : forall s d : sigT Pobj, A.(Morphism) (projT1 s) (projT1 d) -> Type.
 
@@ -46,10 +46,10 @@
     @Pcompose a a b f _ f' (@Pidentity a) ==
     f'.
 
-  Definition SpecializedCategory_sigT : @SpecializedCategory (sigT Pobj).
+  Definition Category_sigT : @Category (sigT Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sigT (@Pmor s d))
           (fun x => existT _ (Identity (C := A) (projT1 x)) (Pidentity x))
           (fun s d d' m1 m2 => existT _ (Compose (C := A) (projT1 m1) (projT1 m2)) (Pcompose (projT2 m1) (projT2 m2)))
@@ -61,8 +61,8 @@
     abstract (intros; simpl_eq; auto with category).
   Defined.
 
-  Definition projT1_functor : SpecializedFunctor SpecializedCategory_sigT A
-    := Build_SpecializedFunctor SpecializedCategory_sigT A
+  Definition projT1_functor : Functor Category_sigT A
+    := Build_Functor Category_sigT A
                                 (@projT1 _ _)
                                 (fun _ _ => @projT1 _ _)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -72,13 +72,13 @@
 Arguments projT1_functor {objA A Pobj Pmor Pidentity Pcompose P_Associativity P_LeftIdentity P_RightIdentity}.
 
 Section sigT_obj.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Type.
 
-  Definition SpecializedCategory_sigT_obj : @SpecializedCategory (sigT Pobj).
+  Definition Category_sigT_obj : @Category (sigT Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => A.(Morphism) (projT1 s) (projT1 d))
           (fun x => Identity (C := A) (projT1 x))
           (fun s d d' m1 m2 => Compose (C := A) m1 m2)
@@ -90,27 +90,27 @@
     abstract (intros; destruct_sig; simpl; auto with category).
   Defined.
 
-  Definition projT1_obj_functor : SpecializedFunctor SpecializedCategory_sigT_obj A
-    := Build_SpecializedFunctor SpecializedCategory_sigT_obj A
+  Definition projT1_obj_functor : Functor Category_sigT_obj A
+    := Build_Functor Category_sigT_obj A
                                 (@projT1 _ _)
                                 (fun s d m => m)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition SpecializedCategory_sigT_obj_as_sigT : @SpecializedCategory (sigT Pobj).
-    refine (@SpecializedCategory_sigT _ A Pobj (fun _ _ _ => unit) (fun _ => tt) (fun _ _ _ _ _ _ _ => tt) _ _ _);
+  Definition Category_sigT_obj_as_sigT : @Category (sigT Pobj).
+    refine (@Category_sigT _ A Pobj (fun _ _ _ => unit) (fun _ => tt) (fun _ _ _ _ _ _ _ => tt) _ _ _);
     abstract (simpl; intros; trivial).
   Defined.
 
-  Definition sigT_functor_obj : SpecializedFunctor SpecializedCategory_sigT_obj_as_sigT SpecializedCategory_sigT_obj
-    := Build_SpecializedFunctor SpecializedCategory_sigT_obj_as_sigT SpecializedCategory_sigT_obj
+  Definition sigT_functor_obj : Functor Category_sigT_obj_as_sigT Category_sigT_obj
+    := Build_Functor Category_sigT_obj_as_sigT Category_sigT_obj
                                 (fun x => x)
                                 (fun _ _ => @projT1 _ _)
                                 (fun _ _ _ _ _ => eq_refl)
                                 (fun _ => eq_refl).
 
-  Definition sigT_functor_obj_inv : SpecializedFunctor SpecializedCategory_sigT_obj SpecializedCategory_sigT_obj_as_sigT
-    := Build_SpecializedFunctor SpecializedCategory_sigT_obj SpecializedCategory_sigT_obj_as_sigT
+  Definition sigT_functor_obj_inv : Functor Category_sigT_obj Category_sigT_obj_as_sigT
+    := Build_Functor Category_sigT_obj Category_sigT_obj_as_sigT
                                 (fun x => x)
                                 (fun _ _ m => existT _ m tt)
                                 (fun _ _ _ _ _ => eq_refl)
@@ -128,7 +128,7 @@
 Arguments projT1_obj_functor {objA A Pobj}.
 
 Section sigT_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pmor : forall s d, A.(Morphism) s d -> Type.
 
   Variable Pidentity : forall x, @Pmor x x (Identity (C := A) _).
@@ -146,10 +146,10 @@
     @Pcompose a a b f _ f' (@Pidentity a) ==
     f'.
 
-  Definition SpecializedCategory_sigT_mor : @SpecializedCategory objA.
+  Definition Category_sigT_mor : @Category objA.
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sigT (@Pmor s d))
           (fun x => existT _ (Identity (C := A) x) (Pidentity x))
           (fun s d d' m1 m2 => existT _ (Compose (C := A) (projT1 m1) (projT1 m2)) (Pcompose (projT2 m1) (projT2 m2)))
@@ -161,8 +161,8 @@
     abstract (intros; simpl_eq; auto with category).
   Defined.
 
-  Definition projT1_mor_functor : SpecializedFunctor SpecializedCategory_sigT_mor A.
-    refine (Build_SpecializedFunctor SpecializedCategory_sigT_mor A
+  Definition projT1_mor_functor : Functor Category_sigT_mor A.
+    refine (Build_Functor Category_sigT_mor A
       (fun x => x)
       (fun s d m => projT1 m)
       _
@@ -171,15 +171,15 @@
     intros; reflexivity.
   Defined.
 
-  Definition SpecializedCategory_sigT_mor_as_sigT : @SpecializedCategory (sigT (fun _ : objA => unit)).
-    apply (@SpecializedCategory_sigT _ A _ (fun s d => @Pmor (projT1 s) (projT1 d)) (fun _ => Pidentity _) (fun _ _ _ _ _ m1 m2 => Pcompose m1 m2));
+  Definition Category_sigT_mor_as_sigT : @Category (sigT (fun _ : objA => unit)).
+    apply (@Category_sigT _ A _ (fun s d => @Pmor (projT1 s) (projT1 d)) (fun _ => Pidentity _) (fun _ _ _ _ _ m1 m2 => Pcompose m1 m2));
       abstract (intros; trivial).
   Defined.
 
-  Definition sigT_functor_mor : SpecializedFunctor SpecializedCategory_sigT_mor_as_sigT SpecializedCategory_sigT_mor.
+  Definition sigT_functor_mor : Functor Category_sigT_mor_as_sigT Category_sigT_mor.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (@projT1 _ _)
           (fun _ _ => @id _)
           _
@@ -189,10 +189,10 @@
     simpl; intros; reflexivity.
   Defined.
 
-  Definition sigT_functor_mor_inv : SpecializedFunctor SpecializedCategory_sigT_mor SpecializedCategory_sigT_mor_as_sigT.
+  Definition sigT_functor_mor_inv : Functor Category_sigT_mor Category_sigT_mor_as_sigT.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           (fun x => existT _ x tt)
           (fun _ _  => @id _)
           _
diff -r 4eb6407c6ca3 SigTInducedFunctors.v
--- a/SigTInducedFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigTInducedFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -12,11 +12,11 @@
 Section T2.
   (* use dummy variables so we don't have to specify the types of
      all these hypotheses *)
-  Context `(dummy0 : @SpecializedCategory_sigT objA A Pobj0 Pmor0 Pidentity0 Pcompose0 P_Associativity0 P_LeftIdentity0 P_RightIdentity0).
-  Context `(dummy1 : @SpecializedCategory_sigT objA A Pobj1 Pmor1 Pidentity1 Pcompose1 P_Associativity1 P_LeftIdentity1 P_RightIdentity1).
+  Context `(dummy0 : @Category_sigT objA A Pobj0 Pmor0 Pidentity0 Pcompose0 P_Associativity0 P_LeftIdentity0 P_RightIdentity0).
+  Context `(dummy1 : @Category_sigT objA A Pobj1 Pmor1 Pidentity1 Pcompose1 P_Associativity1 P_LeftIdentity1 P_RightIdentity1).
 
-  Let sigT_cat0 := @SpecializedCategory_sigT objA A Pobj0 Pmor0 Pidentity0 Pcompose0 P_Associativity0 P_LeftIdentity0 P_RightIdentity0.
-  Let sigT_cat1 := @SpecializedCategory_sigT objA A Pobj1 Pmor1 Pidentity1 Pcompose1 P_Associativity1 P_LeftIdentity1 P_RightIdentity1.
+  Let sigT_cat0 := @Category_sigT objA A Pobj0 Pmor0 Pidentity0 Pcompose0 P_Associativity0 P_LeftIdentity0 P_RightIdentity0.
+  Let sigT_cat1 := @Category_sigT objA A Pobj1 Pmor1 Pidentity1 Pcompose1 P_Associativity1 P_LeftIdentity1 P_RightIdentity1.
 
   Variable P_ObjectOf : forall x, Pobj0 x -> Pobj1 x.
 
@@ -40,10 +40,10 @@
     P_MorphismOf (existT (Pmor0 o o) (Identity (projT1 o)) (Pidentity0 o)) =
     Pidentity1 (InducedT2Functor_sigT_ObjectOf o).
 
-  Definition InducedT2Functor_sigT : SpecializedFunctor sigT_cat0 sigT_cat1.
+  Definition InducedT2Functor_sigT : Functor sigT_cat0 sigT_cat1.
     match goal with
-      | [ |- SpecializedFunctor ?C ?D ] =>
-        refine (Build_SpecializedFunctor C D
+      | [ |- Functor ?C ?D ] =>
+        refine (Build_Functor C D
           InducedT2Functor_sigT_ObjectOf
           InducedT2Functor_sigT_MorphismOf
           _
diff -r 4eb6407c6ca3 SigTSigCategory.v
--- a/SigTSigCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigTSigCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import JMeq ProofIrrelevance.
-Require Export SpecializedCategory Functor SigTCategory.
+Require Export Category Functor SigTCategory.
 Require Import Common Notations FEqualDep.
 
 Set Implicit Arguments.
@@ -13,17 +13,17 @@
 Local Infix "==" := JMeq.
 
 Section sigT_sig_obj_mor.
-  Context `(A : @SpecializedCategory objA).
+  Variable A : Category.
   Variable Pobj : objA -> Type.
   Variable Pmor : forall s d : sigT Pobj, A.(Morphism) (projT1 s) (projT1 d) -> Prop.
 
   Variable Pidentity : forall x, @Pmor x x (Identity (C := A) _).
   Variable Pcompose : forall s d d', forall m1 m2, @Pmor d d' m1 -> @Pmor s d m2 -> @Pmor s d' (Compose (C := A) m1 m2).
 
-  Definition SpecializedCategory_sigT_sig : @SpecializedCategory (sigT Pobj).
+  Definition Category_sigT_sig : @Category (sigT Pobj).
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
           (fun s d => sig (@Pmor s d))
           (fun x => existT _ (Identity (C := A) (projT1 x)) (Pidentity x))
           (fun s d d' m1 m2 => existT _ (Compose (C := A) (proj1_sig m1) (proj1_sig m2)) (Pcompose (proj2_sig m1) (proj2_sig m2)))
@@ -35,8 +35,8 @@
     abstract (intros; simpl_eq; auto with category).
   Defined.
 
-  Let SpecializedCategory_sigT_sig_as_sigT : @SpecializedCategory (sigT Pobj).
-    apply (@SpecializedCategory_sigT _ A _ _ Pidentity Pcompose);
+  Let Category_sigT_sig_as_sigT : @Category (sigT Pobj).
+    apply (@Category_sigT _ A _ _ Pidentity Pcompose);
       abstract (
         simpl; intros;
           match goal with
@@ -49,8 +49,8 @@
       ).
   Defined.
 
-  Definition sigT_sig_functor_sigT : SpecializedFunctor SpecializedCategory_sigT_sig SpecializedCategory_sigT_sig_as_sigT.
-    refine (Build_SpecializedFunctor SpecializedCategory_sigT_sig SpecializedCategory_sigT_sig_as_sigT
+  Definition sigT_sig_functor_sigT : Functor Category_sigT_sig Category_sigT_sig_as_sigT.
+    refine (Build_Functor Category_sigT_sig Category_sigT_sig_as_sigT
       (fun x => x)
       (fun s d m => m)
       _
@@ -59,8 +59,8 @@
     abstract (intros; simpl; destruct_sig; reflexivity).
   Defined.
 
-  Definition sigT_functor_sigT_sig : SpecializedFunctor SpecializedCategory_sigT_sig_as_sigT SpecializedCategory_sigT_sig.
-    refine (Build_SpecializedFunctor SpecializedCategory_sigT_sig_as_sigT SpecializedCategory_sigT_sig
+  Definition sigT_functor_sigT_sig : Functor Category_sigT_sig_as_sigT Category_sigT_sig.
+    refine (Build_Functor Category_sigT_sig_as_sigT Category_sigT_sig
       (fun x => x)
       (fun s d m => m)
       _
@@ -75,6 +75,6 @@
     split; functor_eq; destruct_sig; reflexivity.
   Qed.
 
-  Definition proj1_functor_sigT_sig : SpecializedFunctor SpecializedCategory_sigT_sig A
+  Definition proj1_functor_sigT_sig : Functor Category_sigT_sig A
     := ComposeFunctors projT1_functor sigT_sig_functor_sigT.
 End sigT_sig_obj_mor.
diff -r 4eb6407c6ca3 SigTSigInducedFunctors.v
--- a/SigTSigInducedFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/SigTSigInducedFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -13,11 +13,11 @@
 Section T2.
   (* use dummy variables so we don't have to specify the types of
      all these hypotheses *)
-  Context `(dummy0 : @SpecializedCategory_sigT_sig objA A Pobj0 Pmor0 Pidentity0 Pcompose0).
-  Context `(dummy1 : @SpecializedCategory_sigT_sig objA A Pobj1 Pmor1 Pidentity1 Pcompose1).
+  Context `(dummy0 : @Category_sigT_sig objA A Pobj0 Pmor0 Pidentity0 Pcompose0).
+  Context `(dummy1 : @Category_sigT_sig objA A Pobj1 Pmor1 Pidentity1 Pcompose1).
 
-  Let sigT_sig_cat0 := @SpecializedCategory_sigT_sig objA A Pobj0 Pmor0 Pidentity0 Pcompose0.
-  Let sigT_sig_cat1 := @SpecializedCategory_sigT_sig objA A Pobj1 Pmor1 Pidentity1 Pcompose1.
+  Let sigT_sig_cat0 := @Category_sigT_sig objA A Pobj0 Pmor0 Pidentity0 Pcompose0.
+  Let sigT_sig_cat1 := @Category_sigT_sig objA A Pobj1 Pmor1 Pidentity1 Pcompose1.
 
   Variable P_ObjectOf : forall x, Pobj0 x -> Pobj1 x.
 
@@ -37,7 +37,7 @@
   Let sig_of_sigT' (A : Type) (P : A -> Prop) (X : sigT P) := exist P (projT1 X) (projT2 X).
   Let sigT_of_sig' (A : Type) (P : A -> Prop) (X : sig P) := existT P (proj1_sig X) (proj2_sig X).
 
-  Definition InducedT2Functor_sigT_sig : SpecializedFunctor sigT_sig_cat0 sigT_sig_cat1.
+  Definition InducedT2Functor_sigT_sig : Functor sigT_sig_cat0 sigT_sig_cat1.
     eapply (ComposeFunctors (sigT_functor_sigT_sig _ _ _ _) (ComposeFunctors _ (sigT_sig_functor_sigT _ _ _ _))).
     Grab Existential Variables.
     eapply (@InducedT2Functor_sigT _ A Pobj0 Pmor0 Pidentity0 Pcompose0 _ _ _ Pobj1 Pmor1 Pidentity1 Pcompose1 _ _ _
diff -r 4eb6407c6ca3 SimplicialSets.v
--- a/SimplicialSets.v	Fri May 24 20:09:37 2013 -0400
+++ b/SimplicialSets.v	Fri Jun 21 15:07:08 2013 -0400
@@ -13,7 +13,7 @@
   Definition SimplexCategory := @ComputableCategory nat _ (fun n => [n])%category.
   Local Notation Δ := SimplexCategory.
 
-  Definition SimplicialCategory `(C : SpecializedCategory objC) := (C ^ (OppositeCategory Δ))%category.
+  Definition SimplicialCategory `(C : Category objC) := (C ^ (OppositeCategory Δ))%category.
 
   Definition SimplicialSet := SimplicialCategory SetCat.
   Definition SimplicialType := SimplicialCategory TypeCat.
diff -r 4eb6407c6ca3 SpecializedLaxCommaCategory.v
--- a/SpecializedLaxCommaCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SpecializedLaxCommaCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -12,18 +12,17 @@
 
 Local Open Scope category_scope.
 
-Section LaxSliceSpecializedCategory.
+Section LaxSliceCategory.
   (* [Definition]s are not sort-polymorphic. *)
 
   Variable I : Type.
-  Variable Index2Object : I -> Type.
-  Variable Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i).
+  Variable Index2Cat : I -> Category.
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
-  Let Cat := ComputableCategory Index2Object Index2Cat.
+  Let Cat := ComputableCategory Index2Cat.
 
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   (** Quoting David Spivak:
      David: ok
@@ -48,36 +47,36 @@
        the [α]-part is identity.
        and you've worked both of those cases out already.
        *)
-  (* use a pair, so that it's easily interchangable with [SliceSpecializedCategory] *)
-  Record LaxSliceSpecializedCategory_Object := { LaxSliceSpecializedCategory_Object_Member :> { X : I * unit & SpecializedFunctor (fst X) C } }.
+  (* use a pair, so that it's easily interchangable with [SliceCategory] *)
+  Record LaxSliceCategory_Object := { LaxSliceCategory_Object_Member :> { X : I * unit & Functor (fst X) C } }.
 
   Let SortPolymorphic_Helper (A T : Type) (Build_T : A -> T) := A.
 
-  Definition LaxSliceSpecializedCategory_ObjectT := Eval hnf in SortPolymorphic_Helper Build_LaxSliceSpecializedCategory_Object.
-  Global Identity Coercion LaxSliceSpecializedCategory_Object_Id : LaxSliceSpecializedCategory_ObjectT >-> sigT.
-  Definition Build_LaxSliceSpecializedCategory_Object' (mem : LaxSliceSpecializedCategory_ObjectT) := Build_LaxSliceSpecializedCategory_Object mem.
-  Global Coercion Build_LaxSliceSpecializedCategory_Object' : LaxSliceSpecializedCategory_ObjectT >-> LaxSliceSpecializedCategory_Object.
+  Definition LaxSliceCategory_ObjectT := Eval hnf in SortPolymorphic_Helper Build_LaxSliceCategory_Object.
+  Global Identity Coercion LaxSliceCategory_Object_Id : LaxSliceCategory_ObjectT >-> sigT.
+  Definition Build_LaxSliceCategory_Object' (mem : LaxSliceCategory_ObjectT) := Build_LaxSliceCategory_Object mem.
+  Global Coercion Build_LaxSliceCategory_Object' : LaxSliceCategory_ObjectT >-> LaxSliceCategory_Object.
 
-  Record LaxSliceSpecializedCategory_Morphism (XG X'G' : LaxSliceSpecializedCategory_ObjectT) := { LaxSliceSpecializedCategory_Morphism_Member :>
-    { F : SpecializedFunctor (fst (projT1 XG)) (fst (projT1 X'G')) * unit &
-      SpecializedNaturalTransformation (projT2 XG) (ComposeFunctors (projT2 X'G') (fst F))
+  Record LaxSliceCategory_Morphism (XG X'G' : LaxSliceCategory_ObjectT) := { LaxSliceCategory_Morphism_Member :>
+    { F : Functor (fst (projT1 XG)) (fst (projT1 X'G')) * unit &
+      NaturalTransformation (projT2 XG) (ComposeFunctors (projT2 X'G') (fst F))
     }
   }.
 
-  Definition LaxSliceSpecializedCategory_MorphismT (XG X'G' : LaxSliceSpecializedCategory_ObjectT) :=
-    Eval hnf in SortPolymorphic_Helper (@Build_LaxSliceSpecializedCategory_Morphism XG X'G').
-  Global Identity Coercion LaxSliceSpecializedCategory_Morphism_Id : LaxSliceSpecializedCategory_MorphismT >-> sigT.
-  Definition Build_LaxSliceSpecializedCategory_Morphism' XG X'G' (mem : @LaxSliceSpecializedCategory_MorphismT XG X'G') :=
-    @Build_LaxSliceSpecializedCategory_Morphism _ _ mem.
-  Global Coercion Build_LaxSliceSpecializedCategory_Morphism' : LaxSliceSpecializedCategory_MorphismT >-> LaxSliceSpecializedCategory_Morphism.
+  Definition LaxSliceCategory_MorphismT (XG X'G' : LaxSliceCategory_ObjectT) :=
+    Eval hnf in SortPolymorphic_Helper (@Build_LaxSliceCategory_Morphism XG X'G').
+  Global Identity Coercion LaxSliceCategory_Morphism_Id : LaxSliceCategory_MorphismT >-> sigT.
+  Definition Build_LaxSliceCategory_Morphism' XG X'G' (mem : @LaxSliceCategory_MorphismT XG X'G') :=
+    @Build_LaxSliceCategory_Morphism _ _ mem.
+  Global Coercion Build_LaxSliceCategory_Morphism' : LaxSliceCategory_MorphismT >-> LaxSliceCategory_Morphism.
 
-  Global Arguments LaxSliceSpecializedCategory_Object_Member _ : simpl nomatch.
-  Global Arguments LaxSliceSpecializedCategory_Morphism_Member _ _ _ : simpl nomatch.
-  Global Arguments LaxSliceSpecializedCategory_ObjectT /.
-  Global Arguments LaxSliceSpecializedCategory_MorphismT _ _ /.
+  Global Arguments LaxSliceCategory_Object_Member _ : simpl nomatch.
+  Global Arguments LaxSliceCategory_Morphism_Member _ _ _ : simpl nomatch.
+  Global Arguments LaxSliceCategory_ObjectT /.
+  Global Arguments LaxSliceCategory_MorphismT _ _ /.
 
-  Definition LaxSliceSpecializedCategory_Compose' s d d' (Fα : LaxSliceSpecializedCategory_MorphismT d d') (F'α' : LaxSliceSpecializedCategory_MorphismT s d) :
-    LaxSliceSpecializedCategory_MorphismT s d'.
+  Definition LaxSliceCategory_Compose' s d d' (Fα : LaxSliceCategory_MorphismT d d') (F'α' : LaxSliceCategory_MorphismT s d) :
+    LaxSliceCategory_MorphismT s d'.
     exists (ComposeFunctors (fst (projT1 Fα)) (fst (projT1 F'α')), tt).
     repeat match goal with
              | [ H : _ |- _ ] => unique_pose_with_body (fst (projT1 H))
@@ -91,8 +90,8 @@
     end.
     Grab Existential Variables.
     match goal with
-      | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ C : _ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun x => Identity (C := C) _)
           _
         )
@@ -105,25 +104,25 @@
     ).
   Defined.
 
-  Definition LaxSliceSpecializedCategory_Compose'' s d d' (Fα : LaxSliceSpecializedCategory_MorphismT d d') (F'α' : LaxSliceSpecializedCategory_MorphismT s d) :
-    LaxSliceSpecializedCategory_MorphismT s d'.
-    simpl_definition_by_tac_and_exact (@LaxSliceSpecializedCategory_Compose' s d d' Fα F'α') ltac:(unfold LaxSliceSpecializedCategory_Compose' in *).
+  Definition LaxSliceCategory_Compose'' s d d' (Fα : LaxSliceCategory_MorphismT d d') (F'α' : LaxSliceCategory_MorphismT s d) :
+    LaxSliceCategory_MorphismT s d'.
+    simpl_definition_by_tac_and_exact (@LaxSliceCategory_Compose' s d d' Fα F'α') ltac:(unfold LaxSliceCategory_Compose' in *).
   Defined.
 
   (* Then we clean up a bit with reduction. *)
-  Definition LaxSliceSpecializedCategory_Compose s d d' (Fα : LaxSliceSpecializedCategory_MorphismT d d') (F'α' : LaxSliceSpecializedCategory_MorphismT s d) :
-    LaxSliceSpecializedCategory_MorphismT s d'
-    := Eval cbv beta iota zeta delta [LaxSliceSpecializedCategory_Compose''] in (@LaxSliceSpecializedCategory_Compose'' s d d' Fα F'α').
+  Definition LaxSliceCategory_Compose s d d' (Fα : LaxSliceCategory_MorphismT d d') (F'α' : LaxSliceCategory_MorphismT s d) :
+    LaxSliceCategory_MorphismT s d'
+    := Eval cbv beta iota zeta delta [LaxSliceCategory_Compose''] in (@LaxSliceCategory_Compose'' s d d' Fα F'α').
 
-  Global Arguments LaxSliceSpecializedCategory_Compose _ _ _ _ _ /.
+  Global Arguments LaxSliceCategory_Compose _ _ _ _ _ /.
 
-  Definition LaxSliceSpecializedCategory_Identity o : LaxSliceSpecializedCategory_MorphismT o o.
+  Definition LaxSliceCategory_Identity o : LaxSliceCategory_MorphismT o o.
     exists (IdentityFunctor _, tt).
     eapply (NTComposeT _ (IdentityNaturalTransformation _)).
     Grab Existential Variables.
     match goal with
-      | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ C : _ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun x => Identity (C := C) _)
           _
         )
@@ -135,14 +134,14 @@
       ).
   Defined.
 
-  Global Arguments LaxSliceSpecializedCategory_Identity _ /.
+  Global Arguments LaxSliceCategory_Identity _ /.
 
   Local Ltac lax_slice_t :=
     repeat (
       let H := fresh in intro H; destruct H; simpl in * |-
     );
     unfold projT1, projT2;
-      try unfold LaxSliceSpecializedCategory_Compose at 1; try (symmetry; unfold LaxSliceSpecializedCategory_Compose at 1);
+      try unfold LaxSliceCategory_Compose at 1; try (symmetry; unfold LaxSliceCategory_Compose at 1);
         try apply f_equal (* slow; ~ 3s / goal *);
           simpl_eq (* slow; ~ 4s / goal *);
           nt_eq (* slow; ~ 2s / goal *); clear_refl_eq;
@@ -155,41 +154,41 @@
                     subst;
                     trivial.
 
-  Lemma LaxSliceSpecializedCategory_Associativity : forall (o1 o2 o3 o4 : LaxSliceSpecializedCategory_ObjectT)
-    (m1 : LaxSliceSpecializedCategory_MorphismT o1 o2)
-    (m2 : LaxSliceSpecializedCategory_MorphismT o2 o3)
-    (m3 : LaxSliceSpecializedCategory_MorphismT o3 o4),
-    LaxSliceSpecializedCategory_Compose
-    (LaxSliceSpecializedCategory_Compose m3 m2) m1 =
-    LaxSliceSpecializedCategory_Compose m3
-    (LaxSliceSpecializedCategory_Compose m2 m1).
+  Lemma LaxSliceCategory_Associativity : forall (o1 o2 o3 o4 : LaxSliceCategory_ObjectT)
+    (m1 : LaxSliceCategory_MorphismT o1 o2)
+    (m2 : LaxSliceCategory_MorphismT o2 o3)
+    (m3 : LaxSliceCategory_MorphismT o3 o4),
+    LaxSliceCategory_Compose
+    (LaxSliceCategory_Compose m3 m2) m1 =
+    LaxSliceCategory_Compose m3
+    (LaxSliceCategory_Compose m2 m1).
   Proof.
     abstract lax_slice_t.
   Qed.
 
-  Lemma LaxSliceSpecializedCategory_LeftIdentity : forall (a b : LaxSliceSpecializedCategory_ObjectT)
-    (f : LaxSliceSpecializedCategory_MorphismT a b),
-    LaxSliceSpecializedCategory_Compose
-    (LaxSliceSpecializedCategory_Identity b) f = f.
+  Lemma LaxSliceCategory_LeftIdentity : forall (a b : LaxSliceCategory_ObjectT)
+    (f : LaxSliceCategory_MorphismT a b),
+    LaxSliceCategory_Compose
+    (LaxSliceCategory_Identity b) f = f.
   Proof.
     abstract lax_slice_t.
   Qed.
 
-  Lemma LaxSliceSpecializedCategory_RightIdentity : forall (a b : LaxSliceSpecializedCategory_ObjectT)
-    (f : LaxSliceSpecializedCategory_MorphismT a b),
-    LaxSliceSpecializedCategory_Compose
-    f (LaxSliceSpecializedCategory_Identity a) = f.
+  Lemma LaxSliceCategory_RightIdentity : forall (a b : LaxSliceCategory_ObjectT)
+    (f : LaxSliceCategory_MorphismT a b),
+    LaxSliceCategory_Compose
+    f (LaxSliceCategory_Identity a) = f.
   Proof.
     abstract lax_slice_t.
   Qed.
 
-  Definition LaxSliceSpecializedCategory : @SpecializedCategory LaxSliceSpecializedCategory_Object.
+  Definition LaxSliceCategory : @Category LaxSliceCategory_Object.
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
-          LaxSliceSpecializedCategory_Morphism
-          LaxSliceSpecializedCategory_Identity
-          LaxSliceSpecializedCategory_Compose
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
+          LaxSliceCategory_Morphism
+          LaxSliceCategory_Identity
+          LaxSliceCategory_Compose
           _
           _
           _
@@ -199,60 +198,60 @@
       repeat (
         let H := fresh in intro H; destruct H as [ ]; simpl in * |-
       );
-      unfold LaxSliceSpecializedCategory_Morphism_Member, LaxSliceSpecializedCategory_Object_Member,
-        Build_LaxSliceSpecializedCategory_Morphism', Build_LaxSliceSpecializedCategory_Object';
+      unfold LaxSliceCategory_Morphism_Member, LaxSliceCategory_Object_Member,
+        Build_LaxSliceCategory_Morphism', Build_LaxSliceCategory_Object';
         apply f_equal;
-          apply LaxSliceSpecializedCategory_Associativity ||
-            apply LaxSliceSpecializedCategory_LeftIdentity ||
-              apply LaxSliceSpecializedCategory_RightIdentity
+          apply LaxSliceCategory_Associativity ||
+            apply LaxSliceCategory_LeftIdentity ||
+              apply LaxSliceCategory_RightIdentity
     ).
   Defined.
-End LaxSliceSpecializedCategory.
+End LaxSliceCategory.
 
-Hint Unfold LaxSliceSpecializedCategory_Compose LaxSliceSpecializedCategory_Identity : category.
-Hint Constructors LaxSliceSpecializedCategory_Morphism LaxSliceSpecializedCategory_Object : category.
+Hint Unfold LaxSliceCategory_Compose LaxSliceCategory_Identity : category.
+Hint Constructors LaxSliceCategory_Morphism LaxSliceCategory_Object : category.
 
-Section LaxCosliceSpecializedCategory.
+Section LaxCosliceCategory.
   (* [Definition]s are not sort-polymorphic. *)
 
   Variable I : Type.
-  Context `(Index2Cat : forall i : I, @SpecializedCategory (@Index2Object i)).
+  Context `(Index2Cat : forall i : I, @Category (@Index2Object i)).
 
-  Local Coercion Index2Cat : I >-> SpecializedCategory.
+  Local Coercion Index2Cat : I >-> Category.
 
   Let Cat := ComputableCategory Index2Object Index2Cat.
 
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
-  Record LaxCosliceSpecializedCategory_Object := { LaxCosliceSpecializedCategory_Object_Member :> { X : unit * I & SpecializedFunctor (snd X) C } }.
+  Record LaxCosliceCategory_Object := { LaxCosliceCategory_Object_Member :> { X : unit * I & Functor (snd X) C } }.
 
   Let SortPolymorphic_Helper (A T : Type) (Build_T : A -> T) := A.
 
-  Definition LaxCosliceSpecializedCategory_ObjectT := Eval hnf in SortPolymorphic_Helper Build_LaxCosliceSpecializedCategory_Object.
-  Global Identity Coercion LaxCosliceSpecializedCategory_Object_Id : LaxCosliceSpecializedCategory_ObjectT >-> sigT.
-  Definition Build_LaxCosliceSpecializedCategory_Object' (mem : LaxCosliceSpecializedCategory_ObjectT) := Build_LaxCosliceSpecializedCategory_Object mem.
-  Global Coercion Build_LaxCosliceSpecializedCategory_Object' : LaxCosliceSpecializedCategory_ObjectT >-> LaxCosliceSpecializedCategory_Object.
+  Definition LaxCosliceCategory_ObjectT := Eval hnf in SortPolymorphic_Helper Build_LaxCosliceCategory_Object.
+  Global Identity Coercion LaxCosliceCategory_Object_Id : LaxCosliceCategory_ObjectT >-> sigT.
+  Definition Build_LaxCosliceCategory_Object' (mem : LaxCosliceCategory_ObjectT) := Build_LaxCosliceCategory_Object mem.
+  Global Coercion Build_LaxCosliceCategory_Object' : LaxCosliceCategory_ObjectT >-> LaxCosliceCategory_Object.
 
-  Record LaxCosliceSpecializedCategory_Morphism (XG X'G' : LaxCosliceSpecializedCategory_ObjectT) := { LaxCosliceSpecializedCategory_Morphism_Member :>
-    { F : unit * SpecializedFunctor (snd (projT1 X'G')) (snd (projT1 XG)) &
-      SpecializedNaturalTransformation (ComposeFunctors (projT2 XG) (snd F)) (projT2 X'G')
+  Record LaxCosliceCategory_Morphism (XG X'G' : LaxCosliceCategory_ObjectT) := { LaxCosliceCategory_Morphism_Member :>
+    { F : unit * Functor (snd (projT1 X'G')) (snd (projT1 XG)) &
+      NaturalTransformation (ComposeFunctors (projT2 XG) (snd F)) (projT2 X'G')
     }
   }.
 
-  Definition LaxCosliceSpecializedCategory_MorphismT (XG X'G' : LaxCosliceSpecializedCategory_ObjectT) :=
-    Eval hnf in SortPolymorphic_Helper (@Build_LaxCosliceSpecializedCategory_Morphism XG X'G').
-  Global Identity Coercion LaxCosliceSpecializedCategory_Morphism_Id : LaxCosliceSpecializedCategory_MorphismT >-> sigT.
-  Definition Build_LaxCosliceSpecializedCategory_Morphism' XG X'G' (mem : @LaxCosliceSpecializedCategory_MorphismT XG X'G') :=
-    @Build_LaxCosliceSpecializedCategory_Morphism _ _ mem.
-  Global Coercion Build_LaxCosliceSpecializedCategory_Morphism' : LaxCosliceSpecializedCategory_MorphismT >-> LaxCosliceSpecializedCategory_Morphism.
+  Definition LaxCosliceCategory_MorphismT (XG X'G' : LaxCosliceCategory_ObjectT) :=
+    Eval hnf in SortPolymorphic_Helper (@Build_LaxCosliceCategory_Morphism XG X'G').
+  Global Identity Coercion LaxCosliceCategory_Morphism_Id : LaxCosliceCategory_MorphismT >-> sigT.
+  Definition Build_LaxCosliceCategory_Morphism' XG X'G' (mem : @LaxCosliceCategory_MorphismT XG X'G') :=
+    @Build_LaxCosliceCategory_Morphism _ _ mem.
+  Global Coercion Build_LaxCosliceCategory_Morphism' : LaxCosliceCategory_MorphismT >-> LaxCosliceCategory_Morphism.
 
-  Global Arguments LaxCosliceSpecializedCategory_Object_Member _ : simpl nomatch.
-  Global Arguments LaxCosliceSpecializedCategory_Morphism_Member _ _ _ : simpl nomatch.
-  Global Arguments LaxCosliceSpecializedCategory_ObjectT /.
-  Global Arguments LaxCosliceSpecializedCategory_MorphismT _ _ /.
+  Global Arguments LaxCosliceCategory_Object_Member _ : simpl nomatch.
+  Global Arguments LaxCosliceCategory_Morphism_Member _ _ _ : simpl nomatch.
+  Global Arguments LaxCosliceCategory_ObjectT /.
+  Global Arguments LaxCosliceCategory_MorphismT _ _ /.
 
-  Definition LaxCosliceSpecializedCategory_Compose' s d d' (Fα : LaxCosliceSpecializedCategory_MorphismT d d') (F'α' : LaxCosliceSpecializedCategory_MorphismT s d) :
-    LaxCosliceSpecializedCategory_MorphismT s d'.
+  Definition LaxCosliceCategory_Compose' s d d' (Fα : LaxCosliceCategory_MorphismT d d') (F'α' : LaxCosliceCategory_MorphismT s d) :
+    LaxCosliceCategory_MorphismT s d'.
     exists (tt, ComposeFunctors (snd (projT1 F'α')) (snd (projT1 Fα))).
     repeat match goal with
              | [ H : _ |- _ ] => unique_pose_with_body (snd (projT1 H))
@@ -266,8 +265,8 @@
     end.
     Grab Existential Variables.
     match goal with
-      | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ C : _ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun x => Identity (C := C) _)
           _
         )
@@ -280,25 +279,25 @@
     ).
   Defined.
 
-  Definition LaxCosliceSpecializedCategory_Compose'' s d d' (Fα : LaxCosliceSpecializedCategory_MorphismT d d') (F'α' : LaxCosliceSpecializedCategory_MorphismT s d) :
-    LaxCosliceSpecializedCategory_MorphismT s d'.
-    simpl_definition_by_tac_and_exact (@LaxCosliceSpecializedCategory_Compose' s d d' Fα F'α') ltac:(unfold LaxCosliceSpecializedCategory_Compose' in *).
+  Definition LaxCosliceCategory_Compose'' s d d' (Fα : LaxCosliceCategory_MorphismT d d') (F'α' : LaxCosliceCategory_MorphismT s d) :
+    LaxCosliceCategory_MorphismT s d'.
+    simpl_definition_by_tac_and_exact (@LaxCosliceCategory_Compose' s d d' Fα F'α') ltac:(unfold LaxCosliceCategory_Compose' in *).
   Defined.
 
   (* Then we clean up a bit with reduction. *)
-  Definition LaxCosliceSpecializedCategory_Compose s d d' (Fα : LaxCosliceSpecializedCategory_MorphismT d d') (F'α' : LaxCosliceSpecializedCategory_MorphismT s d) :
-    LaxCosliceSpecializedCategory_MorphismT s d'
-    := Eval cbv beta iota zeta delta [LaxCosliceSpecializedCategory_Compose''] in (@LaxCosliceSpecializedCategory_Compose'' s d d' Fα F'α').
+  Definition LaxCosliceCategory_Compose s d d' (Fα : LaxCosliceCategory_MorphismT d d') (F'α' : LaxCosliceCategory_MorphismT s d) :
+    LaxCosliceCategory_MorphismT s d'
+    := Eval cbv beta iota zeta delta [LaxCosliceCategory_Compose''] in (@LaxCosliceCategory_Compose'' s d d' Fα F'α').
 
-  Global Arguments LaxCosliceSpecializedCategory_Compose _ _ _ _ _ /.
+  Global Arguments LaxCosliceCategory_Compose _ _ _ _ _ /.
 
-  Definition LaxCosliceSpecializedCategory_Identity o : LaxCosliceSpecializedCategory_MorphismT o o.
+  Definition LaxCosliceCategory_Identity o : LaxCosliceCategory_MorphismT o o.
     exists (tt, IdentityFunctor _).
     eapply (NTComposeT _ (IdentityNaturalTransformation _)).
     Grab Existential Variables.
     match goal with
-      | [ C : _ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ C : _ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun x => Identity (C := C) _)
           _
         )
@@ -310,14 +309,14 @@
     ).
   Defined.
 
-  Global Arguments LaxCosliceSpecializedCategory_Identity _ /.
+  Global Arguments LaxCosliceCategory_Identity _ /.
 
   Local Ltac lax_coslice_t :=
     repeat (
       let H := fresh in intro H; destruct H; simpl in * |-
     );
     unfold projT1, projT2;
-      try unfold LaxCosliceSpecializedCategory_Compose at 1; try (symmetry; unfold LaxCosliceSpecializedCategory_Compose at 1);
+      try unfold LaxCosliceCategory_Compose at 1; try (symmetry; unfold LaxCosliceCategory_Compose at 1);
         try apply f_equal (* slow; ~ 3s / goal *);
           simpl_eq (* slow; ~ 4s / goal *);
           nt_eq (* slow; ~ 2s / goal *); clear_refl_eq;
@@ -330,41 +329,41 @@
                     subst;
                     trivial.
 
-  Lemma LaxCosliceSpecializedCategory_Associativity : forall (o1 o2 o3 o4 : LaxCosliceSpecializedCategory_ObjectT)
-    (m1 : LaxCosliceSpecializedCategory_MorphismT o1 o2)
-    (m2 : LaxCosliceSpecializedCategory_MorphismT o2 o3)
-    (m3 : LaxCosliceSpecializedCategory_MorphismT o3 o4),
-    LaxCosliceSpecializedCategory_Compose
-    (LaxCosliceSpecializedCategory_Compose m3 m2) m1 =
-    LaxCosliceSpecializedCategory_Compose m3
-    (LaxCosliceSpecializedCategory_Compose m2 m1).
+  Lemma LaxCosliceCategory_Associativity : forall (o1 o2 o3 o4 : LaxCosliceCategory_ObjectT)
+    (m1 : LaxCosliceCategory_MorphismT o1 o2)
+    (m2 : LaxCosliceCategory_MorphismT o2 o3)
+    (m3 : LaxCosliceCategory_MorphismT o3 o4),
+    LaxCosliceCategory_Compose
+    (LaxCosliceCategory_Compose m3 m2) m1 =
+    LaxCosliceCategory_Compose m3
+    (LaxCosliceCategory_Compose m2 m1).
   Proof.
     abstract lax_coslice_t.
   Qed.
 
-  Lemma LaxCosliceSpecializedCategory_LeftIdentity : forall (a b : LaxCosliceSpecializedCategory_ObjectT)
-    (f : LaxCosliceSpecializedCategory_MorphismT a b),
-    LaxCosliceSpecializedCategory_Compose
-    (LaxCosliceSpecializedCategory_Identity b) f = f.
+  Lemma LaxCosliceCategory_LeftIdentity : forall (a b : LaxCosliceCategory_ObjectT)
+    (f : LaxCosliceCategory_MorphismT a b),
+    LaxCosliceCategory_Compose
+    (LaxCosliceCategory_Identity b) f = f.
   Proof.
     abstract lax_coslice_t.
   Qed.
 
-  Lemma LaxCosliceSpecializedCategory_RightIdentity : forall (a b : LaxCosliceSpecializedCategory_ObjectT)
-    (f : LaxCosliceSpecializedCategory_MorphismT a b),
-    LaxCosliceSpecializedCategory_Compose
-    f (LaxCosliceSpecializedCategory_Identity a) = f.
+  Lemma LaxCosliceCategory_RightIdentity : forall (a b : LaxCosliceCategory_ObjectT)
+    (f : LaxCosliceCategory_MorphismT a b),
+    LaxCosliceCategory_Compose
+    f (LaxCosliceCategory_Identity a) = f.
   Proof.
     abstract lax_coslice_t.
   Qed.
 
-  Definition LaxCosliceSpecializedCategory : @SpecializedCategory LaxCosliceSpecializedCategory_Object.
+  Definition LaxCosliceCategory : @Category LaxCosliceCategory_Object.
     match goal with
-      | [ |- @SpecializedCategory ?obj ] =>
-        refine (@Build_SpecializedCategory obj
-          LaxCosliceSpecializedCategory_Morphism
-          LaxCosliceSpecializedCategory_Identity
-          LaxCosliceSpecializedCategory_Compose
+      | [ |- @Category ?obj ] =>
+        refine (@Build_Category obj
+          LaxCosliceCategory_Morphism
+          LaxCosliceCategory_Identity
+          LaxCosliceCategory_Compose
           _
           _
           _
@@ -374,15 +373,15 @@
       repeat (
         let H := fresh in intro H; destruct H as [ ]; simpl in * |-
       );
-      unfold LaxCosliceSpecializedCategory_Morphism_Member, LaxCosliceSpecializedCategory_Object_Member,
-        Build_LaxCosliceSpecializedCategory_Morphism', Build_LaxCosliceSpecializedCategory_Object';
+      unfold LaxCosliceCategory_Morphism_Member, LaxCosliceCategory_Object_Member,
+        Build_LaxCosliceCategory_Morphism', Build_LaxCosliceCategory_Object';
         apply f_equal;
-          apply LaxCosliceSpecializedCategory_Associativity ||
-            apply LaxCosliceSpecializedCategory_LeftIdentity ||
-              apply LaxCosliceSpecializedCategory_RightIdentity
+          apply LaxCosliceCategory_Associativity ||
+            apply LaxCosliceCategory_LeftIdentity ||
+              apply LaxCosliceCategory_RightIdentity
     ).
   Defined.
-End LaxCosliceSpecializedCategory.
+End LaxCosliceCategory.
 
-Hint Unfold LaxCosliceSpecializedCategory_Compose LaxCosliceSpecializedCategory_Identity : category.
-Hint Constructors LaxCosliceSpecializedCategory_Morphism LaxCosliceSpecializedCategory_Object : category.
+Hint Unfold LaxCosliceCategory_Compose LaxCosliceCategory_Identity : category.
+Hint Constructors LaxCosliceCategory_Morphism LaxCosliceCategory_Object : category.
diff -r 4eb6407c6ca3 Subcategory.v
--- a/Subcategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/Subcategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -6,9 +6,9 @@
 
 Set Universe Polymorphism.
 
-Definition Subcategory := @SpecializedCategory_sig.
+Definition Subcategory := @Category_sig.
 Definition SubcategoryInclusionFunctor := @proj1_sig_functor.
-Definition FullSubcategory := @SpecializedCategory_sig_obj.
+Definition FullSubcategory := @Category_sig_obj.
 Definition FullSubcategoryInclusionFunctor := @proj1_sig_obj_functor.
-Definition WideSubcategory := @SpecializedCategory_sig_mor.
+Definition WideSubcategory := @Category_sig_mor.
 Definition WideSubcategoryInclusionFunctor := @proj1_sig_mor_functor.
diff -r 4eb6407c6ca3 SubobjectClassifier.v
--- a/SubobjectClassifier.v	Fri May 24 20:09:37 2013 -0400
+++ b/SubobjectClassifier.v	Fri Jun 21 15:07:08 2013 -0400
@@ -62,7 +62,7 @@
    See for instance (MacLane-Moerdijk, p. 22).
    *)
 
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Local Reserved Notation "'Ω'".
 
diff -r 4eb6407c6ca3 SumCategory.v
--- a/SumCategory.v	Fri May 24 20:09:37 2013 -0400
+++ b/SumCategory.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Export SpecializedCategory Functor.
+Require Export Category Functor.
 Require Import Common.
 
 Set Implicit Arguments.
@@ -10,10 +10,9 @@
 Set Universe Polymorphism.
 
 Section SumCategory.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
-  Definition SumCategory_Morphism (s d : objC + objD) : Type
+  Definition SumCategory_Morphism (s d : C + D) : Type
     := match (s, d) with
          | (inl s, inl d) => C.(Morphism) s d
          | (inr s, inr d) => D.(Morphism) s d
@@ -30,7 +29,8 @@
 
   Global Arguments SumCategory_Identity _ /.
 
-  Definition SumCategory_Compose (s d d' : C + D) (m1 : SumCategory_Morphism d d') (m2 : SumCategory_Morphism s d) : SumCategory_Morphism s d'.
+  Definition SumCategory_Compose (s d d' : C + D) (m1 : SumCategory_Morphism d d') (m2 : SumCategory_Morphism s d)
+  : SumCategory_Morphism s d'.
     (* XXX NOTE: try to use typeclasses and work up to existance of morphisms here *)
     case s, d, d'; simpl in *; try destruct_to_empty_set;
     eapply Compose; eassumption.
@@ -38,14 +38,14 @@
 
   Global Arguments SumCategory_Compose [_ _ _] _ _ /.
 
-  Definition SumCategory : @SpecializedCategory (objC + objD)%type.
-    refine (@Build_SpecializedCategory _
-                                       SumCategory_Morphism
-                                       SumCategory_Identity
-                                       SumCategory_Compose
-                                       _
-                                       _
-                                       _);
+  Definition SumCategory : Category.
+    refine (@Build_Category (C + D)%type
+                            SumCategory_Morphism
+                            SumCategory_Identity
+                            SumCategory_Compose
+                            _
+                            _
+                            _);
     abstract (
         repeat match goal with
                  | [ H : Empty_set |- _ ] => case H
@@ -59,20 +59,19 @@
 Infix "+" := SumCategory : category_scope.
 
 Section SumCategoryFunctors.
-  Context `(C : @SpecializedCategory objC).
-  Context `(D : @SpecializedCategory objD).
+  Variables C D : Category.
 
-  Definition inl_Functor : SpecializedFunctor C (C + D)
-    := Build_SpecializedFunctor C (C + D)
-                                (@inl _ _)
-                                (fun _ _ m => m)
-                                (fun _ _ _ _ _ => eq_refl)
-                                (fun _ => eq_refl).
+  Definition inl_Functor : Functor C (C + D)
+    := Build_Functor C (C + D)
+                     (@inl _ _)
+                     (fun _ _ m => m)
+                     (fun _ _ _ _ _ => eq_refl)
+                     (fun _ => eq_refl).
 
-  Definition inr_Functor : SpecializedFunctor D (C + D)
-    := Build_SpecializedFunctor D (C + D)
-                                (@inr _ _)
-                                (fun _ _ m => m)
-                                (fun _ _ _ _ _ => eq_refl)
-                                (fun _ => eq_refl).
+  Definition inr_Functor : Functor D (C + D)
+    := Build_Functor D (C + D)
+                     (@inr _ _)
+                     (fun _ _ m => m)
+                     (fun _ _ _ _ _ => eq_refl)
+                     (fun _ => eq_refl).
 End SumCategoryFunctors.
diff -r 4eb6407c6ca3 SumInducedFunctors.v
--- a/SumInducedFunctors.v	Fri May 24 20:09:37 2013 -0400
+++ b/SumInducedFunctors.v	Fri Jun 21 15:07:08 2013 -0400
@@ -11,15 +11,15 @@
 
 Section SumCategoryFunctors.
   Section sum_functor.
-    Context `(C : @SpecializedCategory objC).
-    Context `(C' : @SpecializedCategory objC').
-    Context `(D : @SpecializedCategory objD).
+    Variable C : Category.
+    Context `(C' : @Category objC').
+    Variable D : Category.
 
-    Definition sum_Functor (F : SpecializedFunctor C D) (F' : SpecializedFunctor C' D)
-    : SpecializedFunctor (C + C') D.
+    Definition sum_Functor (F : Functor C D) (F' : Functor C' D)
+    : Functor (C + C') D.
       match goal with
-        | [ |- SpecializedFunctor ?C ?D ] =>
-          refine (Build_SpecializedFunctor C D
+        | [ |- Functor ?C ?D ] =>
+          refine (Build_Functor C D
                                            (fun cc' => match cc' with
                                                          | inl c => F c
                                                          | inr c' => F' c'
@@ -43,13 +43,13 @@
         ).
     Defined.
 
-    Definition sum_Functor_Functorial (F G : SpecializedFunctor C D) (F' G' : SpecializedFunctor C' D)
-               (T : SpecializedNaturalTransformation F G)
-               (T' : SpecializedNaturalTransformation F' G')
-    : SpecializedNaturalTransformation (sum_Functor F F') (sum_Functor G G').
+    Definition sum_Functor_Functorial (F G : Functor C D) (F' G' : Functor C' D)
+               (T : NaturalTransformation F G)
+               (T' : NaturalTransformation F' G')
+    : NaturalTransformation (sum_Functor F F') (sum_Functor G G').
       match goal with
-        | [ |- SpecializedNaturalTransformation ?A ?B ] =>
-          refine (Build_SpecializedNaturalTransformation
+        | [ |- NaturalTransformation ?A ?B ] =>
+          refine (Build_NaturalTransformation
                     A B
                     (fun x => match x with
                                 | inl c => T c
@@ -65,14 +65,14 @@
 
   Section swap_functor.
     Definition sum_swap_Functor
-               `(C : @SpecializedCategory objC)
-               `(D : @SpecializedCategory objD)
-    : SpecializedFunctor (C + D) (D + C)
+               `(C : @Category objC)
+               `(D : @Category objD)
+    : Functor (C + D) (D + C)
       := sum_Functor (inr_Functor _ _) (inl_Functor _ _).
 
     Lemma sum_swap_swap_id
-          `(C : @SpecializedCategory objC)
-          `(D : @SpecializedCategory objD)
+          `(C : @Category objC)
+          `(D : @Category objD)
     : ComposeFunctors (sum_swap_Functor C D) (sum_swap_Functor D C) = IdentityFunctor _.
       apply Functor_eq; repeat intros [?|?]; simpl; trivial.
     Qed.
diff -r 4eb6407c6ca3 TypeclassSimplification.v
--- a/TypeclassSimplification.v	Fri May 24 20:09:37 2013 -0400
+++ b/TypeclassSimplification.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Export LtacReifiedSimplification.
-Require Import SpecializedCategory Functor NaturalTransformation.
+Require Import Category Functor NaturalTransformation.
 
 Set Implicit Arguments.
 
@@ -11,7 +11,7 @@
 
 Section SimplifiedMorphism.
   Section single_category.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     Class SimplifiedMorphism {s d} (m : Morphism C s d) :=
       SimplifyMorphism { reified_morphism : ReifiedMorphism C s d;
@@ -43,11 +43,10 @@
   End single_category.
 
   Section functor.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
-    Variable F : SpecializedFunctor C D.
+    Context {C D : Category}.
+    Variable F : Functor C D.
 
-    Global Instance functor_morphism `(@SimplifiedMorphism objC C s d m)
+    Global Instance functor_morphism s d m (_ : @SimplifiedMorphism C s d m)
     : SimplifiedMorphism (MorphismOf F m) | 50
       := SimplifyMorphism (m := MorphismOf F m)
                           (ReifiedFunctorMorphism F (reified_morphism (m := m)))
@@ -55,10 +54,9 @@
   End functor.
 
   Section natural_transformation.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
-    Variables F G : SpecializedFunctor C D.
-    Variable T : SpecializedNaturalTransformation F G.
+    Context {C D : Category}.
+    Variables F G : Functor C D.
+    Variable T : NaturalTransformation F G.
 
     Global Instance nt_morphism x
     : SimplifiedMorphism (T x) | 50
diff -r 4eb6407c6ca3 TypeclassUnreifiedSimplification.v
--- a/TypeclassUnreifiedSimplification.v	Fri May 24 20:09:37 2013 -0400
+++ b/TypeclassUnreifiedSimplification.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,4 +1,4 @@
-Require Import SpecializedCategory Functor NaturalTransformation.
+Require Import Category Functor NaturalTransformation.
 Require Import Notations SumCategory ProductCategory.
 
 Set Implicit Arguments.
@@ -9,9 +9,11 @@
 
 Set Universe Polymorphism.
 
+Local Open Scope morphism_scope.
+
 Section SimplifiedMorphism.
   Section single_category_definition.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     Class > MorphismSimplifiesTo {s d} (m_orig m_simpl : Morphism C s d) :=
       simplified_morphism_ok :> m_orig = m_simpl.
@@ -29,39 +31,39 @@
     trivial.
 
   Section single_category_instances.
-    Context `{C : SpecializedCategory objC}.
+    Context {C : Category}.
 
     Global Instance LeftIdentitySimplify
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C _ _ m2_orig (Identity _))
-    : MorphismSimplifiesTo (Compose m2_orig m1_orig) m1_simpl | 10.
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C _ _ m2_orig (Identity _))
+    : MorphismSimplifiesTo (m2_orig ∘ m1_orig) m1_simpl | 10.
     t.
     Qed.
 
     Global Instance RightIdentitySimplify
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C _ _ m2_orig (Identity _))
-    : MorphismSimplifiesTo (Compose m1_orig m2_orig) m1_simpl | 10.
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C _ _ m2_orig (Identity _))
+    : MorphismSimplifiesTo (m1_orig ∘ m2_orig) m1_simpl | 10.
     t.
     Qed.
 
     Global Instance ComposeToIdentitySimplify
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C d s m2_orig m2_simpl)
-           `(Compose m2_simpl m1_simpl = Identity _)
-    : MorphismSimplifiesTo (Compose m2_orig m1_orig) (Identity _) | 20.
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C d s m2_orig m2_simpl)
+           `(m2_simpl ∘ m1_simpl = Identity _)
+    : MorphismSimplifiesTo (m2_orig ∘ m1_orig) (Identity _) | 20.
     t.
     Qed.
 
-    Global Instance AssociativitySimplify `(@MorphismSimplifiesTo _ C _ _ (@Compose _ C _ c d m3 (@Compose _ C a b c m2 m1)) m_simpl)
-    : MorphismSimplifiesTo (Compose (Compose m3 m2) m1) m_simpl | 1000.
+    Global Instance AssociativitySimplify `(@MorphismSimplifiesTo C _ _ (@Compose C _ c d m3 (@Compose C a b c m2 m1)) m_simpl)
+    : MorphismSimplifiesTo ((m3 ∘ m2) ∘ m1) m_simpl | 1000.
     t.
     Qed.
 
     Global Instance ComposeSimplify
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C d d' m2_orig m2_simpl)
-    : MorphismSimplifiesTo (Compose m2_orig m1_orig) (Compose m2_simpl m1_simpl) | 5000.
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C d d' m2_orig m2_simpl)
+    : MorphismSimplifiesTo (m2_orig ∘ m1_orig) (m2_simpl ∘ m1_simpl) | 5000.
     t.
     Qed.
 
@@ -71,34 +73,33 @@
   End single_category_instances.
 
   Section sum_prod_category_instances.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
+    Context {C D : Category}.
 
     Global Instance SumCategorySimplify_inl
-           `(@MorphismSimplifiesTo _ C s d m_orig m_simpl)
-    : @MorphismSimplifiesTo _ (C + D) (inl s) (inl d) m_orig m_simpl | 100.
+           `(@MorphismSimplifiesTo C s d m_orig m_simpl)
+    : @MorphismSimplifiesTo (C + D) (inl s) (inl d) m_orig m_simpl | 100.
     t.
     Qed.
 
     Global Instance SumCategorySimplify_inr
-           `(@MorphismSimplifiesTo _ D s d m_orig m_simpl)
-    : @MorphismSimplifiesTo _ (C + D) (inr s) (inr d) m_orig m_simpl | 100.
+           `(@MorphismSimplifiesTo D s d m_orig m_simpl)
+    : @MorphismSimplifiesTo (C + D) (inr s) (inr d) m_orig m_simpl | 100.
     t.
     Qed.
 
     Global Instance SumComposeSimplify_inl
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C d d' m2_orig m2_simpl)
-    : MorphismSimplifiesTo (@Compose _ (C + D) (inl s) (inl d) (inl d') m2_orig m1_orig)
-                           (@Compose _ (C + D) (inl s) (inl d) (inl d') m2_simpl m1_simpl) | 50.
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C d d' m2_orig m2_simpl)
+    : MorphismSimplifiesTo (@Compose (C + D) (inl s) (inl d) (inl d') m2_orig m1_orig)
+                           (@Compose (C + D) (inl s) (inl d) (inl d') m2_simpl m1_simpl) | 50.
     t.
     Qed.
 
     Global Instance SumComposeSimplify_inr
-           `(@MorphismSimplifiesTo _ D s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ D d d' m2_orig m2_simpl)
-    : MorphismSimplifiesTo (@Compose _ (C + D) (inr s) (inr d) (inr d') m2_orig m1_orig)
-                           (@Compose _ (C + D) (inr s) (inr d) (inr d') m2_simpl m1_simpl) | 50.
+           `(@MorphismSimplifiesTo D s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo D d d' m2_orig m2_simpl)
+    : MorphismSimplifiesTo (@Compose (C + D) (inr s) (inr d) (inr d') m2_orig m1_orig)
+                           (@Compose (C + D) (inr s) (inr d) (inr d') m2_simpl m1_simpl) | 50.
     t.
     Qed.
 
@@ -118,36 +119,35 @@
 
 
   Section functor_instances.
-    Context `{C : SpecializedCategory objC}.
-    Context `{D : SpecializedCategory objD}.
-    Variable F : SpecializedFunctor C D.
+    Context {C D : Category}.
+    Variable F : Functor C D.
 
-    Global Instance FIdentityOfSimplify `(@MorphismSimplifiesTo _ C x x m_orig (Identity _))
+    Global Instance FIdentityOfSimplify `(@MorphismSimplifiesTo C x x m_orig (Identity _))
     : MorphismSimplifiesTo (MorphismOf F m_orig) (Identity _) | 30.
     t.
     Qed.
 
     Global Instance FCompositionOfSimplify
-           `(@MorphismSimplifiesTo _ C s d m1_orig m1_simpl)
-           `(@MorphismSimplifiesTo _ C d d' m2_orig m2_simpl)
-           `(@MorphismSimplifiesTo _ _ _ _ (Compose m2_simpl m1_simpl) m_comp)
-           `(@MorphismSimplifiesTo _ _ _ _ (MorphismOf F m_comp) m_Fcomp)
-    : MorphismSimplifiesTo (Compose (MorphismOf F m2_orig) (MorphismOf F m1_orig))
+           `(@MorphismSimplifiesTo C s d m1_orig m1_simpl)
+           `(@MorphismSimplifiesTo C d d' m2_orig m2_simpl)
+           `(@MorphismSimplifiesTo _ _ _ (m2_simpl ∘ m1_simpl) m_comp)
+           `(@MorphismSimplifiesTo _ _ _ (MorphismOf F m_comp) m_Fcomp)
+    : MorphismSimplifiesTo (MorphismOf F m2_orig ∘ MorphismOf F m1_orig)
                            m_Fcomp | 30.
     t.
     Qed.
 
     (** TODO(jgross): Remove this kludge *)
     Global Instance FunctorComposeToIdentitySimplify
-           `(@MorphismSimplifiesTo _ D (F s) (F d) m1_orig (MorphismOf F m1_simpl))
-           `(@MorphismSimplifiesTo _ D (F d) (F s) m2_orig (MorphismOf F m2_simpl))
-           `(Compose m2_simpl m1_simpl = Identity _)
-    : MorphismSimplifiesTo (Compose m2_orig m1_orig) (Identity _) | 20.
+           `(@MorphismSimplifiesTo D (F s) (F d) m1_orig (MorphismOf F m1_simpl))
+           `(@MorphismSimplifiesTo D (F d) (F s) m2_orig (MorphismOf F m2_simpl))
+           `(m2_simpl ∘ m1_simpl = Identity _)
+    : MorphismSimplifiesTo (m2_orig ∘ m1_orig) (Identity _) | 20.
     t.
     Qed.
 
     Global Instance FunctorNoSimplify
-           `(@MorphismSimplifiesTo _ C s d m_orig m_simpl)
+           `(@MorphismSimplifiesTo C s d m_orig m_simpl)
     : MorphismSimplifiesTo (MorphismOf F m_orig) (MorphismOf F m_simpl) | 5000.
     t.
     Qed.
@@ -190,20 +190,18 @@
 (**                Goals which aren't yet solved by [rsimplify_morphisms]     **)
 (*******************************************************************************)
 Section bad1.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Context `(E : SpecializedCategory objE).
-  Variable s : SpecializedFunctor D E.
-  Variable s8 : SpecializedFunctor C D.
-  Variable s6 : SpecializedFunctor D E.
-  Variable s7 : SpecializedFunctor C D.
-  Variable s4 : SpecializedFunctor D E.
-  Variable s5 : SpecializedFunctor C D.
-  Variable s2 : SpecializedNaturalTransformation s s6.
-  Variable s3 : SpecializedNaturalTransformation s8 s7.
-  Variable s0 : SpecializedNaturalTransformation s6 s4.
-  Variable s1 : SpecializedNaturalTransformation s7 s5.
-  Variable x : objC.
+  Variables C D E : Category.
+  Variable s : Functor D E.
+  Variable s8 : Functor C D.
+  Variable s6 : Functor D E.
+  Variable s7 : Functor C D.
+  Variable s4 : Functor D E.
+  Variable s5 : Functor C D.
+  Variable s2 : NaturalTransformation s s6.
+  Variable s3 : NaturalTransformation s8 s7.
+  Variable s0 : NaturalTransformation s6 s4.
+  Variable s1 : NaturalTransformation s7 s5.
+  Variable x : C.
 
   Goal
     (Compose (MorphismOf s4 (Compose (s1 x) (s3 x)))
@@ -221,13 +219,12 @@
 
 
 Section bad2.
-  Context `(C : SpecializedCategory objC).
-  Context `(D : SpecializedCategory objD).
-  Variable F : SpecializedFunctor C D.
-  Variable o1 : objC.
-  Variable o2 : objC.
-  Variable o : objC.
-  Variable o0 : objC.
+  Variables C D : Category.
+  Variable F : Functor C D.
+  Variable o1 : C.
+  Variable o2 : C.
+  Variable o : C.
+  Variable o0 : C.
   Variable m : Morphism C o o1.
   Variable m0 : Morphism C o2 o0.
   Variable x : Morphism C o1 o2.
diff -r 4eb6407c6ca3 Yoneda.v
--- a/Yoneda.v	Fri May 24 20:09:37 2013 -0400
+++ b/Yoneda.v	Fri Jun 21 15:07:08 2013 -0400
@@ -1,5 +1,5 @@
 Require Import FunctionalExtensionality.
-Require Export SpecializedCategory Functor FunctorCategory Hom FunctorAttributes.
+Require Export Category Functor FunctorCategory Hom FunctorAttributes.
 Require Import Common ProductCategory SetCategory.
 
 Set Implicit Arguments.
@@ -22,12 +22,12 @@
     end.
 
 Section Yoneda.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Let COp := OppositeCategory C.
 
   Section Yoneda.
-    Let Yoneda_NT s d (f : COp.(Morphism) s d) : SpecializedNaturalTransformation (CovariantHomFunctor C s) (CovariantHomFunctor C d).
-      refine (Build_SpecializedNaturalTransformation
+    Let Yoneda_NT s d (f : COp.(Morphism) s d) : NaturalTransformation (CovariantHomFunctor C s) (CovariantHomFunctor C d).
+      refine (Build_NaturalTransformation
                 (CovariantHomFunctor C s)
                 (CovariantHomFunctor C d)
                 (fun c : C => (fun m : C.(Morphism) _ _ => Compose m f))
@@ -36,10 +36,10 @@
       abstract (intros; simpl; apply functional_extensionality_dep; intros; auto with morphism).
     Defined.
 
-    Definition Yoneda : SpecializedFunctor COp (TypeCat ^ C).
+    Definition Yoneda : Functor COp (TypeCat ^ C).
       match goal with
-        | [ |- SpecializedFunctor ?C0 ?D0 ] =>
-          refine (Build_SpecializedFunctor C0 D0
+        | [ |- Functor ?C0 ?D0 ] =>
+          refine (Build_Functor C0 D0
             (fun c : COp => CovariantHomFunctor C c : TypeCat ^ C)
             (fun s d (f : COp.(Morphism) s d) => Yoneda_NT s d f)
             _
@@ -51,8 +51,8 @@
   End Yoneda.
 
   Section CoYoneda.
-    Let CoYoneda_NT s d (f : C.(Morphism) s d) : SpecializedNaturalTransformation (ContravariantHomFunctor C s) (ContravariantHomFunctor C d).
-      refine (Build_SpecializedNaturalTransformation
+    Let CoYoneda_NT s d (f : C.(Morphism) s d) : NaturalTransformation (ContravariantHomFunctor C s) (ContravariantHomFunctor C d).
+      refine (Build_NaturalTransformation
                 (ContravariantHomFunctor C s)
                 (ContravariantHomFunctor C d)
                 (fun c : C => (fun m : COp.(Morphism) _ _ => Compose m f))
@@ -61,10 +61,10 @@
       abstract (intros; simpl; apply functional_extensionality_dep; intros; auto with morphism).
     Defined.
 
-    Definition CoYoneda : SpecializedFunctor C (TypeCat ^ COp).
+    Definition CoYoneda : Functor C (TypeCat ^ COp).
       match goal with
-        | [ |- SpecializedFunctor ?C0 ?D0 ] =>
-          refine (Build_SpecializedFunctor C0 D0
+        | [ |- Functor ?C0 ?D0 ] =>
+          refine (Build_Functor C0 D0
             (fun c : C => ContravariantHomFunctor C c : TypeCat ^ COp)
             (fun s d (f : C.(Morphism) s d) => CoYoneda_NT s d f)
             _
@@ -77,8 +77,8 @@
 End Yoneda.
 
 Section YonedaLemma.
-  Context `(C : @SpecializedCategory objC).
-  Let COp := OppositeCategory C : SpecializedCategory _.
+  Variable C : Category.
+  Let COp := OppositeCategory C : Category _.
 
   (* Note: If we use [Yoneda _ c] instead, we get Universe Inconsistencies.  Hmm... *)
   Definition YonedaLemmaMorphism (c : C) (X : TypeCat ^ C) : Morphism TypeCat (Morphism (TypeCat ^ C) (Yoneda C c) X) (X c).
@@ -90,8 +90,8 @@
     simpl; intro Xc.
     hnf.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun c' : C => (fun f : Morphism _ c c' => X.(MorphismOf) f Xc))
           _
         )
@@ -117,7 +117,7 @@
 End YonedaLemma.
 
 Section CoYonedaLemma.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
   Let COp := OppositeCategory C.
 
   Definition CoYonedaLemmaMorphism (c : C) (X : TypeCat ^ COp)
@@ -131,8 +131,8 @@
     simpl; intro Xc.
     hnf.
     match goal with
-      | [ |- SpecializedNaturalTransformation ?F ?G ] =>
-        refine (Build_SpecializedNaturalTransformation F G
+      | [ |- NaturalTransformation ?F ?G ] =>
+        refine (Build_NaturalTransformation F G
           (fun c' : COp => (fun f : COp.(Morphism) c c' => X.(MorphismOf) f Xc))
           _
         )
@@ -158,7 +158,7 @@
 End CoYonedaLemma.
 
 Section FullyFaithful.
-  Context `(C : @SpecializedCategory objC).
+  Variable C : Category.
 
   Definition YonedaEmbedding : FunctorFullyFaithful (Yoneda C).
     unfold FunctorFullyFaithful.
