Timings for Opposite.v
Require Import Basics.Overture Basics.Tactics.
Require Import WildCat.Core.
(** ** Opposite categories *)
Definition op (A : Type) := A.
Notation "A ^op" := (op A).
(** This stops typeclass search from trying to unfold op. *)
#[global] Typeclasses Opaque op.
Instance isgraph_op {A : Type} `{IsGraph A}
: IsGraph A^op.
unfold op; exact (fun a b => b $-> a).
Instance is01cat_op {A : Type} `{Is01Cat A} : Is01Cat A^op.
cbv; exact (fun a b c g f => f $o g).
(** We don't invert 2-cells as this is op on the first level. *)
Instance is2graph_op {A : Type} `{Is2Graph A} : Is2Graph A^op.
intros a b; unfold op in *; cbn; exact _.
Instance is1cat_op {A : Type} `{Is1Cat A} : Is1Cat A^op.
snapply Build_Is1Cat; unfold op in *; cbv in *.
srapply Build_Is0Functor.
srapply Build_Is0Functor.
intros a b c d f g h; exact (cat_assoc_opp h g f).
intros a b c d f g h; exact (cat_assoc h g f).
intros a b f; exact (cat_idr f).
intros a b f; exact (cat_idl f).
Instance is1cat_strong_op A `{Is1Cat_Strong A}
: Is1Cat_Strong (A ^op).
snapply Build_Is1Cat_Strong.
intros a b c d f g h; exact (cat_assoc_opp_strong h g f).
intros a b c d f g h; exact (cat_assoc_strong h g f).
(** Opposite groupoids *)
Instance is0gpd_op A `{Is0Gpd A} : Is0Gpd (A^op).
srapply Build_Is0Gpd; unfold op in *; cbn in *.
Instance op0gpd_fun A `{Is0Gpd A} :
Is0Functor( (fun x => x) : A^op -> A).
srapply Build_Is0Functor; unfold op in *; cbn.
(** ** Opposite functors *)
Instance is0functor_op A B (F : A -> B)
`{IsGraph A, IsGraph B, x : !Is0Functor F}
: Is0Functor (F : A^op -> B^op).
Instance is1functor_op A B (F : A -> B)
`{Is1Cat A, Is1Cat B, !Is0Functor F, !Is1Functor F}
: Is1Functor (F : A^op -> B^op).
apply Build_Is1Functor; cbn.
intros a b; rapply fmap2.
intros a b c f g; exact (fmap_comp F g f).
(** Since [Is01Cat] structures are definitionally involutive (see test/WildCat/Opposite.v), we can use [is0functor_op] to transform in the reverse direction as well. This result makes that much easier to use in practice. *)
Instance is0functor_op' A B (F : A^op -> B^op)
`{IsGraph A, IsGraph B, Fop : !Is0Functor (F : A^op -> B^op)}
: Is0Functor (F : A -> B)
:= is0functor_op A^op B^op F.
(** [Is1Cat] structures are also definitionally involutive. *)
Instance is1functor_op' A B (F : A^op -> B^op)
`{Is1Cat A, Is1Cat B, !Is0Functor (F : A^op -> B^op), Fop2 : !Is1Functor (F : A^op -> B^op)}
: Is1Functor (F : A -> B)
:= is1functor_op A^op B^op F.
Instance hasmorext_op {A : Type} `{H0 : HasMorExt A}
: HasMorExt A^op.
exact (@isequiv_Htpy_path _ _ _ _ _ H0 b a f g).
Instance isinitial_op_isterminal {A : Type} `{Is1Cat A} (x : A)
{t : IsTerminal x} : IsInitial (A := A^op) x
:= t.
Instance isterminal_op_isinitial {A : Type} `{Is1Cat A} (x : A)
{i : IsInitial x} : IsTerminal (A := A^op) x
:= i.