Skip to content

Commit

Permalink
Hint mode
Browse files Browse the repository at this point in the history
  • Loading branch information
Patrick Nicodemus committed Mar 1, 2025
1 parent b14f29a commit 4897236
Show file tree
Hide file tree
Showing 12 changed files with 30 additions and 39 deletions.
2 changes: 1 addition & 1 deletion contrib/SetoidRewrite.v
Original file line number Diff line number Diff line change
Expand Up @@ -121,7 +121,7 @@ Defined.

#[export] Instance gpd_hom_is_proper1 {A : Type} `{Is0Gpd A}
: CMorphisms.Proper
(Hom ==> Hom ==> CRelationClasses.arrow) Hom.
(Hom (A:=A)==> Hom ==> CRelationClasses.arrow) Hom.
Proof.
intros x y eq_xy a b eq_ab f.
refine (transitivity _ eq_ab).
Expand Down
15 changes: 6 additions & 9 deletions theories/Algebra/AbGroups/AbHom.v
Original file line number Diff line number Diff line change
Expand Up @@ -40,10 +40,10 @@ Defined.
(** ** Coequalizers *)

(** Using the cokernel and addition and negation for the homs of abelian groups, we can define the coequalizer of two group homomorphisms as the cokernel of their difference. *)
Definition ab_coeq {A B : AbGroup} (f g : GroupHomomorphism A B)
Definition ab_coeq {A B : AbGroup} (f g : Hom A B)
:= ab_cokernel ((-f) + g).

Definition ab_coeq_in {A B} {f g : A $-> B} : B $-> ab_coeq f g.
Definition ab_coeq_in {A B : AbGroup} {f g : A $-> B} : B $-> ab_coeq f g.
Proof.
snrapply grp_quotient_map.
Defined.
Expand Down Expand Up @@ -97,7 +97,7 @@ Proof.
exact p.
Defined.

Definition functor_ab_coeq {A B} {f g : A $-> B} {A' B'} {f' g' : A' $-> B'}
Definition functor_ab_coeq {A B : AbGroup} {f g : A $-> B} {A' B'} {f' g' : A' $-> B'}
(a : A $-> A') (b : B $-> B') (p : f' $o a $== b $o f) (q : g' $o a $== b $o g)
: ab_coeq f g $-> ab_coeq f' g'.
Proof.
Expand All @@ -109,7 +109,7 @@ Proof.
nrapply ab_coeq_glue.
Defined.

Definition functor2_ab_coeq {A B} {f g : A $-> B} {A' B'} {f' g' : A' $-> B'}
Definition functor2_ab_coeq {A B : AbGroup} {f g : A $-> B} {A' B'} {f' g' : A' $-> B'}
{a a' : A $-> A'} {b b' : B $-> B'}
(p : f' $o a $== b $o f) (q : g' $o a $== b $o g)
(p' : f' $o a' $== b' $o f) (q' : g' $o a' $== b' $o g)
Expand All @@ -121,7 +121,7 @@ Proof.
exact (ap ab_coeq_in (s x)).
Defined.

Definition functor_ab_coeq_compose {A B} {f g : A $-> B}
Definition functor_ab_coeq_compose {A B : AbGroup} {f g : A $-> B}
{A' B'} {f' g' : A' $-> B'}
(a : A $-> A') (b : B $-> B') (p : f' $o a $== b $o f) (q : g' $o a $== b $o g)
{A'' B''} {f'' g'' : A'' $-> B''}
Expand Down Expand Up @@ -206,10 +206,7 @@ Proof.
Defined.

Global Instance is0bifunctor_ab_hom `{Funext}
: Is0Bifunctor (ab_hom : Group^op -> AbGroup -> AbGroup).
Proof.
rapply Build_Is0Bifunctor''.
Defined.
: Is0Bifunctor (ab_hom : Group^op -> AbGroup -> AbGroup) := Build_Is0Bifunctor'' _.

Global Instance is1bifunctor_ab_hom `{Funext}
: Is1Bifunctor (ab_hom : Group^op -> AbGroup -> AbGroup).
Expand Down
10 changes: 2 additions & 8 deletions theories/Algebra/AbSES/BaerSum.v
Original file line number Diff line number Diff line change
Expand Up @@ -52,10 +52,7 @@ Proof.
Defined.

Global Instance is0bifunctor_abses' `{Univalence}
: Is0Bifunctor (AbSES' : AbGroup^op -> AbGroup -> Type).
Proof.
rapply Build_Is0Bifunctor''.
Defined.
: Is0Bifunctor (AbSES' : AbGroup^op -> AbGroup -> Type) := Build_Is0Bifunctor'' _.

Global Instance is1bifunctor_abses' `{Univalence}
: Is1Bifunctor (AbSES' : AbGroup^op -> AbGroup -> Type).
Expand Down Expand Up @@ -230,10 +227,7 @@ Proof.
Defined.

Global Instance is0bifunctor_abses `{Univalence}
: Is0Bifunctor (AbSES : AbGroup^op -> AbGroup -> pType).
Proof.
rapply Build_Is0Bifunctor''.
Defined.
: Is0Bifunctor (AbSES : AbGroup^op -> AbGroup -> pType) := Build_Is0Bifunctor'' _.

Global Instance is1bifunctor_abses `{Univalence}
: Is1Bifunctor (AbSES : AbGroup^op -> AbGroup -> pType).
Expand Down
4 changes: 2 additions & 2 deletions theories/Algebra/AbSES/Core.v
Original file line number Diff line number Diff line change
Expand Up @@ -147,7 +147,7 @@ Definition path_abses_iso `{Univalence} {B A : AbGroup@{u}}

(** A special case of the "short 5-lemma" where the two outer maps are (definitionally) identities. *)
Lemma short_five_lemma {B A : AbGroup@{u}}
{E F : AbSES B A} (phi : GroupHomomorphism E F)
{E F : AbSES B A} (phi : Hom (A:=Group) E F)
(p0 : phi $o inclusion E == inclusion F) (p1 : projection E == projection F $o phi)
: IsEquiv phi.
Proof.
Expand Down Expand Up @@ -185,7 +185,7 @@ Defined.

(** Below we prove that homomorphisms respecting [projection] and [inclusion] correspond to paths in [AbSES B A]. We refer to such homomorphisms simply as path data in [AbSES B A]. *)
Definition abses_path_data {B A : AbGroup@{u}} (E F : AbSES B A)
:= {phi : GroupHomomorphism E F
:= {phi : Hom (A:=Group) E F
& (phi $o inclusion _ == inclusion _)
* (projection _ == projection _ $o phi)}.

Expand Down
5 changes: 1 addition & 4 deletions theories/Algebra/AbSES/Ext.v
Original file line number Diff line number Diff line change
Expand Up @@ -119,10 +119,7 @@ Proof.
Defined.

Global Instance is0bifunctor_abext `{Univalence}
: Is0Bifunctor (A:=AbGroup^op) ab_ext.
Proof.
rapply Build_Is0Bifunctor''.
Defined.
: Is0Bifunctor (A:=AbGroup^op) ab_ext := Build_Is0Bifunctor'' _.

Global Instance is1bifunctor_abext `{Univalence}
: Is1Bifunctor (A:=AbGroup^op) ab_ext.
Expand Down
2 changes: 1 addition & 1 deletion theories/Algebra/Categorical/MonoidObject.v
Original file line number Diff line number Diff line change
Expand Up @@ -205,7 +205,7 @@ Defined.
Section MonoidEnriched.
Context {A : Type} `{HasEquivs A} `{!HasBinaryProducts A}
(unit : A) `{!IsTerminal unit} {x y : A}
`{!HasMorExt A} `{forall x y, IsHSet (x $-> y)}.
`{!HasMorExt A} `{forall x y : A, IsHSet (x $-> y)}.

Section Monoid.
Context `{!IsMonoidObject _ _ y}.
Expand Down
3 changes: 3 additions & 0 deletions theories/Algebra/Groups/Group.v
Original file line number Diff line number Diff line change
Expand Up @@ -771,6 +771,9 @@ Defined.
Global Instance isgraph_group : IsGraph Group
:= Build_IsGraph Group GroupHomomorphism.

Definition isHom_GroupIsomorphism (G H : Group) : GroupIsomorphism G H -> Hom G H := idmap.
Coercion isHom_GroupIsomorphism : GroupIsomorphism >-> Hom.

Global Instance is01cat_group : Is01Cat Group :=
Build_Is01Cat Group _ (@grp_homo_id) (@grp_homo_compose).

Expand Down
2 changes: 1 addition & 1 deletion theories/Algebra/Rings/Module.v
Original file line number Diff line number Diff line change
Expand Up @@ -419,7 +419,7 @@ Global Instance isgraph_leftmodule {R : Ring} : IsGraph (LeftModule R)
:= Build_IsGraph _ LeftModuleHomomorphism.

Global Instance is01cat_leftmodule {R : Ring} : Is01Cat (LeftModule R)
:= Build_Is01Cat _ _ lm_homo_id (@lm_homo_compose R).
:= Build_Is01Cat (LeftModule R) _ lm_homo_id (@lm_homo_compose R).

Global Instance is2graph_leftmodule {R : Ring} : Is2Graph (LeftModule R)
:= fun M N => isgraph_induced (@lm_homo_map R M N).
Expand Down
2 changes: 1 addition & 1 deletion theories/Algebra/Rings/Ring.v
Original file line number Diff line number Diff line change
Expand Up @@ -261,7 +261,7 @@ Global Instance isgraph_ring : IsGraph Ring
:= Build_IsGraph _ RingHomomorphism.

Global Instance is01cat_ring : Is01Cat Ring
:= Build_Is01Cat _ _ rng_homo_id (@rng_homo_compose).
:= Build_Is01Cat Ring _ rng_homo_id (@rng_homo_compose).

Global Instance is2graph_ring : Is2Graph Ring
:= fun A B => isgraph_induced (@rng_homo_map A B : _ -> (group_type _ $-> _)).
Expand Down
2 changes: 1 addition & 1 deletion theories/Algebra/ooGroup.v
Original file line number Diff line number Diff line change
Expand Up @@ -270,7 +270,7 @@ End Subgroups.
Global Instance isgraph_oogroup : IsGraph ooGroup
:= Build_IsGraph _ ooGroupHom.
Global Instance is01cat_oogroup : Is01Cat ooGroup
:= Build_Is01Cat _ _ grouphom_idmap (@grouphom_compose).
:= Build_Is01Cat ooGroup _ grouphom_idmap (@grouphom_compose).
Global Instance is2graph_oogroup : Is2Graph ooGroup
:= is2graph_induced classifying_space.
Global Instance is1cat_oogroup : Is1Cat ooGroup
Expand Down
20 changes: 10 additions & 10 deletions theories/WildCat/Core.v
Original file line number Diff line number Diff line change
Expand Up @@ -9,6 +9,8 @@ Class IsGraph (A : Type) :=
Hom : A -> A -> Type
}.

Hint Mode IsGraph ! : typeclass_instances.

Notation "a $-> b" := (Hom a b).

Definition graph_hfiber {B C : Type} `{IsGraph C} (F : B -> C) (c : C)
Expand Down Expand Up @@ -46,7 +48,7 @@ Global Instance reflexive_GpdHom {A} `{Is0Gpd A}
:= fun a => Id a.

Global Instance reflexive_Hom {A} `{Is01Cat A}
: Reflexive Hom
: Reflexive Hom (A:=A)
:= fun a => Id a.

Definition gpd_comp {A} `{Is0Gpd A} {a b c : A}
Expand All @@ -55,21 +57,21 @@ Definition gpd_comp {A} `{Is0Gpd A} {a b c : A}
Infix "$@" := gpd_comp.

Global Instance transitive_GpdHom {A} `{Is0Gpd A}
: Transitive GpdHom
: Transitive GpdHom (A:=A)
:= fun a b c f g => f $@ g.

Global Instance transitive_Hom {A} `{Is01Cat A}
: Transitive Hom
: Transitive Hom (A:=A)
:= fun a b c f g => g $o f.

Notation "p ^$" := (gpd_rev p).

Global Instance symmetric_GpdHom {A} `{Is0Gpd A}
: Symmetric GpdHom
: Symmetric GpdHom (A:=A)
:= fun a b f => f^$.

Global Instance symmetric_GpdHom' {A} `{Is0Gpd A}
: Symmetric Hom
: Symmetric Hom (A:=A)
:= fun a b f => f^$.

Definition Hom_path {A : Type} `{Is01Cat A} {a b : A} (p : a = b) : (a $-> b).
Expand Down Expand Up @@ -197,8 +199,8 @@ Arguments cat_assoc_opp_strong {_ _ _ _ _ _ _ _ _} f g h.
Arguments cat_idl_strong {_ _ _ _ _ _ _} f.
Arguments cat_idr_strong {_ _ _ _ _ _ _} f.

Definition is1cat_is1cat_strong (A : Type) `{Is1Cat_Strong A}
: Is1Cat A.
Global Instance is1cat_is1cat_strong (A : Type) `{Is1Cat_Strong A}
: Is1Cat A | 1000.
Proof.
srapply Build_Is1Cat.
all: intros a b.
Expand All @@ -212,8 +214,6 @@ Proof.
- intros; apply GpdHom_path, cat_idr_strong.
Defined.

Hint Immediate is1cat_is1cat_strong : typeclass_instances.

(** Initial objects *)
Definition IsInitial {A : Type} `{Is1Cat A} (x : A)
:= forall (y : A), {f : x $-> y & forall g, f $== g}.
Expand Down Expand Up @@ -296,7 +296,7 @@ Class Faithful {A B : Type} (F : A -> B) `{Is1Functor A B F} :=

Section IdentityFunctor.

Global Instance is0functor_idmap {A : Type} `{IsGraph A} : Is0Functor idmap.
Global Instance is0functor_idmap {A : Type} `{IsGraph A} : Is0Functor (A:=A) idmap.
Proof.
by apply Build_Is0Functor.
Defined.
Expand Down
2 changes: 1 addition & 1 deletion theories/WildCat/Paths.v
Original file line number Diff line number Diff line change
Expand Up @@ -24,7 +24,7 @@ Local Existing Instances isgraph_paths is2graph_paths is3graph_paths | 10.

(** Any type has composition and identity morphisms given by path concatenation and reflexivity. *)
Global Instance is01cat_paths (A : Type) : Is01Cat A
:= {| Id := @idpath _ ; cat_comp := fun _ _ _ x y => concat y x |}.
:= { Id := @idpath _ ; cat_comp := fun _ _ _ x y => concat y x }.

(** Any type has a 0-groupoid structure with inverse morphisms given by path inversion. *)
Global Instance is0gpd_paths (A : Type) : Is0Gpd A
Expand Down

0 comments on commit 4897236

Please sign in to comment.