Library UniMath.CategoryTheory.TwoSidedDisplayedCats.Examples.Comma

1. Definition via two-sided displayed categories
  Definition comma_twosided_disp_cat_ob_mor
    : twosided_disp_cat_ob_mor C₁ C₂.
  Show proof.
    simple refine (_ ,, _).
    - exact (λ x y, F x --> G y).
    - exact (λ x₁ x₂ y₁ y₂ h₁ h₂ f g, #F f · h₂ = h₁ · #G g).

  Definition comma_twosided_disp_cat_id_comp
    : twosided_disp_cat_id_comp comma_twosided_disp_cat_ob_mor.
  Show proof.
    split.
    - intros x y xy ; cbn.
      rewrite !functor_id.
      rewrite id_left, id_right.
      apply idpath.
    - intros x₁ x₂ x₃ y₁ y₂ y₃ h₁ h₂ h₃ f₁ f₂ g₁ g₂ hh₁ hh₂ ; cbn in *.
      rewrite !functor_comp.
      rewrite !assoc'.
      rewrite hh₂.
      rewrite !assoc.
      apply maponpaths_2.
      exact hh₁.

  Definition comma_twosided_disp_cat_data
    : twosided_disp_cat_data C₁ C₂.
  Show proof.
    simple refine (_ ,, _).
    - exact comma_twosided_disp_cat_ob_mor.
    - exact comma_twosided_disp_cat_id_comp.

  Definition isaprop_comma_twosided_mor
             {x₁ x₂ : C₁}
             {y₁ y₂ : C₂}
             (xy₁ : comma_twosided_disp_cat_data x₁ y₁)
             (xy₂ : comma_twosided_disp_cat_data x₂ y₂)
             (f : x₁ --> x₂)
             (g : y₁ --> y₂)
    : isaprop (xy₁ -->[ f ][ g ] xy₂).
  Show proof.
    apply homset_property.

  Definition comma_twosided_disp_cat_axioms
    : twosided_disp_cat_axioms comma_twosided_disp_cat_data.
  Show proof.
    repeat split.
    - intro ; intros.
      apply isaprop_comma_twosided_mor.
    - intro ; intros.
      apply isaprop_comma_twosided_mor.
    - intro ; intros.
      apply isaprop_comma_twosided_mor.
    - intro ; intros.
      apply isasetaprop.
      apply isaprop_comma_twosided_mor.

  Definition comma_twosided_disp_cat
    : twosided_disp_cat C₁ C₂.
  Show proof.
    simple refine (_ ,, _).
    - exact comma_twosided_disp_cat_data.
    - exact comma_twosided_disp_cat_axioms.

2. Discreteness and univalence
  Definition comma_twosided_disp_cat_is_iso
    : all_disp_mor_iso comma_twosided_disp_cat.
  Show proof.
    intro ; intros.
    simple refine (_ ,, _ ,, _) ; cbn in *.
    - rewrite !functor_on_inv_from_z_iso.
      use z_iso_inv_on_right.
      rewrite assoc.
      use z_iso_inv_on_left ; cbn.
      exact fg.
    - apply isaprop_comma_twosided_mor.
    - apply isaprop_comma_twosided_mor.

  Definition is_univalent_comma_twosided_disp_cat
    : is_univalent_twosided_disp_cat comma_twosided_disp_cat.
  Show proof.
    intros x₁ x₂ y₁ y₂ p₁ p₂ xy₁ xy₂.
    induction p₁, p₂ ; cbn.
    use isweqimplimpl.
    - intros f.
      pose (p := pr1 f) ; cbn in p.
      rewrite !functor_id in p.
      rewrite id_left, id_right in p.
      exact (!p).
    - apply homset_property.
    - use isaproptotal2.
      + intro.
        apply isaprop_is_iso_twosided_disp.
      + intros.
        apply homset_property.

  Definition discrete_comma_twosided_disp_cat
    : discrete_twosided_disp_cat comma_twosided_disp_cat.
  Show proof.
    repeat split.
    - intro ; intros.
      apply homset_property.
    - exact comma_twosided_disp_cat_is_iso.
    - exact is_univalent_comma_twosided_disp_cat.

3. It is a two-sided fibration
  Definition comma_twosided_opcleaving
    : twosided_opcleaving comma_twosided_disp_cat.
  Show proof.
    intros x₁ x₂ x₃ f g ; cbn in *.
    simple refine (f · #G g ,, _ ,, _) ; cbn.
    - rewrite functor_id.
      rewrite id_left.
      apply idpath.
    - intros x₄ x₅ h k l p.
      use iscontraprop1.
      + use invproofirrelevance.
        intros φ φ.
        use subtypePath.
        {
          intro.
          apply isaset_disp_mor.
        }
        apply isaprop_comma_twosided_mor.
      + cbn in *.
        simple refine (_ ,, _).
        * rewrite id_left, functor_comp, assoc in p.
          exact p.
        * apply isaprop_comma_twosided_mor.

  Definition comma_twosided_cleaving
    : twosided_cleaving comma_twosided_disp_cat.
  Show proof.
    intros x₁ x₂ x₃ f g ; cbn in *.
    simple refine (#F g · f ,, _ ,, _) ; cbn.
    - rewrite functor_id.
      rewrite id_right.
      apply idpath.
    - intros x₄ x₅ h k l p.
      use iscontraprop1.
      + use invproofirrelevance.
        intros φ φ.
        use subtypePath.
        {
          intro.
          apply isaset_disp_mor.
        }
        apply isaprop_comma_twosided_mor.
      + cbn in *.
        simple refine (_ ,, _).
        * rewrite id_right, functor_comp, assoc' in p.
          exact p.
        * apply isaprop_comma_twosided_mor.

  Definition comma_twosided_fibration
    : twosided_fibration comma_twosided_disp_cat.
  Show proof.
    simple refine (_ ,, _ ,, _).
    - exact comma_twosided_opcleaving.
    - exact comma_twosided_cleaving.
    - intro ; intros.
      apply comma_twosided_disp_cat_is_iso.
End CommaTwoSidedDispCat.

4. The representable profunctors
Definition left_repr_twosided_disp_cat
           {C₁ C₂ : category}
           (F : C₁ C₂)
  : twosided_disp_cat C₁ C₂
  := comma_twosided_disp_cat F (functor_identity _).

Definition right_repr_twosided_disp_cat
           {C₁ C₂ : category}
           (F : C₁ C₂)
  : twosided_disp_cat C₂ C₁
  := comma_twosided_disp_cat (functor_identity _) F.