Martin Escardo, 23rd August 2023.
Some counterexamples to injectivity.
We already know that if excluded middle holds then all pointed types
are algebraically injective, and that the converse also holds.
So we can't really give an example of any type which is not
algebraically injective, other than the empty type. The best we can
hope for is to derive a constructive taboo, rather than a
contradiction, from the assumption that a type of interest would be
injective.
Most types one encounters in practice are "not" injective in the above
sense. We can also say "not all types are injective in general", as
there are some β-toposes which do satisfy excluded middle, as well as
some β-toposes which don't, and we intend TypeTopology to apply to all
β-toposes, except when special assumptions are made.
NB. We work with the assumption of algebraic injectivity, rather than
its truncated version (injectivity), but this doesn't matter because
most of our conclusions are propositions, and when they are not we can
consider their truncations, which are also constructive taboos.
More counter-examples are in the module InjectiveTypes.Resizing and in
the module InjectiveTypes.InhabitedTypesTaboo.
\begin{code}
{-# OPTIONS --safe --without-K --lossy-unification #-}
open import UF.Univalence
open import UF.PropTrunc
module InjectiveTypes.CounterExamples
(ua : Univalence)
(pt : propositional-truncations-exist)
where
open PropositionalTruncation pt
open import MLTT.Spartan
open import UF.Equiv
open import UF.Embeddings
open import UF.ClassicalLogic
open import UF.FunExt
open import UF.Retracts
open import UF.Size
open import UF.SubtypeClassifier
open import UF.SubtypeClassifier-Properties
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import UF.UA-FunExt
private
fe : FunExt
fe = Univalence-gives-FunExt ua
fe' : Fun-Ext
fe' {π€} {π₯} = fe π€ π₯
pe : PropExt
pe = Univalence-gives-PropExt ua
pe' : Prop-Ext
pe' {π€} = pe π€
open import Taboos.Decomposability fe
open import InjectiveTypes.Blackboard fe
open import TypeTopology.SimpleTypes fe pt
\end{code}
The two-point type π is not injective in general. It is algebraically
injective if and only if weak excluded middle holds.
\begin{code}
π-ainjective-gives-WEM : ainjective-type π π€ π₯ β typal-WEM π€
π-ainjective-gives-WEM {π€} π-ainj = I
where
d : decomposition π
d = id , (β , refl) , (β , refl)
I : typal-WEM π€
I = decomposition-of-ainjective-type-gives-WEM pe' π π-ainj d
WEM-gives-π-retract-of-Ξ© : typal-WEM π€ β retract π of Ξ© π€
WEM-gives-π-retract-of-Ξ© {π€} wem = II
where
h : (p : Ξ© π€) β is-decidable (Β¬ (p holds)) β π
h p (inl _) = β
h p (inr _) = β
Ξ©-to-π : Ξ© π€ β π
Ξ©-to-π p = h p (wem (p holds))
I : (n : π) (d : is-decidable (Β¬ (π-to-Ξ© n holds))) β h (π-to-Ξ© n) d οΌ n
I β (inl Ο) = refl
I β (inl Ο) = π-elim (Ο β)
I β (inr Ο) = π-elim (Ο unique-from-π)
I β (inr Ο) = refl
d : (p : Ξ© π€) β is-decidable (Β¬ (p holds))
d p = wem (p holds)
II : retract π of (Ξ© π€)
II = (Ξ» p β h p (d p)) ,
π-to-Ξ© ,
(Ξ» n β I n (d (π-to-Ξ© n)))
WEM-gives-π-ainjective : typal-WEM π€ β ainjective-type π π€ π€
WEM-gives-π-ainjective {π€} wem =
retract-of-ainjective π (Ξ© π€) (Ξ©-ainjective pe') (WEM-gives-π-retract-of-Ξ© wem)
WEM-gives-π-aflabby : typal-WEM π€ β aflabby π π€
WEM-gives-π-aflabby wem = ainjective-types-are-aflabby π (WEM-gives-π-ainjective wem)
\end{code}
The simple types are not injective in general. These are the types
formed by starting with β and closing under function types. We can
also add the type π to the base case of the definition to get the same
conclusion.
\begin{code}
simple-typeβ-injective-gives-WEM : (X : π€β Μ )
β simple-typeβ X
β ainjective-type X π€ π€
β typal-WEM π€
simple-typeβ-injective-gives-WEM X s X-ainj =
π-ainjective-gives-WEM
(retract-of-ainjective π X X-ainj
(simple-typesβ-disconnected s))
simple-typeβ-injective-gives-WEM-examples
: (ainjective-type β π€ π€ β typal-WEM π€)
Γ (ainjective-type (β β β) π€ π€ β typal-WEM π€)
Γ (ainjective-type (β β π) π€ π€ β typal-WEM π€)
Γ (ainjective-type ((β β β) β β) π€ π€ β typal-WEM π€)
Γ (ainjective-type ((β β π) β β) π€ π€ β typal-WEM π€)
Γ (ainjective-type (((β β β) β β) β β) π€ π€ β typal-WEM π€)
simple-typeβ-injective-gives-WEM-examples =
simple-typeβ-injective-gives-WEM _ base ,
simple-typeβ-injective-gives-WEM _ (step base base) ,
simple-typeβ-injective-gives-WEM _ (step base baseβ) ,
simple-typeβ-injective-gives-WEM _ (step (step base base) base) ,
simple-typeβ-injective-gives-WEM _ (step (step base baseβ) base) ,
simple-typeβ-injective-gives-WEM _ (step (step (step base base) base) base)
\end{code}
TODO. We can also close under _Γ_ and _+_ to get the same result. We
can also close under Ξ , but maybe not under Ξ£.
If the type β of Dedekind reals is injective then there are
discontinuous functions β β β, e.g. the Heaviside function, which is
also a constructive taboo. Notice that the type β lives in the
universe π€β.
\begin{code}
open import DedekindReals.Type fe' pe' pt
open import DedekindReals.Order fe' pe' pt renaming (_β―_ to _β―β_)
open import Notation.Order
β-ainjective-gives-Heaviside-function : ainjective-type β π€β π€β
β Ξ£ H κ (β β β) ,
((x : β) β (x < 0β β H x οΌ 0β)
Γ (x β₯ 0β β H x οΌ 1β))
β-ainjective-gives-Heaviside-function β-ainj = H , Ξ³
where
j : (Ξ£ x κ β , x < 0β) + (Ξ£ x κ β , x β₯ 0β) β β
j = cases prβ prβ
j-is-embedding : is-embedding j
j-is-embedding = disjoint-cases-embedding prβ prβ
(prβ-is-embedding (Ξ» x β <-is-prop x 0β))
(prβ-is-embedding (Ξ» x β β€-is-prop 0β x))
d
where
d : disjoint-images prβ prβ
d (x , l) (x , b) refl = <β-irreflexive x (β<-β€-trans x 0β x l b)
h : (Ξ£ x κ β , x < 0β) + (Ξ£ x κ β , x β₯ 0β) β β
h = cases (Ξ» _ β 0β) (Ξ» _ β 1β)
H : β β β
H = prβ (β-ainj j j-is-embedding h)
H-extends-h-along-j : β u β H (j u) οΌ h u
H-extends-h-along-j = prβ (β-ainj j j-is-embedding h)
Ξ³ : (x : β) β (x < 0β β H x οΌ 0β)
Γ (x β₯ 0β β H x οΌ 1β)
Ξ³ x = I , II
where
I : x < 0β β H x οΌ 0β
I l = H-extends-h-along-j (inl (x , l))
II : x β₯ 0β β H x οΌ 1β
II b = H-extends-h-along-j (inr (x , b))
\end{code}
But we can do better than that and derive weak excluded middle from
the injectivity of β.
\begin{code}
open import Rationals.Type
open import Rationals.Order
β-ainjective-gives-WEM : ainjective-type β π€ π₯ β typal-WEM π€
β-ainjective-gives-WEM {π€} β-ainj = WEM-gives-typal-WEM fe' XI
where
module _ (P : π€ Μ ) (P-is-prop : is-prop P) where
q : Ξ© π€
q = (P + Β¬ P) , decidability-of-prop-is-prop fe' P-is-prop
β-aflabby : aflabby β π€
β-aflabby = ainjective-types-are-aflabby β β-ainj
f : P + Β¬ P β β
f = cases (Ξ» _ β 0β) (Ξ» _ β 1β)
r : β
r = aflabby-extension β-aflabby q f
I : P β r οΌ 0β
I p = aflabby-extension-property β-aflabby q f (inl p)
II : Β¬ P β r οΌ 1β
II Ξ½ = aflabby-extension-property β-aflabby q f (inr Ξ½)
I-II : r β 0β β r β 1β β π
I-II u v = contrapositive II v (contrapositive I u)
I-IIβ : r β 1β β r οΌ 0β
I-IIβ v = β-is-¬¬-separated r 0β (Ξ» u β I-II u v)
I-IIβ : r β 0β β r οΌ 1β
I-IIβ u = β-is-¬¬-separated r 1β (I-II u)
III : (1/4 < r) β¨ (r < 1/2)
III = β-locatedness r 1/4 1/2 1/4<1/2
IV : 1/4 < r β r οΌ 1β
IV l = I-IIβ IVβ
where
IVβ : r β 0β
IVβ e = β<-irrefl 1/4 IVβ
where
IVβ : 1/4 < 0β
IVβ = transport (1/4 <_) e l
IVβ : 1/4 < 1/4
IVβ = β<-trans 1/4 0β 1/4 IVβ 0<1/4
V : r < 1/2 β r οΌ 0β
V l = I-IIβ Vβ
where
Vβ : r β 1β
Vβ e = β<-irrefl 1/2 Vβ
where
Vβ : 1β < 1/2
Vβ = transport (_< 1/2) e l
Vβ : 1/2 < 1/2
Vβ = β<-trans 1/2 1β 1/2 1/2<1 Vβ
VI : r οΌ 0β β ¬¬ P
VI e Ξ½ = apartness-gives-inequality 0β 1β
β-zero-apart-from-one
(0β οΌβ¨ e β»ΒΉ β©
r οΌβ¨ II Ξ½ β©
1β β)
VII : r οΌ 1β β Β¬ P
VII e p = apartness-gives-inequality 0β 1β
β-zero-apart-from-one
(0β οΌβ¨ (I p)β»ΒΉ β©
r οΌβ¨ e β©
1β β)
VIII : r < 1/2 β ¬¬ P
VIII l = VI (V l)
IX : 1/4 β<β r β Β¬ P
IX l = VII (IV l)
X : ¬ P ⨠¬¬ P
X = β¨-functor IX VIII III
XI : ¬ P + ¬¬ P
XI = exit-β₯β₯ (decidability-of-prop-is-prop fe' (negations-are-props fe')) X
\end{code}
TODO. Probably the converse holds.
The injectivity of ββ, the conatural numbers, or generic convergent
sequence, gives WLPO. Like in the previous example, we first use
injectivity to define a non-continuous function.
\begin{code}
open import CoNaturals.Type
open import Taboos.BasicDiscontinuity (fe π€β π€β)
open import Taboos.WLPO
open import Notation.CanonicalMap
ββ-injective-gives-WLPO : ainjective-type ββ π€β π€β β WLPO
ββ-injective-gives-WLPO ββ-ainj = basic-discontinuity-taboo' f (fβ , fβ)
where
g : β + π β ββ
g (inl _) = ΞΉ 0
g (inr _) = ΞΉ 1
f : ββ β ββ
f = prβ (ββ-ainj ΞΉπ (ΞΉπ-is-embedding fe') g)
f-extends-g-along-ΞΉπ : β u β f (ΞΉπ u) οΌ g u
f-extends-g-along-ΞΉπ = prβ (ββ-ainj ΞΉπ (ΞΉπ-is-embedding fe') g)
fβ : (n : β) β f (ΞΉ n) οΌ ΞΉ 0
fβ n = f-extends-g-along-ΞΉπ (inl n)
fβ : f β οΌ ΞΉ 1
fβ = f-extends-g-along-ΞΉπ (inr β)
\end{code}
The above again illustrates that we can use injectivity to define
discontinuous functions. But we can actually get a stronger
conclusion with a weaker assumption and a simpler proof.
\begin{code}
ββ-injective-gives-WEM : ainjective-type ββ π€ π₯ β typal-WEM π€
ββ-injective-gives-WEM ββ-ainj =
π-ainjective-gives-WEM (retract-of-ainjective π ββ ββ-ainj π-retract-of-ββ)
\end{code}
Added 6 June 2024 by Tom de Jong during a meeting with MartΓn EscardΓ³.
A type with a nontrivial apartness relation cannot be injective unless weak
excluded middle holds.
TODO. We could derive β-ainjective-gives-WEM from the below. (Note the
similarities in the two proofs.)
\begin{code}
open import Apartness.Definition
open import Apartness.Properties pt
open Apartness pt
ainjective-type-with-non-trivial-apartness-gives-WEM
: {X : π€ Μ }
β ainjective-type X π£ π¦
β Nontrivial-Apartness X π₯
β typal-WEM π£
ainjective-type-with-non-trivial-apartness-gives-WEM
{π€} {π£} {π¦} {π₯} {X} ainj ((_β―_ , Ξ±) , ((xβ , xβ) , points-apart))
= WEM-gives-typal-WEM fe' VII
where
module _ (P : π£ Μ ) (P-is-prop : is-prop P) where
X-aflabby : aflabby X π£
X-aflabby = ainjective-types-are-aflabby _ ainj
f : (P + Β¬ P) β X
f = cases (Ξ» _ β xβ) (Ξ» _ β xβ)
q : Ξ© π£
q = (P + Β¬ P) , decidability-of-prop-is-prop fe' P-is-prop
x : X
x = aflabby-extension X-aflabby q f
I : P β x οΌ xβ
I p = aflabby-extension-property X-aflabby q f (inl p)
II : Β¬ P β x οΌ xβ
II Ξ½ = aflabby-extension-property X-aflabby q f (inr Ξ½)
III : x β xβ β Β¬ P
III = contrapositive I
IV : x β xβ β ¬¬ P
IV = contrapositive II
V : xβ β― x β¨ xβ β― x
V = apartness-is-cotransitive _β―_ Ξ± xβ xβ x points-apart
VI : (x β xβ) β¨ (x β xβ)
VI = β¨-functor Ξ½ Ξ½ V
where
Ξ½ : {x y : X} β x β― y β y β x
Ξ½ a refl = apartness-is-irreflexive _β―_ Ξ± _ a
VII : ¬ P + ¬¬ P
VII = β¨-elim (decidability-of-prop-is-prop fe' (negations-are-props fe'))
(inl β III) (inr β IV) VI
\end{code}
In particular, we have the following.
\begin{code}
non-trivial-apartness-on-universe-gives-WEM
: is-univalent π€
β Nontrivial-Apartness (π€ Μ ) π₯
β typal-WEM π€
non-trivial-apartness-on-universe-gives-WEM {π€} {π₯} ua =
ainjective-type-with-non-trivial-apartness-gives-WEM {π€ βΊ} {π€} {π€}
(universes-are-ainjective ua)
non-trivial-apartness-on-universe-iff-WEM
: is-univalent π€
β Nontrivial-Apartness (π€ Μ ) π€ β typal-WEM π€
non-trivial-apartness-on-universe-iff-WEM {π€} ua = f , g
where
f : Nontrivial-Apartness (π€ Μ ) π€ β typal-WEM π€
f = non-trivial-apartness-on-universe-gives-WEM ua
g : typal-WEM π€ β Nontrivial-Apartness (π€ Μ ) π€
g = WEM-gives-that-type-with-two-distinct-points-has-nontrivial-apartnessβΊ
fe'
(universes-are-locally-small ua)
universe-has-two-distinct-points
non-trivial-apartness-on-Ξ©-gives-WEM
: propext π€
β Nontrivial-Apartness (Ξ© π€) π₯
β typal-WEM π€
non-trivial-apartness-on-Ξ©-gives-WEM {π€} {π₯} pe =
ainjective-type-with-non-trivial-apartness-gives-WEM {π€ βΊ} {π€} {π€}
(Ξ©-ainjective pe)
non-trivial-apartness-on-Ξ©-iff-WEM
: propext π€
β funext π€ π€
β Nontrivial-Apartness (Ξ© π€) π€ β typal-WEM π€
non-trivial-apartness-on-Ξ©-iff-WEM {π€} pe fe = f , g
where
f : Nontrivial-Apartness (Ξ© π€) π€ β typal-WEM π€
f = non-trivial-apartness-on-Ξ©-gives-WEM pe
g : typal-WEM π€ β Nontrivial-Apartness (Ξ© π€) π€
g = WEM-gives-that-type-with-two-distinct-points-has-nontrivial-apartnessβΊ
fe'
(Ξ©-is-locally-small pe fe)
((β₯ , β€) , β₯-is-not-β€)
\end{code}
Notice that ainjective-type-with-non-trivial-apartness-gives-WEM
subsumes all the previous examples: the type π, which is a simple
type, the simple types (because they are totally separated and hence
they have a (tight) apartness), the Dedekind reals (with their
standard apartness), ββ (again because it is totally
separated).
TODO. Maybe we can list a few more interesting examples?
\end{code}
Added 27 January 2025 by Tom de Jong. Revised 6 February 2025.
We generalize non-trivial-totally-separated-ainjective-type-gives-¬¬-WEM from
Taboos.Decomposability, where the notion of total separatedness is exploited
directly, to derive ¬¬ WEM from the assumption of a non-trivial injective type
with a tight apartness.
\begin{code}
non-trivial-ainjective-type-with-tight-apartness-gives-¬¬-WEM
: (Ξ£ X κ π€ Μ , ((Β¬ is-prop X) Γ ainjective-type X π₯ π¦ Γ Tight-Apartness X π£))
β ¬¬ typal-WEM π₯
non-trivial-ainjective-type-with-tight-apartness-gives-¬¬-WEM
{π€} {π₯} {π¦} {π£}
(X , X-not-prop , X-inj , (_β―_ , β―-is-apartness , β―-is-tight)) = III
where
I : (x y : X) β x β― y β typal-WEM π₯
I x y a = ainjective-type-with-non-trivial-apartness-gives-WEM X-inj
((_β―_ , β―-is-apartness) , (((x , y) , a)))
II : Β¬ typal-WEM π₯ β is-prop X
II Ξ½ x y = β―-is-tight x y (Ξ» (a : x β― y) β π-elim (Ξ½ (I x y a)))
III : ¬¬ typal-WEM π₯
III Ξ½ = π-elim (X-not-prop (II Ξ½))
open import TypeTopology.TotallySeparated
non-trivial-totally-separated-ainjective-type-gives-¬¬-WEM'
: (Ξ£ X κ π€ Μ , ((Β¬ is-prop X) Γ is-totally-separated X Γ ainjective-type X π₯ π¦))
β ¬¬ typal-WEM π₯
non-trivial-totally-separated-ainjective-type-gives-¬¬-WEM'
(X , X-not-prop , X-tot-sep , X-inj) =
non-trivial-ainjective-type-with-tight-apartness-gives-¬¬-WEM
( X , X-not-prop , X-inj
, _β―β_ , β―β-is-apartness
, totally-separated-gives-totally-separatedβ X-tot-sep)
where
open total-separatedness-via-apartness pt
\end{code}
Added 16 October 2025 by Tom de Jong, formalizing a proof sketches
of 4 and 8 September 2025.
The following theorem is instantiated below to show that the injectivity of the
infinite dimensional real projective space RPβ implies a weak choice principle
known as the world's simplest axiom of choice.
\begin{code}
open import InjectiveTypes.Subtypes fe
open import UF.ExitPropTrunc
open split-support-and-collapsibility pt
family-has-unspecified-split-support-if-total-space-of-truncation-is-ainjective
: (D : π€ Μ )
β ainjective-type D π₯ π¦
β (T : D β π£ Μ )
β ainjective-type (Ξ£ d κ D , β₯ T d β₯) (π£ β π₯') π¦'
β (d : D) β β₯ has-split-support (T d) β₯
family-has-unspecified-split-support-if-total-space-of-truncation-is-ainjective
D D-inj T E-inj d = I
where
E = Ξ£ d κ D , β₯ T d β₯
lem : Ξ£ f κ (D β D) , ((x : D) β β₯ T (f x) β₯)
Γ ((x : D) β β₯ T x β₯ β f x οΌ x)
lem = necessary-condition-for-injectivity-of-subtype
D
(Ξ» x β β₯ T x β₯)
(Ξ» x β β₯β₯-is-prop)
E-inj
f : D β D
f = prβ lem
fβ : β₯ T (f d) β₯
fβ = prβ (prβ lem) d
fβ : β₯ T d β₯ β f d οΌ d
fβ = prβ (prβ lem) d
I : β₯ (β₯ T d β₯ β T d) β₯
I = β₯β₯-functor II fβ
where
II : T (f d) β β₯ T d β₯ β T d
II t Ο = transport T (fβ Ο) t
open import SyntheticHomotopyTheory.RP-infinity pt
open import UF.Choice
open world's-simplest-axiom-of-choice fe pt
RPβ-ainjective-implies-WSAC : ainjective-type RPβ π₯ π¦
β WSAC' π€β
RPβ-ainjective-implies-WSAC RPβ-inj =
family-has-unspecified-split-support-if-total-space-of-truncation-is-ainjective
(π€β Μ ) (universes-are-ainjective (ua π€β)) (Ξ» X β X β π) RPβ-inj
\end{code}