Martin Escardo and Tom de Jong 7th - 21st November 2024.
We improve the universe levels of Lemma 14 of [1], which corresponds
to `embedding-retract` from InjectiveTypes/Blackboard.
We use this to remove resizing assumption of Theorem 51 of [1], which
characterizes the algebraically injective types as the retracts of the
algebras of the lifiting monad (also known as the partial-map
classifier monad.
[1] M.H. EscardΓ³. Injective type in univalent mathematics.
https://doi.org/10.1017/S0960129520000225
\begin{code}
{-# OPTIONS --safe --without-K #-}
open import UF.FunExt
module InjectiveTypes.CharacterizationViaLifting (fe : FunExt) where
open import InjectiveTypes.Blackboard fe
open import InjectiveTypes.OverSmallMaps fe
open import MLTT.Spartan
open import UF.Equiv
open import UF.Size
open import UF.Subsingletons
private
fe' : Fun-Ext
fe' {π€} {π₯} = fe π€ π₯
\end{code}
We first improve the universe levels of Blackboard.ainjectivity-of-Lifting.
\begin{code}
open import UF.Univalence
module ainjectivity-of-Lifting'
(π£ : Universe)
(ua : is-univalent π£)
where
private
pe : propext π£
pe = univalence-gives-propext ua
open ainjectivity-of-Lifting π£
open import Lifting.UnivalentWildCategory π£
open import UF.Retracts
Ξ·-is-small-map : {X : π€ Μ } β (Ξ· βΆ (X β π X)) is π£ small-map
Ξ·-is-small-map {π€} {X} l = is-defined l ,
β-sym (Ξ·-fiber-same-as-is-defined X pe fe' fe' fe' l)
\end{code}
The following improves the universe levels of Lemma 50 of [1].
\begin{code}
ainjective-is-retract-of-free-π-algebra' : ({π€} π₯ {π¦} : Universe)
(D : π€ Μ )
β ainjective-type D (π£ β π₯) π¦
β retract D of (π D)
ainjective-is-retract-of-free-π-algebra' {π€} π₯ D =
embedding-retract' π₯ D (π D) Ξ·
(Ξ·-is-embedding' π€ D ua fe')
Ξ·-is-small-map
ainjectives-in-terms-of-free-π-algebras'
: (D : π€ Μ ) β ainjective-type D π£ π£ β (Ξ£ X κ π€ Μ , retract D of (π X))
ainjectives-in-terms-of-free-π-algebras' {π€} D = a , b
where
a : ainjective-type D π£ π£ β Ξ£ X κ π€ Μ , retract D of (π X)
a i = D , ainjective-is-retract-of-free-π-algebra' π£ D i
b : (Ξ£ X κ π€ Μ , retract D of (π X)) β ainjective-type D π£ π£
b (X , r) = retract-of-ainjective D (π X) (free-π-algebra-ainjective ua X) r
\end{code}
A particular case of interest that arises in practice is the following.
\begin{code}
ainjectives-in-terms-of-free-π-algebrasβΊ
: (D : π£ βΊ Μ ) β ainjective-type D π£ π£ β (Ξ£ X κ π£ βΊ Μ , retract D of (π X))
ainjectives-in-terms-of-free-π-algebrasβΊ
= ainjectives-in-terms-of-free-π-algebras'
_ : {X : π£ βΊ Μ } β type-of (π X) οΌ π£ βΊ Μ
_ = refl
\end{code}
The following removes the resizing assumption of Theorem 51 of [1].
\begin{code}
ainjectives-in-terms-of-π-algebras
: (D : π€ Μ ) β ainjective-type D π£ π£ β (Ξ£ A κ π£ βΊ β π€ Μ , π-alg A Γ retract D of A)
ainjectives-in-terms-of-π-algebras {π€} D = a , b
where
a : ainjective-type D π£ π£ β (Ξ£ A κ π£ βΊ β π€ Μ , π-alg A Γ retract D of A)
a i = π D ,
π-algebra-gives-alg (free-π-algebra ua D) ,
ainjective-is-retract-of-free-π-algebra' π£ D i
b : (Ξ£ A κ π£ βΊ β π€ Μ , π-alg A Γ retract D of A) β ainjective-type D π£ π£
b (A , Ξ± , Ο) = retract-of-ainjective D A (π-alg-ainjective pe A Ξ±) Ο
\end{code}
Particular case of interest:
\begin{code}
ainjectives-in-terms-of-π-algebrasβΊ
: (D : π£ βΊ Μ ) β ainjective-type D π£ π£ β (Ξ£ A κ π£ βΊ Μ , π-alg A Γ retract D of A)
ainjectives-in-terms-of-π-algebrasβΊ = ainjectives-in-terms-of-π-algebras
\end{code}