Library UniMath.CategoryTheory.EnrichedCats.RezkCompletion.EnrichedRezkCompletion
Require Import UniMath.Foundations.All.
Require Import UniMath.MoreFoundations.All.
Require Import UniMath.CategoryTheory.Core.Prelude.
Require Import UniMath.CategoryTheory.opp_precat.
Require Import UniMath.CategoryTheory.OppositeCategory.Core.
Require Import UniMath.CategoryTheory.Subcategory.Core.
Require Import UniMath.CategoryTheory.Subcategory.Full.
Require Import UniMath.CategoryTheory.EnrichedCats.Enrichment.
Require Import UniMath.CategoryTheory.EnrichedCats.EnrichmentFunctor.
Require Import UniMath.CategoryTheory.EnrichedCats.EnrichmentTransformation.
Require Import UniMath.CategoryTheory.EnrichedCats.FullyFaithful.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.OppositeEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.FunctorCategory.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.SelfEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.ImageEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.Yoneda.
Require Import UniMath.CategoryTheory.EnrichedCats.YonedaLemma.
Require Import UniMath.CategoryTheory.Monoidal.Categories.
Require Import UniMath.CategoryTheory.Monoidal.Structure.Symmetric.
Require Import UniMath.CategoryTheory.Monoidal.Structure.Closed.
Require Import UniMath.CategoryTheory.limits.equalizers.
Require Import UniMath.CategoryTheory.limits.products.
Local Open Scope cat.
Local Open Scope moncat.
Section EnrichedRezkCompletion.
Context {V : sym_mon_closed_cat}
{C : category}
(E : enrichment C V)
(EqV : Equalizers V)
(PV : Products C V)
(PV' : Products (C × C) V)
(HV : is_univalent V).
Require Import UniMath.MoreFoundations.All.
Require Import UniMath.CategoryTheory.Core.Prelude.
Require Import UniMath.CategoryTheory.opp_precat.
Require Import UniMath.CategoryTheory.OppositeCategory.Core.
Require Import UniMath.CategoryTheory.Subcategory.Core.
Require Import UniMath.CategoryTheory.Subcategory.Full.
Require Import UniMath.CategoryTheory.EnrichedCats.Enrichment.
Require Import UniMath.CategoryTheory.EnrichedCats.EnrichmentFunctor.
Require Import UniMath.CategoryTheory.EnrichedCats.EnrichmentTransformation.
Require Import UniMath.CategoryTheory.EnrichedCats.FullyFaithful.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.OppositeEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.FunctorCategory.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.SelfEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.ImageEnriched.
Require Import UniMath.CategoryTheory.EnrichedCats.Examples.Yoneda.
Require Import UniMath.CategoryTheory.EnrichedCats.YonedaLemma.
Require Import UniMath.CategoryTheory.Monoidal.Categories.
Require Import UniMath.CategoryTheory.Monoidal.Structure.Symmetric.
Require Import UniMath.CategoryTheory.Monoidal.Structure.Closed.
Require Import UniMath.CategoryTheory.limits.equalizers.
Require Import UniMath.CategoryTheory.limits.products.
Local Open Scope cat.
Local Open Scope moncat.
Section EnrichedRezkCompletion.
Context {V : sym_mon_closed_cat}
{C : category}
(E : enrichment C V)
(EqV : Equalizers V)
(PV : Products C V)
(PV' : Products (C × C) V)
(HV : is_univalent V).
Definition enriched_rezk_completion
: univalent_category.
Show proof.
Definition enriched_rezk_completion_enrichment
: enrichment enriched_rezk_completion V.
Show proof.
: univalent_category.
Show proof.
use make_univalent_category.
- exact (full_img_sub_precategory (enriched_yoneda_functor E)).
- use is_univalent_full_sub_category.
use is_univalent_enriched_functor_cat.
exact HV.
- exact (full_img_sub_precategory (enriched_yoneda_functor E)).
- use is_univalent_full_sub_category.
use is_univalent_enriched_functor_cat.
exact HV.
Definition enriched_rezk_completion_enrichment
: enrichment enriched_rezk_completion V.
Show proof.
Definition enriched_rezk_completion_map
: C ⟶ enriched_rezk_completion
:= functor_full_img _.
Definition enriched_rezk_completion_map_enrichment
: functor_enrichment
enriched_rezk_completion_map
E
enriched_rezk_completion_enrichment.
Show proof.
Proposition is_essentially_surjective_enriched_rezk_completion_map
: essentially_surjective enriched_rezk_completion_map.
Show proof.
Proposition is_fully_faithful_enriched_rezk_completion_map
: fully_faithful_enriched_functor enriched_rezk_completion_map_enrichment.
Show proof.
Proposition enriched_weak_equivalence_enriched_rezk_completion_map
: enriched_weak_equivalence enriched_rezk_completion_map_enrichment.
Show proof.
: C ⟶ enriched_rezk_completion
:= functor_full_img _.
Definition enriched_rezk_completion_map_enrichment
: functor_enrichment
enriched_rezk_completion_map
E
enriched_rezk_completion_enrichment.
Show proof.
Proposition is_essentially_surjective_enriched_rezk_completion_map
: essentially_surjective enriched_rezk_completion_map.
Show proof.
Proposition is_fully_faithful_enriched_rezk_completion_map
: fully_faithful_enriched_functor enriched_rezk_completion_map_enrichment.
Show proof.
exact (fully_faithful_enriched_factorization_precomp
(enriched_yoneda_enrichment E EqV PV PV')
enriched_rezk_completion_map_enrichment
_
(image_factorization_enriched_commutes_inv _)
(fully_faithful_enriched_yoneda _ _ _ _)
image_incl_enrichment_fully_faithful).
(enriched_yoneda_enrichment E EqV PV PV')
enriched_rezk_completion_map_enrichment
_
(image_factorization_enriched_commutes_inv _)
(fully_faithful_enriched_yoneda _ _ _ _)
image_incl_enrichment_fully_faithful).
Proposition enriched_weak_equivalence_enriched_rezk_completion_map
: enriched_weak_equivalence enriched_rezk_completion_map_enrichment.
Show proof.
split.
- exact is_essentially_surjective_enriched_rezk_completion_map.
- exact is_fully_faithful_enriched_rezk_completion_map.
- exact is_essentially_surjective_enriched_rezk_completion_map.
- exact is_fully_faithful_enriched_rezk_completion_map.
Definition enriched_rezk_completion_bundled
: ∑ (RC : univalent_category)
(ERC : enrichment RC V)
(F : C ⟶ RC)
(EF : functor_enrichment F E ERC),
enriched_weak_equivalence EF
:= enriched_rezk_completion
,,
enriched_rezk_completion_enrichment
,,
enriched_rezk_completion_map
,,
enriched_rezk_completion_map_enrichment
,,
enriched_weak_equivalence_enriched_rezk_completion_map.
End EnrichedRezkCompletion.
: ∑ (RC : univalent_category)
(ERC : enrichment RC V)
(F : C ⟶ RC)
(EF : functor_enrichment F E ERC),
enriched_weak_equivalence EF
:= enriched_rezk_completion
,,
enriched_rezk_completion_enrichment
,,
enriched_rezk_completion_map
,,
enriched_rezk_completion_map_enrichment
,,
enriched_weak_equivalence_enriched_rezk_completion_map.
End EnrichedRezkCompletion.