Library UniMath.CategoryTheory.Hyperdoctrines.PartialEqRels.PERSubobjectClassifier

1. The partial setoid representing the subobject classifier

  Definition omega_per_form
    : form ((Ω ×h Ω) : ty H)
    := let Γ := Ω ×h Ω : ty H in
       let φ := Prf [ π (tm_var Γ) ] in
       let ψ := Prf [ π (tm_var Γ) ] in
       φ ψ.

  Arguments omega_per_form /.

  Proposition omega_per_axioms
    : per_axioms omega_per_form.
  Proof.
    split.
    - unfold per_symm_axiom ; cbn.
      do 2 use forall_intro.
      use impl_intro.
      use weaken_right.
      hypersimplify.
      pose (φ := Prf [ π (tm_var (((𝟙 : ty H) ×h Ω) ×h Ω))) ] ).
      pose (ψ := Prf [ π (tm_var (((𝟙 : ty H) ×h Ω) ×h Ω)) ]).
      fold φ ψ.
      use iff_sym.
      apply hyperdoctrine_hyp.
    - unfold per_trans_axiom ; cbn.
      do 3 use forall_intro.
      use impl_intro.
      use weaken_right.
      use impl_intro.
      hypersimplify.
      cbn.
      pose (φ := Prf [ π (tm_var ((((𝟙 : ty H) ×h Ω) ×h Ω) ×h Ω)))) ]).
      pose (ψ := Prf [ π (tm_var ((((𝟙 : ty H) ×h Ω) ×h Ω) ×h Ω))) ]).
      pose (χ := Prf [ π (tm_var ((((𝟙 : ty H) ×h Ω) ×h Ω) ×h Ω)) ]).
      fold φ ψ χ.
      use iff_trans.
      + exact ψ.
      + use weaken_left.
        apply hyperdoctrine_hyp.
      + use weaken_right.
        apply hyperdoctrine_hyp.
  Qed.

  Definition omega_per
    : per (Ω : ty H).
  Proof.
    use make_per.
    - exact omega_per_form.
    - exact omega_per_axioms.
  Defined.

  Definition omega_partial_setoid
    : partial_setoid H.
  Proof.
    use make_partial_setoid.
    - exact Ω.
    - exact omega_per.
  Defined.

2. Accessors for the partial equivalence relation

  Proposition eq_in_omega_partial_setoid
              {Γ : ty H}
              {Δ : form Γ}
              {ω ω : tm Γ omega_partial_setoid}
              (p : Δ Prf [ ω ] Prf [ ω ])
    : Δ ω ~ ω.
  Proof.
    unfold partial_setoid_formula ; cbn.
    hypersimplify.
    exact p.
  Qed.

  Proposition from_eq_in_omega_partial_setoid
              {Γ : ty H}
              {Δ : form Γ}
              {ω ω : tm Γ omega_partial_setoid}
              (p : Δ ω ~ ω)
    : Δ Prf [ ω ] Prf [ ω ].
  Proof.
    refine (hyperdoctrine_cut p _).
    unfold partial_setoid_formula ; cbn.
    hypersimplify.
    apply hyperdoctrine_hyp.
  Qed.

  Proposition from_eq_in_omega_partial_setoid_left
              {Γ : ty H}
              {Δ : form Γ}
              {ω ω : tm Γ omega_partial_setoid}
              (p : Δ ω ~ ω)
              (q : Δ Prf [ ω ])
    : Δ Prf [ ω ].
  Proof.
    use (iff_elim_left (from_eq_in_omega_partial_setoid p)).
    exact q.
  Qed.

  Proposition from_eq_in_omega_partial_setoid_right
              {Γ : ty H}
              {Δ : form Γ}
              {ω ω : tm Γ omega_partial_setoid}
              (p : Δ ω ~ ω)
              (q : Δ Prf [ ω ])
    : Δ Prf [ ω ].
  Proof.
    use (iff_elim_right (from_eq_in_omega_partial_setoid p)).
    exact q.
  Qed.

3. The map representing truth

  Definition omega_partial_setoid_true_form
    : form (eq_partial_setoid 𝟙 ×h omega_partial_setoid)
    := Prf [ π (tm_var _) ].

  Arguments omega_partial_setoid_true_form /.

  Proposition omega_partial_setoid_true_laws
    : partial_setoid_morphism_laws omega_partial_setoid_true_form.
  Proof.
    repeat split.
    - unfold partial_setoid_mor_dom_defined_law ; cbn.
      do 2 use forall_intro.
      use impl_intro.
      use weaken_right.
      hypersimplify.
      pose (ω := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω))).
      pose (t₁ := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω)))).
      pose (t₂ := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω)))).
      fold ω t₁ t₂.
      use eq_in_eq_partial_setoid.
      use hyperdoctrine_unit_tm_eq.
    - unfold partial_setoid_mor_cod_defined_law ; cbn.
      do 2 use forall_intro.
      use impl_intro.
      use weaken_right.
      hypersimplify.
      pose (ω := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω))).
      pose (t₁ := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω)))).
      pose (t₂ := π (tm_var (((𝟙 : ty H) ×h 𝟙) ×h Ω)))).
      fold ω t₁ t₂.
      use eq_in_omega_partial_setoid.
      apply iff_refl.
    - unfold partial_setoid_mor_eq_defined_law ; cbn.
      do 4 (use forall_intro).
      use impl_intro.
      use weaken_right.
      do 2 use impl_intro.
      hypersimplify.
      pose (t₁ := π (tm_var (((((𝟙 : ty H) ×h 𝟙) ×h 𝟙) ×h Ω) ×h Ω)))))).
      pose (t₂ := π (tm_var (((((𝟙 : ty H) ×h 𝟙) ×h 𝟙) ×h Ω) ×h Ω))))).
      pose := π (tm_var (((((𝟙 : ty H) ×h 𝟙) ×h 𝟙) ×h Ω) ×h Ω)))).
      pose := π (tm_var (((((𝟙 : ty H) ×h 𝟙) ×h 𝟙) ×h Ω) ×h Ω))).
      fold t₁ t₂ ω ω.
      use hyp_ltrans.
      use weaken_right.
      use iff_elim_left.
      + exact (Prf [ ω ]).
      + use weaken_left.
        use from_eq_in_omega_partial_setoid.
        apply hyperdoctrine_hyp.
      + use weaken_right.
        apply hyperdoctrine_hyp.
    - unfold partial_setoid_mor_unique_im_law ; cbn.
      do 3 use forall_intro.
      use impl_intro.
      use weaken_right.
      use impl_intro.
      hypersimplify.
      pose := π (tm_var ((((𝟙 : ty H) ×h 𝟙) ×h Ω) ×h Ω)))).
      pose := π (tm_var ((((𝟙 : ty H) ×h 𝟙) ×h Ω) ×h Ω))).
      fold ω ω.
      use eq_in_omega_partial_setoid.
      use iff_true_true.
      + use weaken_left.
        apply hyperdoctrine_hyp.
      + use weaken_right.
        apply hyperdoctrine_hyp.
    - unfold partial_setoid_mor_hom_exists_law ; cbn.
      use forall_intro.
      use impl_intro.
      use weaken_right.
      pose (t₁ := π (tm_var ((𝟙 : ty H) ×h 𝟙))).
      pose (t₂ := π (tm_var ((𝟙 : ty H) ×h 𝟙))).
      use exists_intro.
      {
        exact (tripos_form_to_tm ).
      }
      hypersimplify.
      fold t₁ t₂.
      rewrite tripos_form_to_tm_Prf.
      apply truth_intro.
  Qed.

  Definition omega_partial_setoid_true
    : partial_setoid_morphism (eq_partial_setoid 𝟙) omega_partial_setoid.
  Proof.
    use make_partial_setoid_morphism.
    - exact omega_partial_setoid_true_form.
    - exact omega_partial_setoid_true_laws.
  Defined.

4. The universal mapping property of the subobject classifier

  Section UMP.
    Context {X Y : partial_setoid H}
            (m : Monic (category_of_partial_setoids H) X Y).

    Let φ : partial_setoid_morphism X Y := pr1 m.
    Let Hφ : isMonic (C := category_of_partial_setoids H) φ := pr2 m.

4.1. Maps to the subobject classifier

    Definition subobject_classifier_partial_setoid_map_form
      : form (Y ×h omega_partial_setoid)
      := let x := π (tm_var ((Y ×h Ω) ×h X)) in
         let y := π (tm_var (Y ×h Ω)) in
         let ω := π (tm_var (Y ×h Ω)) in
         y ~ y
         
         (h (φ [ x , y [ π (tm_var _) ]tm ]) Prf [ ω ]).

    Arguments subobject_classifier_partial_setoid_map_form /.

    Proposition subobject_classifier_partial_setoid_map_laws
      : partial_setoid_morphism_laws
          subobject_classifier_partial_setoid_map_form.
    Proof.
      repeat split.
      - unfold partial_setoid_mor_dom_defined_law ; cbn.
        do 2 (use forall_intro).
        use impl_intro.
        use weaken_right.
        hypersimplify_form.
        hypersimplify.
        pose (y := π (tm_var ((𝟙 ×h Y) ×h Ω)))).
        pose (ω := π (tm_var ((𝟙 ×h Y) ×h Ω))).
        fold y ω.
        use weaken_left.
        apply hyperdoctrine_hyp.
      - unfold partial_setoid_mor_cod_defined_law ; cbn.
        do 2 (use forall_intro).
        use impl_intro.
        use weaken_right.
        hypersimplify_form.
        hypersimplify.
        pose (y := π (tm_var ((𝟙 ×h Y) ×h Ω)))).
        pose (ω := π (tm_var ((𝟙 ×h Y) ×h Ω))).
        fold y ω.
        use eq_in_omega_partial_setoid.
        apply iff_refl.
      - unfold partial_setoid_mor_eq_defined_law ; cbn.
        do 4 use forall_intro.
        use impl_intro.
        use weaken_right.
        do 2 use impl_intro.
        hypersimplify_form.
        hypersimplify.
        pose (y₁ := π (tm_var ((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω)))))).
        pose (y₂ := π (tm_var ((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω))))).
        pose := π (tm_var ((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω)))).
        pose := π (tm_var ((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω))).
        fold y₁ y₂ ω ω.
        use conj_intro.
        + do 2 use weaken_left.
          exact (partial_setoid_refl_r (hyperdoctrine_hyp _)).
        + use iff_intro.
          * use (from_eq_in_omega_partial_setoid_left := ω) (ω := ω)).
            ** do 2 use weaken_left.
               use weaken_right.
               apply hyperdoctrine_hyp.
            ** use (iff_elim_left
                      (weaken_left (weaken_right (weaken_right (hyperdoctrine_hyp _) _) _) _)).
               use hyp_sym.
               use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
               rewrite conj_subst.
               use hyp_ltrans.
               use weaken_right.
               hypersimplify_form.
               unfold y₁, y₂, ω, ω ; clear y₁ y₂ ω ω.
               hypersimplify.
               pose (y₁ := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))))))).
               pose (y₂ := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X)))))).
               pose := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))))).
               pose := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X)))).
               pose (x := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))).
               cbn.
               fold y₁ y₂ ω ω x.
               use exists_intro.
               {
                 exact x.
               }
               hypersimplify.
               fold y₁.
               use (partial_setoid_mor_eq_defined φ).
               *** exact x.
               *** exact y₂.
               *** use weaken_right.
                   use (partial_setoid_mor_dom_defined φ x y₂).
                   apply hyperdoctrine_hyp.
               *** use partial_setoid_sym.
                   do 3 use weaken_left.
                   apply hyperdoctrine_hyp.
               *** use weaken_right.
                   apply hyperdoctrine_hyp.
          * use weaken_cut.
            ** exact (h (φ [ π (tm_var _) , π (tm_var _))))) ])).
            ** use iff_elim_right.
               *** exact (Prf [ ω ]).
               *** use weaken_left.
                   do 2 use weaken_right.
                   apply hyperdoctrine_hyp.
               *** use from_eq_in_omega_partial_setoid_left.
                   **** exact ω.
                   **** do 2 use weaken_left.
                        use weaken_right.
                        use partial_setoid_sym.
                        apply hyperdoctrine_hyp.
                   **** use weaken_right.
                        apply hyperdoctrine_hyp.
            ** use hyp_sym.
               use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
               rewrite conj_subst.
               use hyp_ltrans.
               use weaken_right.
               hypersimplify_form.
               unfold y₁, y₂, ω, ω ; clear y₁ y₂ ω ω.
               hypersimplify.
               pose (y₁ := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))))))).
               pose (y₂ := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X)))))).
               pose := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))))).
               pose := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X)))).
               pose (x := π (tm_var (((((𝟙 ×h Y) ×h Y) ×h Ω) ×h Ω) ×h X))).
               cbn.
               fold y₁ y₂ ω ω x.
               use exists_intro.
               {
                 exact x.
               }
               hypersimplify.
               fold y₂.
               use (partial_setoid_mor_eq_defined
                      φ
                      _ _
                      (weaken_right (hyperdoctrine_hyp _) _)).
               *** use weaken_right.
                   use (partial_setoid_mor_dom_defined φ x y₁).
                   apply hyperdoctrine_hyp.
               *** do 4 use weaken_left.
                   apply hyperdoctrine_hyp.
      - unfold partial_setoid_mor_unique_im_law ; cbn.
        do 3 use forall_intro.
        use impl_intro.
        use weaken_right.
        hypersimplify_form.
        hypersimplify.
        pose (y := π (tm_var (((𝟙 ×h Y) ×h Ω) ×h Ω))))).
        pose := π (tm_var (((𝟙 ×h Y) ×h Ω) ×h Ω)))).
        pose := π (tm_var (((𝟙 ×h Y) ×h Ω) ×h Ω))).
        fold y ω ω.
        use impl_intro.
        use eq_in_omega_partial_setoid.
        use iff_intro.
        + use (iff_elim_left
                 (weaken_left (weaken_right (weaken_right (hyperdoctrine_hyp _) _) _) _)).
          use (iff_elim_right
                 (weaken_left (weaken_left (weaken_right (hyperdoctrine_hyp _) _) _) _)).
          use weaken_right.
          apply hyperdoctrine_hyp.
        + use (iff_elim_left
                 (weaken_left (weaken_left (weaken_right (hyperdoctrine_hyp _) _) _) _)).
          use (iff_elim_right
                 (weaken_left (weaken_right (weaken_right (hyperdoctrine_hyp _) _) _) _)).
          use weaken_right.
          apply hyperdoctrine_hyp.
      - unfold partial_setoid_mor_hom_exists_law ; cbn.
        use forall_intro.
        use impl_intro.
        use weaken_right.
        pose (y := π (tm_var (𝟙 ×h Y))).
        fold y.
        use exists_intro.
        {
          exact (tripos_form_to_tm (h (φ [ π (tm_var _) , π (tm_var _)) ]))).
        }
        hypersimplify_form.
        rewrite partial_setoid_subst.
        unfold y.
        hypersimplify.
        rewrite tripos_form_to_tm_Prf.
        use conj_intro.
        + apply hyperdoctrine_hyp.
        + apply iff_refl.
    Qed.

    Definition subobject_classifier_partial_setoid_map
      : partial_setoid_morphism Y omega_partial_setoid.
    Proof.
      use make_partial_setoid_morphism.
      - exact subobject_classifier_partial_setoid_map_form.
      - exact subobject_classifier_partial_setoid_map_laws.
    Defined.

    Proposition subobject_classifier_partial_setoid_comm
      : partial_setoid_comp_morphism
          φ
          subobject_classifier_partial_setoid_map
        =
        partial_setoid_comp_morphism
          (partial_setoid_morphism_to_terminal X)
          omega_partial_setoid_true.
    Proof.
      use eq_partial_setoid_morphism ; cbn.
      - use (exists_elim (hyperdoctrine_hyp _)).
        use weaken_right.
        rewrite exists_subst.
        use exists_intro.
        {
          exact !!.
        }
        hypersimplify_form.
        hypersimplify.
        pose (x := π (tm_var ((X ×h Ω) ×h Y)))).
        pose (ω := π (tm_var ((X ×h Ω) ×h Y)))).
        pose (y := π (tm_var ((X ×h Ω) ×h Y))).
        fold x ω y.
        use conj_intro.
        + use (partial_setoid_mor_dom_defined φ).
          * exact y.
          * use weaken_left.
            apply hyperdoctrine_hyp.
        + use (iff_elim_left (weaken_right (weaken_right (hyperdoctrine_hyp _) _) _)).
          use exists_intro.
          * exact x.
          * hypersimplify.
            use weaken_left.
            apply hyperdoctrine_hyp.
      - use (exists_elim (hyperdoctrine_hyp _)).
        use weaken_right.
        hypersimplify.
        use (exists_elim (partial_setoid_mor_hom_exists φ (weaken_left (hyperdoctrine_hyp _) _))).
        rewrite exists_subst.
        use exists_intro.
        {
          exact (π (tm_var _)).
        }
        hypersimplify_form.
        hypersimplify.
        pose (x := π (tm_var (((X ×h Ω) ×h 𝟙) ×h Y))))).
        pose (ω := π (tm_var (((X ×h Ω) ×h 𝟙) ×h Y))))).
        pose (y := π (tm_var (((X ×h Ω) ×h 𝟙) ×h Y))).
        fold x ω y.
        repeat use conj_intro.
        + use weaken_right.
          apply hyperdoctrine_hyp.
        + use (partial_setoid_mor_cod_defined φ).
          * exact x.
          * use weaken_right.
            apply hyperdoctrine_hyp.
        + use impl_intro.
          do 2 use weaken_left.
          use weaken_right.
          apply hyperdoctrine_hyp.
        + use impl_intro.
          use exists_intro.
          * exact x.
          * hypersimplify.
            use weaken_left.
            use weaken_right.
            apply hyperdoctrine_hyp.
    Qed.

4.2. The map gives rise to a pullback square

    Section PullbackUMP.
      Context {W : partial_setoid H}
              (ψ : partial_setoid_morphism W Y)
              (ψ : partial_setoid_morphism W (eq_partial_setoid 𝟙))
              (q : partial_setoid_comp_morphism
                     ψ
                     subobject_classifier_partial_setoid_map
                   =
                   partial_setoid_comp_morphism
                     ψ
                     omega_partial_setoid_true).

      Definition is_pullback_subobject_classifier_partial_setoid_map_form
        : form (W ×h X)
        := let w := π (tm_var ((W ×h X) ×h Y))) in
           let x := π (tm_var ((W ×h X) ×h Y))) in
           let y := π (tm_var ((W ×h X) ×h Y)) in
           (h (φ [ x , y ] ψ [ w , y ])).

      Arguments is_pullback_subobject_classifier_partial_setoid_map_form /.

      Proposition is_pullback_subobject_classifier_partial_setoid_map_laws
        : partial_setoid_morphism_laws
            is_pullback_subobject_classifier_partial_setoid_map_form.
      Proof.
        repeat split.
        - unfold partial_setoid_mor_dom_defined_law ; cbn.
          do 2 use forall_intro.
          use impl_intro.
          use weaken_right.
          rewrite exists_subst.
          use (exists_elim (hyperdoctrine_hyp _)).
          use weaken_right.
          hypersimplify_form.
          hypersimplify.
          pose (w := π (tm_var (((𝟙 ×h W) ×h X) ×h Y))))).
          pose (x := π (tm_var (((𝟙 ×h W) ×h X) ×h Y)))).
          pose (y := π (tm_var (((𝟙 ×h W) ×h X) ×h Y))).
          fold w x y.
          use weaken_right.
          use (partial_setoid_mor_dom_defined ψ w y).
          apply hyperdoctrine_hyp.
        - unfold partial_setoid_mor_cod_defined_law ; cbn.
          do 2 use forall_intro.
          use impl_intro.
          use weaken_right.
          rewrite exists_subst.
          use (exists_elim (hyperdoctrine_hyp _)).
          use weaken_right.
          hypersimplify_form.
          hypersimplify.
          pose (w := π (tm_var (((𝟙 ×h W) ×h X) ×h Y))))).
          pose (x := π (tm_var (((𝟙 ×h W) ×h X) ×h Y)))).
          pose (y := π (tm_var (((𝟙 ×h W) ×h X) ×h Y))).
          fold w x y.
          use weaken_left.
          use (partial_setoid_mor_dom_defined φ x y).
          apply hyperdoctrine_hyp.
        - unfold partial_setoid_mor_eq_defined_law ; cbn.
          do 4 use forall_intro.
          use impl_intro.
          use weaken_right.
          do 2 use impl_intro.
          use hyp_sym.
          rewrite exists_subst.
          use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
          rewrite conj_subst.
          use hyp_ltrans.
          use weaken_right.
          hypersimplify_form.
          hypersimplify.
          pose (w₁ := π (tm_var (((((𝟙 ×h W) ×h W) ×h X) ×h X) ×h Y))))))).
          pose (w₂ := π (tm_var (((((𝟙 ×h W) ×h W) ×h X) ×h X) ×h Y)))))).
          pose (x₁ := π (tm_var (((((𝟙 ×h W) ×h W) ×h X) ×h X) ×h Y))))).
          pose (x₂ := π (tm_var (((((𝟙 ×h W) ×h W) ×h X) ×h X) ×h Y)))).
          pose (y := π (tm_var (((((𝟙 ×h W) ×h W) ×h X) ×h X) ×h Y))).
          fold w₁ w₂ x₁ x₂ y.
          use exists_intro.
          {
            exact y.
          }
          hypersimplify.
          fold x₂ w₂.
          use conj_intro.
          + use (partial_setoid_mor_eq_defined φ).
            * exact x₁.
            * exact y.
            * use weaken_left.
              use weaken_right.
              apply hyperdoctrine_hyp.
            * do 2 use weaken_right.
              use (partial_setoid_mor_cod_defined ψ w₁).
              apply hyperdoctrine_hyp.
            * use weaken_right.
              use weaken_left.
              apply hyperdoctrine_hyp.
          + use (partial_setoid_mor_eq_defined ψ).
            * exact w₁.
            * exact y.
            * do 2 use weaken_left.
              apply hyperdoctrine_hyp.
            * do 2 use weaken_right.
              use (partial_setoid_mor_cod_defined ψ w₁).
              apply hyperdoctrine_hyp.
            * do 2 use weaken_right.
              apply hyperdoctrine_hyp.
        - unfold partial_setoid_mor_unique_im_law ; cbn.
          rewrite !exists_subst.
          do 3 use forall_intro.
          use impl_intro.
          use weaken_right.
          use (exists_elim (hyperdoctrine_hyp _)).
          use weaken_right.
          hypersimplify_form.
          use impl_intro.
          use hyp_sym.
          use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
          rewrite conj_subst.
          use hyp_ltrans.
          use weaken_right.
          hypersimplify_form.
          hypersimplify.
          pose (w := π (tm_var (((((𝟙 ×h W) ×h X) ×h X) ×h Y) ×h Y))))))).
          pose (x₁ := π (tm_var (((((𝟙 ×h W) ×h X) ×h X) ×h Y) ×h Y)))))).
          pose (x₂ := π (tm_var (((((𝟙 ×h W) ×h X) ×h X) ×h Y) ×h Y))))).
          pose (y₁ := π (tm_var (((((𝟙 ×h W) ×h X) ×h X) ×h Y) ×h Y)))).
          pose (y₂ := π (tm_var (((((𝟙 ×h W) ×h X) ×h X) ×h Y) ×h Y))).
          fold w x₁ x₂ y₁ y₂.
          pose (Δ := (φ [ x₁, y₁ ] ψ [ w, y₁ ]) φ [ x₂, y₂ ] ψ [ w, y₂ ]).
          pose (Δ' := φ [ x₁, y₁ ] φ [ x₂, y₁ ]).
          assert (r : Δ Δ').
          {
            unfold Δ, Δ'.
            use conj_intro.
            + do 2 use weaken_left.
              apply hyperdoctrine_hyp.
            + use (partial_setoid_mor_eq_defined φ).
              * exact x₂.
              * exact y₂.
              * use weaken_right.
                use weaken_left.
                use (partial_setoid_mor_dom_defined φ x₂ y₂).
                apply hyperdoctrine_hyp.
              * use (partial_setoid_mor_unique_im ψ).
                ** exact w.
                ** do 2 use weaken_right.
                   apply hyperdoctrine_hyp.
                ** use weaken_left.
                   use weaken_right.
                   apply hyperdoctrine_hyp.
              * use weaken_right.
                use weaken_left.
                apply hyperdoctrine_hyp.
          }
          refine (hyperdoctrine_cut r _).
          unfold Δ'.
          use (partial_setoid_mono_eq φ Hφ).
          + exact y₁.
          + use weaken_left.
            apply hyperdoctrine_hyp.
          + use weaken_right.
            apply hyperdoctrine_hyp.
        - unfold partial_setoid_mor_hom_exists_law ; cbn.
          use forall_intro.
          use impl_intro.
          use weaken_right.
          pose (w := π (tm_var (𝟙 ×h W))).
          fold w.
          pose (from_eq_partial_setoid_morphism_b
                  q
                  (t₁ := w) (t₂ := tripos_form_to_tm )
                  (Δ := w ~ w)).
          cbn -[tripos_form_to_tm] in h.
          rewrite !exists_subst in h.
          use (exists_elim (h _)).
          + hypersimplify.
            use exists_intro.
            {
              exact !!.
            }
            hypersimplify.
            rewrite tripos_form_to_tm_Prf.
            use conj_intro ; [ | apply truth_intro ].
            use (exists_elim (partial_setoid_mor_hom_exists ψ (hyperdoctrine_hyp _))).
            use weaken_right.
            unfold w.
            hypersimplify.
            use (hyperdoctrine_eq_transportf _ _ (hyperdoctrine_hyp _)).
            use hyperdoctrine_eq_pair_right.
            apply hyperdoctrine_unit_tm_eq.
          + unfold w.
            hypersimplify_form.
            hypersimplify.
            rewrite <- hyperdoctrine_comp_subst.
            rewrite tripos_form_to_tm_Prf.
            hypersimplify_form.
            refine (weaken_cut _ _).
            {
              do 3 use weaken_right.
              use (iff_elim_right (hyperdoctrine_hyp _)).
              apply truth_intro.
            }
            refine (exists_elim _ _).
            {
              use weaken_right.
              apply hyperdoctrine_hyp.
            }
            rewrite !conj_subst.
            use hyp_sym.
            use hyp_rtrans.
            use weaken_left.
            do 3 use hyp_rtrans.
            use weaken_left.
            hypersimplify_form.
            hypersimplify.
            use exists_intro.
            {
              exact (π (tm_var _)).
            }
            rewrite exists_subst.
            use exists_intro.
            {
              exact (π (tm_var _))).
            }
            hypersimplify.
            clear w h.
            pose (x := π (tm_var (((𝟙 ×h W) ×h Y) ×h X))).
            pose (y := π (tm_var (((𝟙 ×h W) ×h Y) ×h X)))).
            pose (w := π (tm_var (((𝟙 ×h W) ×h Y) ×h X))))).
            fold x y w.
            use conj_intro.
            * do 3 use weaken_left.
              apply hyperdoctrine_hyp.
            * use weaken_left.
              use weaken_right.
              apply hyperdoctrine_hyp.
      Qed.

      Definition is_pullback_subobject_classifier_partial_setoid_map
        : partial_setoid_morphism W X.
      Proof.
        use make_partial_setoid_morphism.
        - exact is_pullback_subobject_classifier_partial_setoid_map_form.
        - exact is_pullback_subobject_classifier_partial_setoid_map_laws.
      Defined.

      Proposition is_pullback_subobject_classifier_partial_setoid_map_pr1
        : partial_setoid_comp_morphism
            is_pullback_subobject_classifier_partial_setoid_map
            φ
          =
          ψ.
      Proof.
        use eq_partial_setoid_morphism ; cbn.
        - use (exists_elim (hyperdoctrine_hyp _)).
          use weaken_right.
          rewrite exists_subst.
          use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
          rewrite conj_subst.
          use hyp_ltrans.
          use weaken_right.
          hypersimplify.
          pose (w := π (tm_var (((W ×h Y) ×h X) ×h Y))))).
          pose (y₁ := π (tm_var (((W ×h Y) ×h X) ×h Y))))).
          pose (x := π (tm_var (((W ×h Y) ×h X) ×h Y)))).
          pose (y₂ := π (tm_var (((W ×h Y) ×h X) ×h Y))).
          fold w x y₁ y₂.
          rewrite (hyperdoctrine_pair_eta (π (tm_var _)))).
          fold w y₁.
          use (partial_setoid_mor_eq_defined ψ).
          + exact w.
          + exact y₂.
          + use (partial_setoid_mor_dom_defined ψ).
            * exact y₂.
            * do 2 use weaken_right.
              apply hyperdoctrine_hyp.
          + use (partial_setoid_mor_unique_im φ).
            * exact x.
            * use weaken_right.
              use weaken_left.
              apply hyperdoctrine_hyp.
            * use weaken_left.
              apply hyperdoctrine_hyp.
          + do 2 use weaken_right.
            apply hyperdoctrine_hyp.
        - use (exists_elim
                 (partial_setoid_mor_hom_exists
                    is_pullback_subobject_classifier_partial_setoid_map
                    _)).
          + exact (π (tm_var _)).
          + use (partial_setoid_mor_dom_defined ψ).
            * exact (π (tm_var _)).
            * rewrite <- hyperdoctrine_pair_eta.
              hypersimplify.
              apply hyperdoctrine_hyp.
          + cbn.
            rewrite exists_subst.
            use hyp_sym.
            use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
            rewrite conj_subst.
            use hyp_ltrans.
            use weaken_right.
            hypersimplify.
            pose (w := π (tm_var (((W ×h Y) ×h X) ×h Y))))).
            pose (y₁ := π (tm_var (((W ×h Y) ×h X) ×h Y))))).
            pose (x := π (tm_var (((W ×h Y) ×h X) ×h Y)))).
            pose (y₂ := π (tm_var (((W ×h Y) ×h X) ×h Y))).
            fold w x y₁ y₂.
            use exists_intro.
            {
              exact x.
            }
            hypersimplify_form.
            assert (ψ [ w, y₁ ] φ [ x, y₂ ] ψ [ w, y₂ ] φ [ x, y₁ ]) as r.
            {
              use (partial_setoid_mor_eq_defined φ).
              * exact x.
              * exact y₂.
              * use weaken_right.
                use weaken_left.
                use (partial_setoid_mor_dom_defined φ x y₂).
                apply hyperdoctrine_hyp.
              * use partial_setoid_sym.
                use (partial_setoid_mor_unique_im ψ).
                ** exact w.
                ** use weaken_left.
                   apply hyperdoctrine_hyp.
                ** do 2 use weaken_right.
                   apply hyperdoctrine_hyp.
              * use weaken_right.
                use weaken_left.
                apply hyperdoctrine_hyp.
            }
            use conj_intro.
            * use exists_intro.
              {
                exact y₁.
              }
              hypersimplify.
              fold w.
              rewrite (hyperdoctrine_pair_eta (π (tm_var _)))).
              fold w y₁.
              use conj_intro.
              ** exact r.
              ** use weaken_left.
                 apply hyperdoctrine_hyp.
            * hypersimplify.
              fold y₁.
              rewrite (hyperdoctrine_pair_eta (π (tm_var _)))).
              fold w y₁.
              exact r.
      Qed.

      Proposition is_pullback_subobject_classifier_partial_setoid_map_pr2
        : partial_setoid_comp_morphism
            is_pullback_subobject_classifier_partial_setoid_map
            (partial_setoid_morphism_to_terminal X)
          =
          ψ.
      Proof.
        apply (TerminalArrowEq (T := terminal_partial_setoid H)).
      Qed.

      Context {ζ : partial_setoid_morphism W X}
              (ζp : partial_setoid_comp_morphism ζ φ = ψ)
              (ζq : partial_setoid_comp_morphism
                      ζ
                      (partial_setoid_morphism_to_terminal X)
                    =
                    ψ).

      Proposition is_pullback_subobject_classifier_partial_setoid_unique
        : ζ = is_pullback_subobject_classifier_partial_setoid_map.
      Proof.
        use eq_partial_setoid_morphism ; cbn.
        - use (exists_elim (partial_setoid_mor_hom_exists φ _)).
          + exact (π (tm_var _)).
          + use (partial_setoid_mor_cod_defined ζ).
            {
              exact (π (tm_var _)).
            }
            rewrite <- hyperdoctrine_pair_eta.
            hypersimplify.
            apply hyperdoctrine_hyp.
          + rewrite exists_subst.
            pose (w := π (tm_var ((W ×h X) ×h Y)))).
            pose (x := π (tm_var ((W ×h X) ×h Y)))).
            pose (y := π (tm_var ((W ×h X) ×h Y))).
            use exists_intro.
            {
              exact y.
            }
            hypersimplify.
            fold w x y.
            rewrite (hyperdoctrine_pair_eta (π (tm_var _))).
            fold w x.
            use conj_intro.
            * use weaken_right.
              apply hyperdoctrine_hyp.
            * use (from_eq_partial_setoid_morphism_f ζp) ; cbn.
              hypersimplify_form.
              use exists_intro.
              {
                exact x.
              }
              hypersimplify.
              apply hyperdoctrine_hyp.
        - use (exists_elim (hyperdoctrine_hyp _)).
          use weaken_right.
          pose (w := π (tm_var ((W ×h X) ×h Y)))).
          pose (x := π (tm_var ((W ×h X) ×h Y)))).
          pose (y := π (tm_var ((W ×h X) ×h Y))).
          fold w x y.
          rewrite (hyperdoctrine_pair_eta (π (tm_var _))).
          fold w x.
          refine (weaken_cut
                    (from_eq_partial_setoid_morphism_b
                       ζp
                       (weaken_right (hyperdoctrine_hyp _) _))
                    _).
          cbn.
          hypersimplify_form.
          use hyp_sym.
          use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
          rewrite conj_subst.
          use hyp_ltrans.
          use weaken_right.
          unfold w, x, y.
          clear w x y.
          hypersimplify.
          pose (w := π (tm_var (((W ×h X) ×h Y) ×h X))))).
          pose (x₁ := π (tm_var (((W ×h X) ×h Y) ×h X))))).
          pose (y := π (tm_var (((W ×h X) ×h Y) ×h X)))).
          pose (x₂ := π (tm_var (((W ×h X) ×h Y) ×h X))).
          fold w x₁ x₂ y.
          use (partial_setoid_mor_eq_defined ζ).
          + exact w.
          + exact x₂.
          + use weaken_left.
            use weaken_right.
            use (partial_setoid_mor_dom_defined ψ w y).
            apply hyperdoctrine_hyp.
          + use partial_setoid_sym.
            use (partial_setoid_mono_eq φ Hφ).
            * exact y.
            * do 2 use weaken_left.
              apply hyperdoctrine_hyp.
            * do 2 use weaken_right.
              apply hyperdoctrine_hyp.
          + use weaken_right.
            use weaken_left.
            apply hyperdoctrine_hyp.
      Qed.
    End PullbackUMP.

    Definition is_pullback_subobject_classifier_partial_setoid
      : isPullback
          (C := category_of_partial_setoids H)
          subobject_classifier_partial_setoid_comm.
    Proof.
      intros W ψ ψ q.
      use make_iscontr.
      - simple refine (_ ,, _ ,, _).
        + exact (is_pullback_subobject_classifier_partial_setoid_map ψ ψ q).
        + exact (is_pullback_subobject_classifier_partial_setoid_map_pr1 ψ ψ q).
        + exact (is_pullback_subobject_classifier_partial_setoid_map_pr2 ψ ψ q).
      - abstract
          (intros ζ ;
           use subtypePath ;
           [ intro ;
             apply isapropdirprod ;
             apply homset_property
           | ] ;
           induction ζ as [ ζ [ ζp ζq ]] ;
           exact (is_pullback_subobject_classifier_partial_setoid_unique _ _ _ ζp)).
    Defined.

4.3. Uniqueness

    Context (χ : partial_setoid_morphism Y omega_partial_setoid)
            (p : partial_setoid_comp_morphism φ χ
                 =
                 partial_setoid_comp_morphism
                   (partial_setoid_morphism_to_terminal X)
                   omega_partial_setoid_true)
            (Hχ : isPullback
                    (C := category_of_partial_setoids H)
                    p).

    Lemma subobject_classifier_partial_setoid_map_unique_lem
      : (χ Prf [π (tm_var (Y ×h Ω))]
         
         h (φ [ π (tm_var ((Y ×h Ω) ×h X)) , π (tm_var ((Y ×h Ω) ×h X))) ])).
    Proof.
      pose (y := π (tm_var (Y ×h Ω))).
      pose (ω := π (tm_var (Y ×h Ω))).
      fold y ω.
      assert (χ Prf [ω] y ~ y) as q₁.
      {
        use weaken_left.
        rewrite <- (hyperdoctrine_id_subst χ).
        rewrite (hyperdoctrine_pair_eta (tm_var (Y ×h omega_partial_setoid))).
        cbn.
        fold y ω.
        exact (partial_setoid_mor_dom_defined χ y ω (hyperdoctrine_hyp _)).
      }
      pose (fy := point_partial_setoid_morphism H (χ Prf [ω]) y q₁).
      assert (partial_setoid_comp_morphism fy χ
              =
              partial_setoid_comp_morphism
                (partial_setoid_morphism_to_terminal _)
                omega_partial_setoid_true)
        as feq.
      {
        use eq_partial_setoid_morphism ; cbn.
        - use (exists_elim (hyperdoctrine_hyp _)).
          rewrite conj_subst.
          use weaken_right.
          rewrite exists_subst.
          use exists_intro ; [ apply !! | ].
          unfold y, ω.
          hypersimplify_form.
          hypersimplify.
          rewrite (hyperdoctrine_pair_eta (π (tm_var (((Y ×h Ω) ×h Ω) ×h Y))))).
          refine (transportb (λ x, _ (x ~ x) _) _ _).
          {
            apply hyperdoctrine_pair_eta.
          }
          clear y ω q₁ fy.
          pose (y₁ := π (tm_var (((Y ×h Ω) ×h Ω) ×h Y))))).
          pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h Y))))).
          pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h Y)))).
          pose (y₂ := π (tm_var (((Y ×h Ω) ×h Ω) ×h Y))).
          fold y₁ y₂ ω ω.
          hypersimplify.
          use conj_intro.
          + unfold partial_setoid_formula.
            cbn.
            hypersimplify.
            repeat use conj_intro.
            * apply hyperdoctrine_refl.
            * do 3 use weaken_left.
              apply hyperdoctrine_hyp.
            * do 2 use weaken_left.
              use weaken_right.
              apply hyperdoctrine_hyp.
          + use from_eq_in_omega_partial_setoid_left.
            * exact ω.
            * use (partial_setoid_mor_unique_im χ).
              ** exact y₁.
              ** do 3 use weaken_left.
                 apply hyperdoctrine_hyp.
              ** use partial_setoid_mor_eq_defined.
                 *** exact y₂.
                 *** exact ω.
                 *** use weaken_left.
                     use weaken_right.
                     apply hyperdoctrine_hyp.
                 *** use weaken_right.
                     exact (partial_setoid_mor_cod_defined χ _ _ (hyperdoctrine_hyp _)).
                 *** use weaken_right.
                     apply hyperdoctrine_hyp.
            * do 2 use weaken_left.
              use weaken_right.
              apply hyperdoctrine_hyp.
        - use (exists_elim (hyperdoctrine_hyp _)).
          rewrite conj_subst.
          use weaken_right.
          rewrite exists_subst.
          use exists_intro.
          {
            exact (π (tm_var _)))).
          }
          unfold y, ω.
          hypersimplify_form.
          hypersimplify.
          rewrite (hyperdoctrine_pair_eta (π (tm_var (((Y ×h Ω) ×h Ω) ×h 𝟙))))).
          refine (transportb (λ x, _ ((_ [x] _) _) _) (hyperdoctrine_pair_eta _) _).
          clear y ω q₁ fy.
          pose (y := π (tm_var (((Y ×h Ω) ×h Ω) ×h 𝟙))))).
          pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h 𝟙))))).
          pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h 𝟙)))).
          pose (t := π (tm_var (((Y ×h Ω) ×h Ω) ×h 𝟙))).
          fold y t ω ω.
          unfold partial_setoid_formula.
          cbn.
          hypersimplify.
          use hyp_ltrans.
          use weaken_right.
          repeat use conj_intro.
          + do 2 use weaken_left.
            apply hyperdoctrine_hyp.
          + use weaken_left.
            use weaken_right.
            apply hyperdoctrine_hyp.
          + do 2 use weaken_left.
            exact (partial_setoid_mor_dom_defined χ _ _ (hyperdoctrine_hyp _)).
          + use (partial_setoid_mor_eq_defined χ).
            * exact y.
            * exact ω.
            * do 2 use weaken_left.
              exact (partial_setoid_mor_dom_defined χ _ _ (hyperdoctrine_hyp _)).
            * use eq_in_omega_partial_setoid.
              use iff_true_true.
              ** use weaken_left.
                 use weaken_right.
                 apply hyperdoctrine_hyp.
              ** use weaken_right.
                 apply hyperdoctrine_hyp.
            * do 2 use weaken_left.
              apply hyperdoctrine_hyp.
      }
      pose (PullbackArrow
              (make_Pullback _ Hχ)
              (formula_to_partial_setoid H (χ Prf [ω]))
              fy
              (partial_setoid_morphism_to_terminal _)
              feq)
        as g.
      cbn in g.
      refine (weaken_cut _ _).
      {
        simple refine (partial_setoid_mor_hom_exists g _).
        cbn.
        apply tm_var.
        unfold partial_setoid_formula.
        cbn.
        hypersimplify.
        repeat use conj_intro.
        + use hyperdoctrine_refl.
        + use weaken_left.
          apply hyperdoctrine_hyp.
        + use weaken_right.
          apply hyperdoctrine_hyp.
      }
      cbn.
      use hyp_sym.
      use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
      rewrite conj_subst.
      use hyp_ltrans.
      use weaken_right.
      hypersimplify.
      pose (y' := π (tm_var ((Y ×h Ω) ×h X)))).
      pose (ω' := π (tm_var ((Y ×h Ω) ×h X)))).
      pose (x' := π (tm_var ((Y ×h Ω) ×h X))).
      rewrite (hyperdoctrine_pair_eta (π (tm_var ((Y ×h Ω) ×h X)))).
      unfold y, ω.
      fold x' ω' y'.
      hypersimplify.
      pose (@from_eq_partial_setoid_morphism_b
              _ _ _ _ _
              (PullbackArrow_PullbackPr1 (make_Pullback _ Hχ) _ _ _ feq)
              _
              ((χ [ y', ω' ] Prf [ω']) g [ y', ω' ⟩, x' ])
              (π (tm_var _))
              (π (tm_var _)))).
      cbn in h.
      refine (hyperdoctrine_cut (h _) _).
      {
        unfold y, ω.
        hypersimplify_form.
        hypersimplify.
        rewrite (hyperdoctrine_pair_eta (π (tm_var ((Y ×h Ω) ×h X)))).
        fold x' y' ω'.
        hypersimplify.
        repeat use conj_intro.
        + do 2 use weaken_left.
          apply hyperdoctrine_hyp.
        + use weaken_left.
          use weaken_right.
          apply hyperdoctrine_hyp.
        + do 2 use weaken_left.
          exact (partial_setoid_mor_dom_defined χ _ _ (hyperdoctrine_hyp _)).
      }
      rewrite exists_subst.
      use (exists_elim (hyperdoctrine_hyp _)).
      use weaken_right.
      hypersimplify.
      unfold x', y'.
      hypersimplify.
      use exists_intro.
      {
        exact (π (tm_var _)).
      }
      hypersimplify.
      use weaken_right.
      apply hyperdoctrine_hyp.
    Qed.

    Proposition subobject_classifier_partial_setoid_map_unique
      : χ = subobject_classifier_partial_setoid_map.
    Proof.
      use eq_partial_setoid_morphism ; cbn.
      - pose (y := π (tm_var (Y ×h Ω))).
        pose (ω := π (tm_var (Y ×h Ω))).
        fold y ω.
        use conj_intro.
        + use (partial_setoid_mor_dom_defined χ y ω).
          unfold y, ω.
          rewrite <- hyperdoctrine_pair_eta.
          hypersimplify.
          apply hyperdoctrine_hyp.
        + use iff_intro.
          * use hyp_sym.
            use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
            rewrite conj_subst.
            use hyp_ltrans.
            use weaken_right.
            hypersimplify_form.
            unfold y, ω ; clear y ω.
            hypersimplify.
            pose (y := π (tm_var ((Y ×h Ω) ×h X)))).
            pose (ω := π (tm_var ((Y ×h Ω) ×h X)))).
            pose (x := π (tm_var ((Y ×h Ω) ×h X))).
            fold y ω x.
            rewrite (hyperdoctrine_pair_eta (π (tm_var _))).
            fold y ω.
            pose (from_eq_partial_setoid_morphism_f
                    p
                    (t₁ := x) (t₂ := ω)
                    (Δ := χ [ y, ω ] φ [ x, y ]))
              as r.
            cbn in r.
            rewrite !exists_subst in r.
            use (exists_elim (r _)) ; clear r.
            ** use exists_intro.
               {
                 exact y.
               }
               hypersimplify.
               use hyp_sym.
               apply hyperdoctrine_hyp.
            ** unfold y, ω, x ; clear x y ω.
               hypersimplify_form.
               hypersimplify.
               pose (y := π (tm_var (((Y ×h Ω) ×h X) ×h 𝟙))))).
               pose (ω := π (tm_var (((Y ×h Ω) ×h X) ×h 𝟙))))).
               pose (x := π (tm_var (((Y ×h Ω) ×h X) ×h 𝟙)))).
               fold x y ω.
               do 2 use weaken_right.
               apply hyperdoctrine_hyp.
          * unfold ω, y.
            hypersimplify.
            apply subobject_classifier_partial_setoid_map_unique_lem.
      - use (exists_elim
               (partial_setoid_mor_hom_exists
                  χ
                  (weaken_left (hyperdoctrine_hyp _) _))).
        hypersimplify_form.
        hypersimplify.
        pose (y := π (tm_var ((Y ×h Ω) ×h Ω)))).
        pose := π (tm_var ((Y ×h Ω) ×h Ω)))).
        pose := π (tm_var ((Y ×h Ω) ×h Ω))).
        cbn.
        fold y ω ω.
        rewrite (hyperdoctrine_pair_eta (π (tm_var ((Y ×h Ω) ×h Ω)))).
        fold y ω.
        use (partial_setoid_mor_eq_defined χ).
        + exact y.
        + exact ω.
        + do 2 use weaken_left.
          apply hyperdoctrine_hyp.
        + use partial_setoid_sym.
          use eq_in_omega_partial_setoid.
          use iff_intro.
          * pose (((y ~ y
                     (h (φ [ π (tm_var _) , π (tm_var _))) ]) Prf [ ω ]))
                     χ [ y, ω ]) Prf [ ω ])
              as Δ.
            pose (y ~ y
                   χ [ y, ω ]
                   (h (φ [ π (tm_var _) , π (tm_var _))) ])))
              as Δ'.
            assert (Δ Δ') as r.
            {
              unfold Δ, Δ'.
              repeat use conj_intro.
              ** do 3 use weaken_left.
                 apply hyperdoctrine_hyp.
              ** use weaken_left.
                 use weaken_right.
                 apply hyperdoctrine_hyp.
              ** use iff_elim_right.
                 *** exact (Prf [ ω ]).
                 *** do 2 use weaken_left.
                     use weaken_right.
                     apply hyperdoctrine_hyp.
                 *** use weaken_right.
                     apply hyperdoctrine_hyp.
            }
            refine (hyperdoctrine_cut r _).
            unfold Δ'.
            use hyp_rtrans.
            use hyp_sym.
            use (exists_elim (weaken_left (hyperdoctrine_hyp _) _)).
            rewrite conj_subst.
            use hyp_ltrans.
            use weaken_right.
            unfold y, ω, ω, Δ, Δ' ; clear r Δ Δ' y ω ω.
            hypersimplify_form.
            hypersimplify.
            pose (y := π (tm_var (((Y ×h Ω) ×h Ω) ×h X))))).
            pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h X))))).
            pose := π (tm_var (((Y ×h Ω) ×h Ω) ×h X)))).
            pose (x := π (tm_var (((Y ×h Ω) ×h Ω) ×h X))).
            cbn.
            fold y ω ω x.
            pose (from_eq_partial_setoid_morphism_f
                    p
                    (t₁ := x) (t₂ := ω)
                    (Δ := (y ~ y χ [ y, ω ]) φ [ x, y ]))
              as r.
            cbn in r.
            rewrite !exists_subst in r.
            use (exists_elim (r _)).
            ** use exists_intro.
               {
                 exact y.
               }
               hypersimplify.
               use hyp_ltrans.
               use weaken_right.
               use hyp_sym.
               apply hyperdoctrine_hyp.
            ** hypersimplify_form.
               hypersimplify.
               do 2 use weaken_right.
               apply hyperdoctrine_hyp.
          * refine (iff_elim_left _ _).
            {
              do 2 use weaken_left.
              use weaken_right.
              apply hyperdoctrine_hyp.
            }
            do 2 use hyp_ltrans.
            do 2 use weaken_right.
            pose (hyperdoctrine_proof_subst
                     y , ω
                    subobject_classifier_partial_setoid_map_unique_lem)
              as q.
            refine (hyperdoctrine_cut _ (hyperdoctrine_cut q _)).
            ** hypersimplify.
               apply hyperdoctrine_hyp.
            ** unfold y.
               hypersimplify.
               apply hyperdoctrine_hyp.
        + use weaken_right.
          apply hyperdoctrine_hyp.
    Qed.
  End UMP.

5. The subobject classifier of partial setoids

  Definition subobject_classifier_partial_setoid
    : subobject_classifier (terminal_partial_setoid H).
  Proof.
    use make_subobject_classifier_cat.
    - exact omega_partial_setoid.
    - exact omega_partial_setoid_true.
    - intros X Y m.
      use make_iscontr.
      + simple refine (_ ,, _ ,, _).
        * exact (subobject_classifier_partial_setoid_map m).
        * exact (subobject_classifier_partial_setoid_comm m).
        * exact (is_pullback_subobject_classifier_partial_setoid m).
      + abstract
          (intros χ ;
           use subtypePath ;
           [ intro ;
             use isaproptotal2 ;
             [ intro ; apply isaprop_isPullback
             | intros ; apply homset_property ]
           | ] ;
           induction χ as [ χ [ p Hχ ] ] ;
           cbn ;
           exact (subobject_classifier_partial_setoid_map_unique m χ p Hχ)).
  Defined.
End TriposSubobjectClassifier.

Arguments omega_per_form H /.
Arguments omega_partial_setoid_true_form H /.
Arguments subobject_classifier_partial_setoid_map_form {H X Y} m /.
Arguments is_pullback_subobject_classifier_partial_setoid_map_form {H X Y} m {W} ψ /.