Library UniMath.Bicategories.ComprehensionCat.HPropMono
Require Import UniMath.MoreFoundations.All.
Require Import UniMath.CategoryTheory.Core.Prelude.
Require Import UniMath.CategoryTheory.DisplayedCats.Core.
Require Import UniMath.CategoryTheory.DisplayedCats.Constructions.
Require Import UniMath.CategoryTheory.DisplayedCats.Examples.Sigma.
Require Import UniMath.CategoryTheory.DisplayedCats.Functors.
Require Import UniMath.CategoryTheory.DisplayedCats.Univalence.
Require Import UniMath.CategoryTheory.DisplayedCats.Fibrations.
Require Import UniMath.CategoryTheory.DisplayedCats.Total.
Require Import UniMath.CategoryTheory.DisplayedCats.Fiber.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.FiberCod.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.CodMorphisms.
Require Import UniMath.CategoryTheory.Equivalences.Core.
Require Import UniMath.CategoryTheory.Limits.Terminal.
Require Import UniMath.CategoryTheory.Limits.Pullbacks.
Require Import UniMath.CategoryTheory.Monics.
Require Import UniMath.CategoryTheory.WeakEquivalences.Core.
Require Import UniMath.CategoryTheory.WeakEquivalences.Mono.
Require Import UniMath.Bicategories.Core.Bicat.
Import Bicat.Notations.
Require Import UniMath.Bicategories.ComprehensionCat.BicatOfCompCat.
Require Import UniMath.Bicategories.ComprehensionCat.CompCatNotations.
Require Import UniMath.Bicategories.ComprehensionCat.ComprehensionEso.
Require Import UniMath.Bicategories.ComprehensionCat.DFLCompCat.
Require Import UniMath.Bicategories.ComprehensionCat.DFLCompCatNotations.
Local Open Scope cat.
Local Open Scope comp_cat.
Section MonoVSHProp.
Context {C : dfl_full_comp_cat}.
Require Import UniMath.CategoryTheory.Core.Prelude.
Require Import UniMath.CategoryTheory.DisplayedCats.Core.
Require Import UniMath.CategoryTheory.DisplayedCats.Constructions.
Require Import UniMath.CategoryTheory.DisplayedCats.Examples.Sigma.
Require Import UniMath.CategoryTheory.DisplayedCats.Functors.
Require Import UniMath.CategoryTheory.DisplayedCats.Univalence.
Require Import UniMath.CategoryTheory.DisplayedCats.Fibrations.
Require Import UniMath.CategoryTheory.DisplayedCats.Total.
Require Import UniMath.CategoryTheory.DisplayedCats.Fiber.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.FiberCod.
Require Import UniMath.CategoryTheory.DisplayedCats.Codomain.CodMorphisms.
Require Import UniMath.CategoryTheory.Equivalences.Core.
Require Import UniMath.CategoryTheory.Limits.Terminal.
Require Import UniMath.CategoryTheory.Limits.Pullbacks.
Require Import UniMath.CategoryTheory.Monics.
Require Import UniMath.CategoryTheory.WeakEquivalences.Core.
Require Import UniMath.CategoryTheory.WeakEquivalences.Mono.
Require Import UniMath.Bicategories.Core.Bicat.
Import Bicat.Notations.
Require Import UniMath.Bicategories.ComprehensionCat.BicatOfCompCat.
Require Import UniMath.Bicategories.ComprehensionCat.CompCatNotations.
Require Import UniMath.Bicategories.ComprehensionCat.ComprehensionEso.
Require Import UniMath.Bicategories.ComprehensionCat.DFLCompCat.
Require Import UniMath.Bicategories.ComprehensionCat.DFLCompCatNotations.
Local Open Scope cat.
Local Open Scope comp_cat.
Section MonoVSHProp.
Context {C : dfl_full_comp_cat}.
Definition is_hprop_ty
{Γ : C}
(A : ty Γ)
(lhs : tm ((Γ & A) & (A [[ π _ ]])) (A [[ π _ ]] [[ π _ ]])
:= comp_cat_tm_var (Γ & A) (A [[π A]]))
(rhs : tm ((Γ & A) & (A [[ π _ ]])) (A [[ π _ ]] [[ π _ ]])
:= comp_cat_tm_var Γ A [[ π _ ]]tm)
: UU
:= tm (Γ & A & (A [[ π A ]])) (dfl_ext_identity_type lhs rhs).
Proposition isaprop_is_hprop_ty
{Γ : C}
(A : ty Γ)
: isaprop (is_hprop_ty A).
Show proof.
{Γ : C}
(A : ty Γ)
(lhs : tm ((Γ & A) & (A [[ π _ ]])) (A [[ π _ ]] [[ π _ ]])
:= comp_cat_tm_var (Γ & A) (A [[π A]]))
(rhs : tm ((Γ & A) & (A [[ π _ ]])) (A [[ π _ ]] [[ π _ ]])
:= comp_cat_tm_var Γ A [[ π _ ]]tm)
: UU
:= tm (Γ & A & (A [[ π A ]])) (dfl_ext_identity_type lhs rhs).
Proposition isaprop_is_hprop_ty
{Γ : C}
(A : ty Γ)
: isaprop (is_hprop_ty A).
Show proof.
use invproofirrelevance.
intros t₁ t₂.
pose (dfl_eq_subst_equalizer_tm (identity _) _ _ (t₁ [[ _ ]]tm) (t₂ [[ _ ]]tm)) as p.
rewrite !id_sub_comp_cat_tm in p.
pose proof (maponpaths (λ z, z ↑ id_subst_ty_inv _) p) as q.
simpl in q.
rewrite !comp_coerce_comp_cat_tm in q.
refine (_ @ q @ _).
- refine (!(id_coerce_comp_cat_tm _) @ _).
apply maponpaths_2.
refine (!_).
apply (z_iso_inv_after_z_iso (id_subst_ty_iso _)).
- refine (_ @ id_coerce_comp_cat_tm _).
apply maponpaths_2.
apply (z_iso_inv_after_z_iso (id_subst_ty_iso _)).
intros t₁ t₂.
pose (dfl_eq_subst_equalizer_tm (identity _) _ _ (t₁ [[ _ ]]tm) (t₂ [[ _ ]]tm)) as p.
rewrite !id_sub_comp_cat_tm in p.
pose proof (maponpaths (λ z, z ↑ id_subst_ty_inv _) p) as q.
simpl in q.
rewrite !comp_coerce_comp_cat_tm in q.
refine (_ @ q @ _).
- refine (!(id_coerce_comp_cat_tm _) @ _).
apply maponpaths_2.
refine (!_).
apply (z_iso_inv_after_z_iso (id_subst_ty_iso _)).
- refine (_ @ id_coerce_comp_cat_tm _).
apply maponpaths_2.
apply (z_iso_inv_after_z_iso (id_subst_ty_iso _)).
The following is the displayed category of propositions. We also construct the
inclusion from this displayed category to the displayed category of all types.
This functor is fully faithful.
Definition disp_cat_hprop_ty
: disp_cat C
:= sigma_disp_cat
(disp_full_sub
(total_category (disp_cat_of_types C))
(λ ΓA, is_hprop_ty (pr2 ΓA))).
Proposition is_univalent_disp_disp_cat_hprop_ty
: is_univalent_disp disp_cat_hprop_ty.
Show proof.
Definition disp_functor_hprop_ty_to_ty
: disp_functor
(functor_identity _)
disp_cat_hprop_ty
(disp_cat_of_types C)
:= sigmapr1_disp_functor _.
Proposition disp_functor_hprop_ty_to_ty_ff
: disp_functor_ff disp_functor_hprop_ty_to_ty.
Show proof.
: disp_cat C
:= sigma_disp_cat
(disp_full_sub
(total_category (disp_cat_of_types C))
(λ ΓA, is_hprop_ty (pr2 ΓA))).
Proposition is_univalent_disp_disp_cat_hprop_ty
: is_univalent_disp disp_cat_hprop_ty.
Show proof.
use is_univalent_sigma_disp.
- apply disp_univalent_category_is_univalent_disp.
- use disp_full_sub_univalent.
intro ; simpl.
apply isaprop_is_hprop_ty.
- apply disp_univalent_category_is_univalent_disp.
- use disp_full_sub_univalent.
intro ; simpl.
apply isaprop_is_hprop_ty.
Definition disp_functor_hprop_ty_to_ty
: disp_functor
(functor_identity _)
disp_cat_hprop_ty
(disp_cat_of_types C)
:= sigmapr1_disp_functor _.
Proposition disp_functor_hprop_ty_to_ty_ff
: disp_functor_ff disp_functor_hprop_ty_to_ty.
Show proof.
intros Γ Δ A B s.
use isweq_iso.
- exact (λ ff, ff ,, tt).
- abstract
(simpl ;
intro sff ;
induction sff as [ sf x ] ;
induction x ;
apply idpath).
- abstract
(intro sff ; simpl ;
apply idpath).
use isweq_iso.
- exact (λ ff, ff ,, tt).
- abstract
(simpl ;
intro sff ;
induction sff as [ sf x ] ;
induction x ;
apply idpath).
- abstract
(intro sff ; simpl ;
apply idpath).
2. A type `A` is a proposition iff `π A` is a monomorphism
Proposition mono_ty_to_hprop_ty
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
: is_hprop_ty A.
Show proof.
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
: is_hprop_ty A.
Show proof.
unfold is_hprop_ty.
use dfl_ext_identity_type_tm.
use (MorphismsIntoPullbackEqual (isPullback_Pullback (comp_cat_pullback _ _))).
- rewrite !assoc'.
etrans.
{
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
refine (!_).
etrans.
{
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
use (MorphismsIntoPullbackEqual (isPullback_Pullback (comp_cat_pullback _ _))).
+ rewrite !assoc'.
etrans.
{
do 2 apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
use HA.
rewrite !assoc'.
refine (!_).
etrans.
{
apply maponpaths.
apply PullbackSqrCommutes.
}
apply idpath.
+ rewrite !assoc'.
etrans.
{
do 2 apply maponpaths.
apply comp_cat_tm_eq.
}
rewrite id_right.
apply idpath.
- rewrite !assoc'.
etrans.
{
apply maponpaths.
apply comp_cat_tm_eq.
}
refine (!_).
etrans.
{
apply maponpaths.
apply comp_cat_tm_eq.
}
apply idpath.
use dfl_ext_identity_type_tm.
use (MorphismsIntoPullbackEqual (isPullback_Pullback (comp_cat_pullback _ _))).
- rewrite !assoc'.
etrans.
{
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
refine (!_).
etrans.
{
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
use (MorphismsIntoPullbackEqual (isPullback_Pullback (comp_cat_pullback _ _))).
+ rewrite !assoc'.
etrans.
{
do 2 apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
use HA.
rewrite !assoc'.
refine (!_).
etrans.
{
apply maponpaths.
apply PullbackSqrCommutes.
}
apply idpath.
+ rewrite !assoc'.
etrans.
{
do 2 apply maponpaths.
apply comp_cat_tm_eq.
}
rewrite id_right.
apply idpath.
- rewrite !assoc'.
etrans.
{
apply maponpaths.
apply comp_cat_tm_eq.
}
refine (!_).
etrans.
{
apply maponpaths.
apply comp_cat_tm_eq.
}
apply idpath.
Next we show that the projection of every proposition is a monomorphism, and
here we use that identity types are extensional. From extensionality, we obtain
an equality between two terms (that are variables). We perform a suitable
substitution, and that gives us the desired equality.
Section HPropToMono.
Context {Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
{Δ : C}
{t₁ t₂ : Δ --> Γ & A}
(p : t₁ · π A = t₂ · π A).
Definition hprop_ty_to_mono_ty_subst
: Δ --> Γ & A & (A [[ π _ ]]).
Show proof.
Proposition hprop_ty_to_mono_ty_eq
: t₁ = t₂.
Show proof.
Proposition hprop_ty_to_mono_ty
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: isMonic (π A).
Show proof.
Definition hprop_ty_to_monic
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: Monic _ (Γ & A) Γ.
Show proof.
Definition is_hprop_ty_weq_mono_ty
{Γ : C}
(A : ty Γ)
: is_hprop_ty A ≃ isMonic (π A).
Show proof.
Context {Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
{Δ : C}
{t₁ t₂ : Δ --> Γ & A}
(p : t₁ · π A = t₂ · π A).
Definition hprop_ty_to_mono_ty_subst
: Δ --> Γ & A & (A [[ π _ ]]).
Show proof.
Proposition hprop_ty_to_mono_ty_eq
: t₁ = t₂.
Show proof.
pose proof (maponpaths
(λ z, hprop_ty_to_mono_ty_subst
· pr1 z
· PullbackPr1 (comp_cat_pullback _ _)
· PullbackPr1 (comp_cat_pullback _ _))
(tm_ext_identity_to_eq HA))
as q.
simpl in q.
refine (_ @ q @ _) ; clear q.
- refine (!_).
etrans.
{
rewrite !assoc'.
apply maponpaths.
rewrite !assoc.
rewrite PullbackArrow_PullbackPr1.
apply id_left.
}
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
- etrans.
{
rewrite !assoc'.
apply maponpaths.
rewrite !assoc.
rewrite PullbackArrow_PullbackPr1.
rewrite !assoc'.
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _)).
End HPropToMono.(λ z, hprop_ty_to_mono_ty_subst
· pr1 z
· PullbackPr1 (comp_cat_pullback _ _)
· PullbackPr1 (comp_cat_pullback _ _))
(tm_ext_identity_to_eq HA))
as q.
simpl in q.
refine (_ @ q @ _) ; clear q.
- refine (!_).
etrans.
{
rewrite !assoc'.
apply maponpaths.
rewrite !assoc.
rewrite PullbackArrow_PullbackPr1.
apply id_left.
}
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
- etrans.
{
rewrite !assoc'.
apply maponpaths.
rewrite !assoc.
rewrite PullbackArrow_PullbackPr1.
rewrite !assoc'.
apply maponpaths.
apply (PullbackArrow_PullbackPr1 (comp_cat_pullback _ _)).
}
rewrite id_right.
apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _)).
Proposition hprop_ty_to_mono_ty
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: isMonic (π A).
Show proof.
Definition hprop_ty_to_monic
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: Monic _ (Γ & A) Γ.
Show proof.
Definition is_hprop_ty_weq_mono_ty
{Γ : C}
(A : ty Γ)
: is_hprop_ty A ≃ isMonic (π A).
Show proof.
use weqimplimpl.
- exact hprop_ty_to_mono_ty.
- exact mono_ty_to_hprop_ty.
- apply isaprop_is_hprop_ty.
- apply isapropisMonic.
- exact hprop_ty_to_mono_ty.
- exact mono_ty_to_hprop_ty.
- apply isaprop_is_hprop_ty.
- apply isapropisMonic.
Now we conclude that all displayed morphisms in the displayed category of
propositions are equal.
Proposition locally_propositional_disp_cat_hprop_ty
: locally_propositional disp_cat_hprop_ty.
Show proof.
: locally_propositional disp_cat_hprop_ty.
Show proof.
intros Γ Δ s A B.
use invproofirrelevance.
intros ff gg.
use subtypePath.
{
intro.
apply isapropunit.
}
use (invmaponpathsweq
(disp_functor_ff_weq
_
(full_comp_cat_comprehension_fully_faithful C)
_ _ _)).
use subtypePath.
{
intro.
apply homset_property.
}
use (hprop_ty_to_mono_ty (pr2 B)).
simpl.
etrans.
{
apply comprehension_functor_mor_comm.
}
refine (!_).
apply comprehension_functor_mor_comm.
use invproofirrelevance.
intros ff gg.
use subtypePath.
{
intro.
apply isapropunit.
}
use (invmaponpathsweq
(disp_functor_ff_weq
_
(full_comp_cat_comprehension_fully_faithful C)
_ _ _)).
use subtypePath.
{
intro.
apply homset_property.
}
use (hprop_ty_to_mono_ty (pr2 B)).
simpl.
etrans.
{
apply comprehension_functor_mor_comm.
}
refine (!_).
apply comprehension_functor_mor_comm.
3. A type `A` is a proposition iff all terms of type `A` in every context are equal
Section UniqueTermsToMono.
Context {Γ : C}
{A : ty Γ}
(HA : ∏ (Δ : C) (s : Δ --> Γ), isaprop (tm Δ (A [[ s ]])))
{Δ : C}
{t₁ t₂ : Δ --> Γ & A}
(p : t₁ · π A = t₂ · π A).
Definition all_terms_eq_to_mono_type_lhs
: tm Δ (A [[ t₁ · π A ]]).
Show proof.
Definition all_terms_eq_to_mono_type_rhs
: tm Δ (A [[ t₁ · π A ]]).
Show proof.
Proposition all_terms_eq_to_mono_type_lhs_rhs_eq
: all_terms_eq_to_mono_type_lhs
=
all_terms_eq_to_mono_type_rhs.
Show proof.
Proposition all_terms_eq_to_mono_type_eq
: t₁ = t₂.
Show proof.
Context {Γ : C}
{A : ty Γ}
(HA : ∏ (Δ : C) (s : Δ --> Γ), isaprop (tm Δ (A [[ s ]])))
{Δ : C}
{t₁ t₂ : Δ --> Γ & A}
(p : t₁ · π A = t₂ · π A).
Definition all_terms_eq_to_mono_type_lhs
: tm Δ (A [[ t₁ · π A ]]).
Show proof.
use make_comp_cat_tm.
- use (PullbackArrow (comp_cat_pullback _ _)).
+ exact t₁.
+ apply identity.
+ abstract
(rewrite id_left ;
apply idpath).
- abstract
(apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _))).
- use (PullbackArrow (comp_cat_pullback _ _)).
+ exact t₁.
+ apply identity.
+ abstract
(rewrite id_left ;
apply idpath).
- abstract
(apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _))).
Definition all_terms_eq_to_mono_type_rhs
: tm Δ (A [[ t₁ · π A ]]).
Show proof.
use make_comp_cat_tm.
- use (PullbackArrow (comp_cat_pullback _ _)).
+ exact t₂.
+ apply identity.
+ abstract
(rewrite id_left ;
rewrite p ;
apply idpath).
- abstract
(apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _))).
- use (PullbackArrow (comp_cat_pullback _ _)).
+ exact t₂.
+ apply identity.
+ abstract
(rewrite id_left ;
rewrite p ;
apply idpath).
- abstract
(apply (PullbackArrow_PullbackPr2 (comp_cat_pullback _ _))).
Proposition all_terms_eq_to_mono_type_lhs_rhs_eq
: all_terms_eq_to_mono_type_lhs
=
all_terms_eq_to_mono_type_rhs.
Show proof.
exact (proofirrelevance
_
(HA Δ (t₁ · π A))
all_terms_eq_to_mono_type_lhs
all_terms_eq_to_mono_type_rhs).
_
(HA Δ (t₁ · π A))
all_terms_eq_to_mono_type_lhs
all_terms_eq_to_mono_type_rhs).
Proposition all_terms_eq_to_mono_type_eq
: t₁ = t₂.
Show proof.
pose (maponpaths
(λ z, pr1 z · PullbackPr1 (comp_cat_pullback _ _))
all_terms_eq_to_mono_type_lhs_rhs_eq)
as q.
simpl in q.
refine (_ @ q @ _).
- rewrite PullbackArrow_PullbackPr1.
apply idpath.
- rewrite PullbackArrow_PullbackPr1.
apply idpath.
End UniqueTermsToMono.(λ z, pr1 z · PullbackPr1 (comp_cat_pullback _ _))
all_terms_eq_to_mono_type_lhs_rhs_eq)
as q.
simpl in q.
refine (_ @ q @ _).
- rewrite PullbackArrow_PullbackPr1.
apply idpath.
- rewrite PullbackArrow_PullbackPr1.
apply idpath.
To prove the reverse, we first show that all terms in a propositon are equal. Then
it suffices to prove that the type `A [ s ]` is a proposition, so it is enough
to show that the projection of `A [ s ]` is a monomorphism. Here we use that
monomorphisms are closed under pullback
Proposition isaprop_tm_hprop_ty
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: isaprop (tm Γ A).
Show proof.
Proposition mono_ty_subst_all_terms_eq
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
{Δ : C}
(s : Δ --> Γ)
: isaprop (tm Δ (A [[ s ]])).
Show proof.
Definition mono_ty_weq_all_terms_eq
{Γ : C}
{A : ty Γ}
: isMonic (π A) ≃ ∏ (Δ : C) (s : Δ --> Γ), isaprop (tm Δ (A [[ s ]])).
Show proof.
{Γ : C}
{A : ty Γ}
(HA : is_hprop_ty A)
: isaprop (tm Γ A).
Show proof.
use invproofirrelevance.
intros t₁ t₂.
use eq_comp_cat_tm.
use (hprop_ty_to_mono_ty HA).
rewrite !comp_cat_tm_eq.
apply idpath.
intros t₁ t₂.
use eq_comp_cat_tm.
use (hprop_ty_to_mono_ty HA).
rewrite !comp_cat_tm_eq.
apply idpath.
Proposition mono_ty_subst_all_terms_eq
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
{Δ : C}
(s : Δ --> Γ)
: isaprop (tm Δ (A [[ s ]])).
Show proof.
use isaprop_tm_hprop_ty.
use mono_ty_to_hprop_ty.
pose (m := (π A ,, HA) : Monic _ _ _).
exact (MonicPullbackisMonic _ m s (comp_cat_pullback A s)).
use mono_ty_to_hprop_ty.
pose (m := (π A ,, HA) : Monic _ _ _).
exact (MonicPullbackisMonic _ m s (comp_cat_pullback A s)).
Definition mono_ty_weq_all_terms_eq
{Γ : C}
{A : ty Γ}
: isMonic (π A) ≃ ∏ (Δ : C) (s : Δ --> Γ), isaprop (tm Δ (A [[ s ]])).
Show proof.
use weqimplimpl.
- exact mono_ty_subst_all_terms_eq.
- intros HA Δ t₁ t₂ p.
exact (all_terms_eq_to_mono_type_eq HA p).
- apply isapropisMonic.
- do 2 (use impred ; intro).
apply isapropisaprop.
- exact mono_ty_subst_all_terms_eq.
- intros HA Δ t₁ t₂ p.
exact (all_terms_eq_to_mono_type_eq HA p).
- apply isapropisMonic.
- do 2 (use impred ; intro).
apply isapropisaprop.
Proposition subsingleton_to_mono_ty
{Γ : C}
{A : ty Γ}
(HA : isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A))
: isMonic (π A).
Show proof.
Proposition mono_ty_to_subsingleton
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
: isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A).
Show proof.
Definition subsingleton_weq_mono_ty
{Γ : C}
(A : ty Γ)
: isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A)
≃
isMonic (π A).
Show proof.
{Γ : C}
{A : ty Γ}
(HA : isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A))
: isMonic (π A).
Show proof.
rewrite <- (comp_cat_comp_mor_comm
(TerminalArrow (dfl_full_comp_cat_terminal Γ)
A)).
use isMonic_comp.
- use from_is_monic_cod_fib.
pose (m := (_ ,, HA) : Monic _ _ _).
refine (weak_equiv_preserves_monic
(F := fiber_functor (comp_cat_comprehension C) Γ)
_
m).
use weak_equiv_from_equiv.
exact (fiber_functor_comprehension_adj_equiv _ Γ).
- apply is_iso_isMonic.
apply dfl_full_comp_cat_extend_unit.
(TerminalArrow (dfl_full_comp_cat_terminal Γ)
A)).
use isMonic_comp.
- use from_is_monic_cod_fib.
pose (m := (_ ,, HA) : Monic _ _ _).
refine (weak_equiv_preserves_monic
(F := fiber_functor (comp_cat_comprehension C) Γ)
_
m).
use weak_equiv_from_equiv.
exact (fiber_functor_comprehension_adj_equiv _ Γ).
- apply is_iso_isMonic.
apply dfl_full_comp_cat_extend_unit.
Proposition mono_ty_to_subsingleton
{Γ : C}
{A : ty Γ}
(HA : isMonic (π A))
: isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A).
Show proof.
intros A' t₁ t₂ p.
use (invmaponpathsweq
(disp_functor_ff_weq
_
(full_comp_cat_comprehension_fully_faithful C)
_ _ _)).
use eq_mor_cod_fib ; simpl.
use HA.
exact (mor_eq _ @ !(mor_eq _)).
use (invmaponpathsweq
(disp_functor_ff_weq
_
(full_comp_cat_comprehension_fully_faithful C)
_ _ _)).
use eq_mor_cod_fib ; simpl.
use HA.
exact (mor_eq _ @ !(mor_eq _)).
Definition subsingleton_weq_mono_ty
{Γ : C}
(A : ty Γ)
: isMonic (TerminalArrow (dfl_full_comp_cat_terminal Γ) A)
≃
isMonic (π A).
Show proof.
use weqimplimpl.
- exact subsingleton_to_mono_ty.
- exact mono_ty_to_subsingleton.
- apply isapropisMonic.
- apply isapropisMonic.
- exact subsingleton_to_mono_ty.
- exact mono_ty_to_subsingleton.
- apply isapropisMonic.
- apply isapropisMonic.
Proposition is_hprop_ty_subst
{Γ Δ : C}
(s : Γ --> Δ)
{A : ty Δ}
(HA : is_hprop_ty A)
: is_hprop_ty (A [[ s ]]).
Show proof.
{Γ Δ : C}
(s : Γ --> Δ)
{A : ty Δ}
(HA : is_hprop_ty A)
: is_hprop_ty (A [[ s ]]).
Show proof.
use mono_ty_to_hprop_ty.
exact (MonicPullbackisMonic _ (hprop_ty_to_monic HA) _ (comp_cat_pullback A s)).
exact (MonicPullbackisMonic _ (hprop_ty_to_monic HA) _ (comp_cat_pullback A s)).
Proposition is_hprop_ty_unit_type
(Γ : C)
: is_hprop_ty (dfl_full_comp_cat_unit Γ).
Show proof.
Proposition is_hprop_ty_dfl_ext_identity_type
{Γ : C}
{A : ty Γ}
(t₁ t₂ : tm Γ A)
: is_hprop_ty (dfl_ext_identity_type t₁ t₂).
Show proof.
(Γ : C)
: is_hprop_ty (dfl_full_comp_cat_unit Γ).
Show proof.
Proposition is_hprop_ty_dfl_ext_identity_type
{Γ : C}
{A : ty Γ}
(t₁ t₂ : tm Γ A)
: is_hprop_ty (dfl_ext_identity_type t₁ t₂).
Show proof.
use mono_ty_to_hprop_ty.
use all_terms_eq_to_mono_type_eq.
intros Δ s.
use invproofirrelevance.
intros p₁ p₂.
apply dfl_eq_subst_equalizer_tm.
End MonoVSHProp.use all_terms_eq_to_mono_type_eq.
intros Δ s.
use invproofirrelevance.
intros p₁ p₂.
apply dfl_eq_subst_equalizer_tm.