Library UniMath.CategoryTheory.categories.preorder_categories

Category generated by a preorder

Require Import UniMath.Foundations.HLevels.
Require Import UniMath.Foundations.Sets.
Require Import UniMath.Foundations.Propositions.
Require Import UniMath.CategoryTheory.Core.Categories.
Require Import UniMath.CategoryTheory.Core.Isos.
Require Import UniMath.CategoryTheory.Core.Univalence.
Local Open Scope cat.

Section po_category_def.
Context {X : UU}.

Precategory over a preorder
Category over preorder
If the preorder is antisymmetric and X is a set, the category is univalent
Context (xisaset : isaset X).

Lemma antisymm_po_category_isoiseq (PO : po X) {A B : (po_category PO)}
  (poasymm : isantisymm PO) (isoAB : z_iso A B) : A = B.
Show proof.
  apply poasymm.
  apply (z_iso_mor isoAB).
  apply (inv_from_z_iso isoAB).

Lemma antisymm_po_category_isweq (PO : po X) {A B : po_category PO}
      (poasymm : isantisymm PO) : isweq (λ p : A = B, idtoiso p).
Show proof.
  use isweq_iso.
  - apply (antisymm_po_category_isoiseq PO poasymm).
  - intro eq.
    apply proofirrelevance.
    apply xisaset.
  - intro iso.
    use z_iso_eq.
    apply proofirrelevance.
    apply propproperty.

Theorem po_category_is_univalent_iff_is_antisymm (PO : po X) :
  is_univalent (po_category PO) <-> isantisymm PO.
Show proof.
  split.
  - intros isuni a b relab relba.
    apply (isotoid _ isuni).
    exists relab.
    exists relba.
    apply make_is_inverse_in_precat; apply po_homsets_isaprop.
  - intro poasymm.
    intros ? ?.
    apply (antisymm_po_category_isweq PO poasymm).

Definition antisymm_po_univalent_category (PO : po X) (poasymm : isantisymm PO) :
  univalent_category.
  use make_univalent_category.
  - exact (po_category PO).
  - apply po_category_is_univalent_iff_is_antisymm.
    exact poasymm.
Defined.

End po_category_def.