X-Git-Url: http://git.megacz.com/?p=coq-categories.git;a=blobdiff_plain;f=src%2FSubcategories_ch7_1.v;h=25391a12eda6a431ed4bac1e99f4d432314bc903;hp=e6b8aa4361d776657419b2a9f8317b468ebedddb;hb=20641452e40570b4bfc9429ca57b0cffca6eccfb;hpb=a0b31d2cc2b6cf7184efe4ff01ad682749f779ad diff --git a/src/Subcategories_ch7_1.v b/src/Subcategories_ch7_1.v index e6b8aa4..25391a1 100644 --- a/src/Subcategories_ch7_1.v +++ b/src/Subcategories_ch7_1.v @@ -3,7 +3,7 @@ (****************************************************************************) Generalizable All Variables. -Require Import Preamble. +Require Import Notations. Require Import Categories_ch1_3. Require Import Functors_ch1_4. Require Import Isomorphisms_ch1_5. @@ -11,38 +11,62 @@ Require Import OppositeCategories_ch1_6_2. Require Import NaturalTransformations_ch7_4. Require Import NaturalIsomorphisms_ch7_5. -(* Any morphism-predicate which is closed under composition and - * passing to identity morphisms (of either the domain or codomain) - * - * We could recycle the "predicate on morphisms" to determine the - * "predicate on objects", but this causes technical difficulties with - * Coq *) -Class SubCategory `(C:Category Ob Hom)(Pobj:Ob->Type)(Pmor:forall a b:Ob, Pobj a -> Pobj b -> (a~>b) ->Type) : Type := -{ sc_id_included : forall (a:Ob)(pa:Pobj a), Pmor _ _ pa pa (id a) -; sc_comp_included : forall (a b c:Ob)(pa:Pobj a)(pb:Pobj b)(pc:Pobj c) f g, - (Pmor _ _ pa pb f) -> (Pmor _ _ pb pc g) -> (Pmor _ _ pa pc (f>>>g)) +(* + * See the README for an explanation of why there is "WideSubcategory" + * and "FullSubcategory" but no "Subcategory" + *) + +(* a full subcategory requires nothing more than a predicate on objects *) +Class FullSubcategory `(C:Category)(Pobj:C->Type) := { }. + +(* the category construction for full subcategories is simpler: *) +Instance FullSubCategoriesAreCategories `(fsc:@FullSubcategory Ob Hom C Pobj) + : Category (sigT Pobj) (fun dom ran => (projT1 dom)~~{C}~~>(projT1 ran)) := +{ id := fun t => id (projT1 t) +; eqv := fun a b f g => eqv _ _ f g +; comp := fun a b c f g => f >>> g }. + intros; apply Build_Equivalence. unfold Reflexive. + intros; reflexivity. + unfold Symmetric; intros; simpl; symmetry; auto. + unfold Transitive; intros; simpl. transitivity y; auto. + intros; unfold Proper. unfold respectful. intros. simpl. apply comp_respects. auto. auto. + intros; simpl. apply left_identity. + intros; simpl. apply right_identity. + intros; simpl. apply associativity. + Defined. +Coercion FullSubCategoriesAreCategories : FullSubcategory >-> Category. -(* every category is a subcategory of itself! *) +(* every category is a subcategory of itself *) +(* Instance IdentitySubCategory `(C:Category Ob Hom) : SubCategory C (fun _ => True) (fun _ _ _ _ _ => True). intros; apply Build_SubCategory. intros; auto. intros; auto. Defined. - -(* a full subcategory requires nothing more than a predicate on objects *) -Definition FullSubcategory `(C:Category)(Pobj:C->Type) : SubCategory C Pobj (fun _ _ _ _ _ => True). - apply Build_SubCategory; intros; auto. +*) +(* the inclusion operation from a subcategory to its host is a functor *) +Instance FullSubcategoryInclusionFunctor `(SP:@FullSubcategory Ob Hom C Pobj) + : Functor SP C (fun x => projT1 x) := + { fmor := fun dom ran f => f }. + intros. unfold eqv in H. simpl in H. auto. + intros. simpl. reflexivity. + intros. simpl. reflexivity. Defined. -Section SubCategoriesAreCategories. - (* Any such predicate determines a category *) - Instance SubCategoriesAreCategories `(C:Category Ob Hom)`(SP:@SubCategory _ _ C Pobj Pmor) - : Category (sigT Pobj) (fun dom ran => sigT (fun f => Pmor (projT1 dom) (projT1 ran) (projT2 dom) (projT2 ran) f)) := - { id := fun t => existT (fun f => Pmor _ _ _ _ f) (id (projT1 t)) (sc_id_included _ (projT2 t)) +(* a wide subcategory includes all objects, so it requires nothing more than a predicate on each hom-set *) +Class WideSubcategory `(C:Category Ob Hom)(Pmor:forall a b:Ob, (a~>b) ->Type) : Type := +{ wsc_id_included : forall (a:Ob), Pmor a a (id a) +; wsc_comp_included : forall (a b c:Ob) f g, (Pmor a b f) -> (Pmor b c g) -> (Pmor a c (f>>>g)) +}. + +(* the category construction for full subcategories is simpler: *) +Instance WideSubCategoriesAreCategories `{C:Category(Ob:=Ob)}{Pmor}(wsc:WideSubcategory C Pmor) + : Category Ob (fun x y => sigT (Pmor x y)) := + { id := fun t => existT _ (id t) (@wsc_id_included _ _ _ _ wsc t) ; eqv := fun a b f g => eqv _ _ (projT1 f) (projT1 g) - ; comp := fun a b c f g => existT (fun f => Pmor _ _ _ _ f) (projT1 f >>> projT1 g) - (sc_comp_included _ _ _ (projT2 a) (projT2 b) (projT2 c) _ _ (projT2 f) (projT2 g)) + ; comp := fun a b c f g => existT (Pmor a c) (projT1 f >>> projT1 g) + (@wsc_comp_included _ _ _ _ wsc _ _ _ _ _ (projT2 f) (projT2 g)) }. intros; apply Build_Equivalence. unfold Reflexive. intros; reflexivity. @@ -53,55 +77,90 @@ Section SubCategoriesAreCategories. intros; simpl. apply right_identity. intros; simpl. apply associativity. Defined. -End SubCategoriesAreCategories. -Coercion SubCategoriesAreCategories : SubCategory >-> Category. +Coercion WideSubCategoriesAreCategories : WideSubcategory >-> Category. -(* the inclusion operation from a subcategory to its host is a functor *) -Instance InclusionFunctor `(C:Category Ob Hom)`(SP:@SubCategory _ _ C Pobj Pmor) - : Functor SP C (fun x => projT1 x) := - { fmor := fun dom ran f => projT1 f }. - intros. unfold eqv in H. simpl in H. auto. - intros. simpl. reflexivity. - intros. simpl. reflexivity. - Defined. +(* the full image of a functor is a full subcategory *) +Section FullImage. -Definition FullImage `(F:Functor(c1:=C)(c2:=D)(fobj:=Fobj)) := FullSubcategory D (fun d => { c:C & (Fobj c)=d }). + Context `(F:Functor(c1:=C)(c2:=D)). -(* any functor may be restricted to its image *) -Section RestrictToImage. - Context `(F:Functor(c1:=C)(c2:=D)(fobj:=Fobj)). - Definition RestrictToImage_fobj : C -> FullImage F. - intros. - exists (F X). - exists X. - reflexivity. - Defined. - Definition RestrictToImage_fmor a b (f:a~>b) : (RestrictToImage_fobj a)~~{FullImage F}~~>(RestrictToImage_fobj b). - exists (F \ f); auto. + Instance FullImage : Category C (fun x y => (F x)~~{D}~~>(F y)) := + { id := fun t => id (F t) + ; eqv := fun x y f g => eqv(Category:=D) _ _ f g + ; comp := fun x y z f g => comp(Category:=D) _ _ _ f g + }. + intros; apply Build_Equivalence. unfold Reflexive. + intros; reflexivity. + unfold Symmetric; intros; simpl; symmetry; auto. + unfold Transitive; intros; simpl. transitivity y; auto. + intros; unfold Proper. unfold respectful. intros. simpl. apply comp_respects. auto. auto. + intros; simpl. apply left_identity. + intros; simpl. apply right_identity. + intros; simpl. apply associativity. + Defined. + + Instance FullImage_InclusionFunctor : Functor FullImage D (fun x => F x) := + { fmor := fun x y f => f }. + intros; auto. + intros; simpl; reflexivity. + intros; simpl; reflexivity. Defined. - Instance RestrictToImage : Functor C (FullImage F) RestrictToImage_fobj := - { fmor := fun a b f => RestrictToImage_fmor a b f }. + + Instance RestrictToImage : Functor C FullImage (fun x => x) := + { fmor := fun a b f => F \ f }. intros; simpl; apply fmor_respects; auto. intros; simpl; apply fmor_preserves_id; auto. intros; simpl; apply fmor_preserves_comp; auto. Defined. - Lemma RestrictToImage_splits : F ≃ (RestrictToImage >>>> InclusionFunctor _ _). - exists (fun A => iso_id (F A)). - intros; simpl. + + Lemma RestrictToImage_splits : F ~~~~ (RestrictToImage >>>> FullImage_InclusionFunctor). + unfold EqualFunctors; simpl; intros; apply heq_morphisms_intro. + apply fmor_respects. + auto. + Defined. + + Lemma RestrictToImage_splits_niso : F ≃ (RestrictToImage >>>> FullImage_InclusionFunctor). + unfold IsomorphicFunctors. + exists (fun A => iso_id (fobj A)). + intros. + simpl. setoid_rewrite left_identity. setoid_rewrite right_identity. reflexivity. Qed. -End RestrictToImage. -Instance func_opSubcat `(c1:Category Ob1 Hom1)`(c2:Category Ob Hom)`(SP:@SubCategory _ _ c2 Pobj Pmor) +End FullImage. + +(* any functor may be restricted to a subcategory of its domain *) +Section RestrictDomain. + + Context `{C:Category}. + Context `{D:Category}. + Context `(F:!Functor C D fobj). + Context {Pmor}(S:WideSubcategory C Pmor). + + Instance RestrictDomain : Functor S D fobj := + { fmor := fun a b f => F \ (projT1 f) }. + intros; destruct f; destruct f'; simpl in *. + apply fmor_respects; auto. + intros. simpl. apply fmor_preserves_id. + intros; simpl; destruct f; destruct g; simpl in *. + apply fmor_preserves_comp. + Defined. + +End RestrictDomain. + +(* +Instance func_opSubcat `(c1:Category)`(c2:Category)`(SP:@SubCategory _ _ c2 Pobj Pmor) {fobj}(F:Functor c1⁽ºᑭ⁾ SP fobj) : Functor c1 SP⁽ºᑭ⁾ fobj := { fmor := fun a b f => fmor F f }. intros. apply (@fmor_respects _ _ _ _ _ _ _ F _ _ f f' H). intros. apply (@fmor_preserves_id _ _ _ _ _ _ _ F a). intros. apply (@fmor_preserves_comp _ _ _ _ _ _ _ F _ _ g _ f). Defined. +*) +(* (* if a functor's range falls within a subcategory, then it is already a functor into that subcategory *) Section FunctorWithRangeInSubCategory. Context `(Cat1:Category O1 Hom1). @@ -160,6 +219,7 @@ Section FunctorWithRangeInSubCategory. *) End Opposite. End FunctorWithRangeInSubCategory. +*) (* Definition 7.1: faithful functors *) @@ -178,6 +238,8 @@ Definition EssentiallySurjectiveOnObjects `(F:Functor(c1:=C1)(c2:=C2)(fobj:=Fobj forall o2:C2, { o1:C1 & (F o1) ≅ o2 }. (* Definition 7.1: (essentially) injective on objects *) +(* TODO *) + Class ConservativeFunctor `(F:Functor(c1:=C1)(c2:=C2)(fobj:=Fobj)) := { cf_reflect_iso : forall (a b:C1), (F a) ≅ (F b) -> a ≅ b ; cf_reflect_iso1 : forall a b (i:(F a) ≅ (F b)), F \ #(cf_reflect_iso a b i) ~~ #i @@ -190,7 +252,7 @@ Definition WeaklyMonic `{D:Category} {Fobj} (F:@Functor _ _ C _ _ D Fobj) := forall - `{E:Category} + Eob EHom (E:@Category Eob EHom) `{G :@Functor _ _ E _ _ C Gobj'} `{H :@Functor _ _ E _ _ C Hobj'}, G >>>> F ≃ H >>>> F @@ -211,20 +273,10 @@ Section FullFaithfulFunctor_section. auto. Qed. - Definition ff_functor_section_fobj (a:FullImage F) : C1 := projT1 (projT2 a). - - Definition ff_functor_section_fmor {a b:FullImage F} (f:a~~{FullImage F}~~>b) - : (ff_functor_section_fobj a)~~{C1}~~>(ff_functor_section_fobj b). - destruct a as [ a1 [ a2 a3 ] ]. - subst. - unfold ff_functor_section_fobj. - simpl. - destruct b as [ b1 [ b2 b3 ] ]. - subst. - unfold ff_functor_section_fobj. - simpl. - apply (@ff_invert _ _ _ _ _ _ _ _ F_full). - apply f. + Definition ff_functor_section_fmor {a b:FullImage F} (f:a~~{FullImage F}~~>b) : a~~{C1}~~>b. + set (@ff_invert _ _ _ _ _ _ _ _ F_full _ _ f) as f'. + destruct f'. + apply x. Defined. Lemma ff_functor_section_respectful {a2 b2 c2 : C1} @@ -247,116 +299,67 @@ Section FullFaithfulFunctor_section. reflexivity. Qed. - Instance ff_functor_section_functor : Functor (FullImage F) C1 ff_functor_section_fobj := + Instance ff_functor_section_functor : Functor (FullImage F) C1 (fun x => x) := { fmor := fun a b f => ff_functor_section_fmor f }. - abstract (intros; - destruct a; destruct b; destruct s; destruct s0; destruct f; destruct f'; simpl in *; - subst; simpl; set (F_full x1 x2 x3) as ff1; set (F_full x1 x2 x4) as ff2; destruct ff1; destruct ff2; - apply F_faithful; - etransitivity; [ apply e | idtac ]; - symmetry; - etransitivity; [ apply e0 | idtac ]; - symmetry; auto). - abstract (intros; - simpl; - destruct a as [ a1 [ a2 a3 ] ]; - subst; - simpl; - apply ff_functor_section_id_preserved). - abstract (intros; - destruct a as [ a1 [ a2 a3 ] ]; - destruct b as [ b1 [ b2 b3 ] ]; - destruct c as [ c1 [ c2 c3 ] ]; - subst; - simpl in *; - destruct f; - destruct g; - simpl in *; - apply ff_functor_section_respectful). + intros. + unfold ff_functor_section_fmor; simpl. + destruct (F_full a b f). + destruct (F_full a b f'). + apply F_faithful. + setoid_rewrite e0. + setoid_rewrite e. + auto. + intros; simpl; subst. + apply ff_functor_section_id_preserved. + intros; simpl in *. + apply ff_functor_section_respectful. Defined. - +(* Lemma ff_functor_section_splits_helper (a2 b2:C1)(f:existT (fun d : C2, {c : C1 & Fobj c = d}) (Fobj a2) (existT (fun c : C1, Fobj c = Fobj a2) a2 (eq_refl _)) ~~{ FullImage F }~~> existT (fun d : C2, {c : C1 & Fobj c = d}) (Fobj b2) (existT (fun c : C1, Fobj c = Fobj b2) b2 (eq_refl _))) - : F \ (let (x1, _) := F_full a2 b2 (let (x1, _) := f in x1) in x1) ~~ projT1 f. - destruct f. + : F \ (let (x1, _) := F_full a2 b2 f in x1) ~~ f. simpl. - set (F_full a2 b2 x) as qq. + set (F_full a2 b2 f) as qq. destruct qq. apply e. Qed. - +*) Lemma ff_functor_section_splits : (ff_functor_section_functor >>>> RestrictToImage F) ~~~~ functor_id _. - unfold EqualFunctors. - intros. - simpl. - destruct a as [ a1 [ a2 a3 ] ]. - destruct b as [ b1 [ b2 b3 ] ]. - subst. - simpl in *. - inversion f; subst. - inversion f'; subst. - simpl in *. - apply heq_morphisms_intro. - simpl. - etransitivity; [ idtac | apply H ]. - clear H. - clear f'. - apply ff_functor_section_splits_helper. + unfold EqualFunctors; intros; simpl. + unfold ff_functor_section_fmor; simpl. + destruct (F_full a b f). + idtac. + apply (@heq_morphisms_intro _ _ (FullImage F) a b). + unfold eqv; simpl. + setoid_rewrite e. + apply H. Qed. - Definition ff_functor_section_splits_niso_helper a : ((ff_functor_section_functor >>>> RestrictToImage F) a ≅ (functor_id (FullImage F)) a). - intros; simpl. - unfold functor_fobj. - unfold ff_functor_section_fobj. - unfold RestrictToImage_fobj. - destruct a as [ a1 [ a2 a3 ] ]. - simpl. - subst. - unfold functor_fobj. - apply iso_id. - Defined. - Lemma ff_functor_section_splits_niso : (ff_functor_section_functor >>>> RestrictToImage F) ≃ functor_id _. intros; simpl. - exists ff_functor_section_splits_niso_helper. - intros. - simpl in *. - destruct A as [ a1 [ a2 a3 ] ]. - destruct B as [ b1 [ b2 b3 ] ]. - simpl. - destruct f; subst. - simpl. + exists iso_id; intros. + symmetry. + unfold functor_comp; simpl. setoid_rewrite left_identity. setoid_rewrite right_identity. - set (F_full a2 b2 x) as qr. - destruct qr. - symmetry; auto. + unfold ff_functor_section_fmor. + destruct (F_full A B f). + auto. Qed. - Definition ff_functor_section_splits_niso_helper' a - : ((RestrictToImage F >>>> ff_functor_section_functor) a ≅ (functor_id _) a). - intros; simpl. - unfold functor_fobj. - unfold ff_functor_section_fobj. - unfold RestrictToImage_fobj. - simpl. - apply iso_id. - Defined. - Lemma ff_functor_section_splits_niso' : (RestrictToImage F >>>> ff_functor_section_functor) ≃ functor_id _. intros; simpl. - exists ff_functor_section_splits_niso_helper'. - intros. - simpl in *. + exists iso_id; intros. + symmetry. + unfold functor_comp; simpl. setoid_rewrite left_identity. setoid_rewrite right_identity. - set (F_full _ _ (F \ f)) as qr. - destruct qr. - apply F_faithful in e. - symmetry. + unfold ff_functor_section_fmor. + destruct (F_full A B (F \ f)). + apply F_faithful. auto. Qed. @@ -393,13 +396,13 @@ Section FullFaithfulFunctor_section. apply if_right_identity. eapply if_inv. apply (if_comp(F2:=G>>>>(RestrictToImage F >>>> ff_functor_section_functor))). - apply (@if_respects _ _ _ _ _ _ _ _ _ _ G _ G _ (functor_id C1) _ (RestrictToImage F >>>> ff_functor_section_functor)). + apply (if_respects G G (functor_id C1) (RestrictToImage F >>>> ff_functor_section_functor)). apply if_id. apply if_inv. apply ff_functor_section_splits_niso'. apply if_inv. apply (if_comp(F2:=H0>>>>(RestrictToImage F >>>> ff_functor_section_functor))). - apply (@if_respects _ _ _ _ _ _ _ _ _ _ H0 _ H0 _ (functor_id C1) _ (RestrictToImage F >>>> ff_functor_section_functor)). + apply (if_respects H0 H0 (functor_id C1) (RestrictToImage F >>>> ff_functor_section_functor)). apply if_id. apply if_inv. apply ff_functor_section_splits_niso'. @@ -407,7 +410,7 @@ Section FullFaithfulFunctor_section. ((H0 >>>> (RestrictToImage F >>>> ff_functor_section_functor)) ≃ ((H0 >>>> RestrictToImage F) >>>> ff_functor_section_functor)). apply if_inv. - apply if_associativity. + apply (if_associativity H0 (RestrictToImage F) ff_functor_section_functor). apply (if_comp H2). clear H2. apply if_inv. @@ -415,15 +418,20 @@ Section FullFaithfulFunctor_section. ((G >>>> (RestrictToImage F >>>> ff_functor_section_functor)) ≃ ((G >>>> RestrictToImage F) >>>> ff_functor_section_functor)). apply if_inv. - apply if_associativity. + apply (if_associativity G (RestrictToImage F) ff_functor_section_functor). apply (if_comp H2). clear H2. - apply if_respects. + apply (if_respects (G >>>> RestrictToImage F) (H0 >>>> RestrictToImage F) + ff_functor_section_functor ff_functor_section_functor). apply if_fullimage. apply H1. - apply if_id. + simpl. + exists (ni_id _). + intros. + simpl. + setoid_rewrite left_identity. + setoid_rewrite right_identity. + reflexivity. Qed. - Opaque ff_functor_section_splits_niso_helper. - End FullFaithfulFunctor_section.