Martin Escardo \begin{code} {-# OPTIONS --without-K --exact-split --safe --auto-inline #-} module UF.Embeddings where open import MLTT.Spartan open import MLTT.Plus-Properties open import UF.Base open import UF.Subsingletons open import UF.Equiv open import UF.EquivalenceExamples open import UF.LeftCancellable open import UF.Yoneda open import UF.Retracts open import UF.FunExt open import UF.Lower-FunExt open import UF.Subsingletons-FunExt open import UF.Univalence open import UF.UA-FunExt is-embedding : {X : π€ Μ } {Y : π₯ Μ } β (X β Y) β π€ β π₯ Μ is-embedding f = β y β is-prop (fiber f y) being-embedding-is-prop : funext (π€ β π₯) (π€ β π₯) β {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-prop (is-embedding f) being-embedding-is-prop {π€} {π₯} fe f = Ξ -is-prop (lower-funext π€ π€ fe) (Ξ» x β being-prop-is-prop fe) embedding-criterion : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β ((x : X) β is-prop (fiber f (f x))) β is-embedding f embedding-criterion f Ο .(f x) (x , refl) = Ο x (x , refl) embedding-criterion' : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β ((x x' : X) β (f x οΌ f x') β (x οΌ x')) β is-embedding f embedding-criterion' {π€} {π₯} {X} {Y} f e = embedding-criterion f (Ξ» x' β equiv-to-prop (a x') (singleton-types'-are-props x')) where a : (x' : X) β fiber f (f x') β (Ξ£ x κ X , x οΌ x') a x' = Ξ£-cong (Ξ» x β e x x') vv-equivs-are-embeddings : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-vv-equiv f β is-embedding f vv-equivs-are-embeddings f e y = singletons-are-props (e y) equivs-are-embeddings : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-equiv f β is-embedding f equivs-are-embeddings f e = vv-equivs-are-embeddings f (equivs-are-vv-equivs f e) embeddings-with-sections-are-vv-equivs : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding f β has-section f β is-vv-equiv f embeddings-with-sections-are-vv-equivs f i (g , Ξ·) y = pointed-props-are-singletons (g y , Ξ· y) (i y) embeddings-with-sections-are-equivs : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding f β has-section f β is-equiv f embeddings-with-sections-are-equivs f i h = vv-equivs-are-equivs f (embeddings-with-sections-are-vv-equivs f i h) _βͺ_ : π€ Μ β π₯ Μ β π€ β π₯ Μ X βͺ Y = Ξ£ f κ (X β Y) , is-embedding f Subtypes' : (π€ {π₯} : Universe) β π₯ Μ β π€ βΊ β π₯ Μ Subtypes' π€ {π₯} Y = Ξ£ X κ π€ Μ , X βͺ Y Subtypes : π€ Μ β π€ βΊ Μ Subtypes {π€} Y = Subtypes' π€ Y etofun : {X : π€ Μ } {Y : π₯ Μ } β (X βͺ Y) β (X β Y) etofun = prβ is-embedding-etofun : {X : π€ Μ } {Y : π₯ Μ } β (e : X βͺ Y) β is-embedding (etofun e) is-embedding-etofun = prβ equivs-embedding : {X : π€ Μ } {Y : π₯ Μ } β X β Y β X βͺ Y equivs-embedding e = β e β , equivs-are-embeddings β e β (ββ-is-equiv e) embeddings-are-lc : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding f β left-cancellable f embeddings-are-lc f e {x} {x'} p = ap prβ (e (f x) (x , refl) (x' , (p β»ΒΉ))) is-embedding' : {X : π€ Μ } {Y : π₯ Μ } β (X β Y) β π€ β π₯ Μ is-embedding' f = β x x' β is-equiv (ap f {x} {x'}) embedding-embedding' : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding f β is-embedding' f embedding-embedding' {π€} {π₯} {X} {Y} f ise = g where b : (x : X) β is-singleton (fiber f (f x)) b x = (x , refl) , ise (f x) (x , refl) c : (x : X) β is-singleton (fiber' f (f x)) c x = retract-of-singleton (β-gives-β· (fiber-lemma f (f x))) (b x) g : (x x' : X) β is-equiv (ap f {x} {x'}) g x = universality-equiv x refl (central-point-is-universal (Ξ» x' β f x οΌ f x') (center (c x)) (centrality (c x))) embedding-criterion-converse : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding f β (x' x : X) β (f x' οΌ f x) β (x' οΌ x) embedding-criterion-converse f e x' x = β-sym (ap f {x'} {x} , embedding-embedding' f e x' x) embedding'-embedding : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-embedding' f β is-embedding f embedding'-embedding {π€} {π₯} {X} {Y} f ise = g where e : (x : X) β is-central (Ξ£ x' κ X , f x οΌ f x') (x , refl) e x = universal-element-is-central (x , refl) (equiv-universality x refl (ise x)) h : (x : X) β is-prop (fiber' f (f x)) h x Ο Ο = Ο οΌβ¨ (e x Ο)β»ΒΉ β© (x , refl) οΌβ¨ e x Ο β© Ο β g' : (y : Y) β is-prop (fiber' f y) g' y (x , p) = transport (Ξ» - β is-prop (Ξ£ x' κ X , - οΌ f x')) (p β»ΒΉ) (h x) (x , p) g : (y : Y) β is-prop (fiber f y) g y = left-cancellable-reflects-is-prop β fiber-lemma f y β (section-lc _ (equivs-are-sections _ (ββ-is-equiv (fiber-lemma f y )))) (g' y) prβ-is-embedding : {X : π€ Μ } {Y : X β π₯ Μ } β ((x : X) β is-prop (Y x)) β is-embedding (prβ {π€} {π₯} {X} {Y}) prβ-is-embedding f x ((.x , y') , refl) ((.x , y'') , refl) = g where g : (x , y') , refl οΌ (x , y'') , refl g = ap (Ξ» - β (x , -) , refl) (f x y' y'') prβ-lc-bis : {X : π€ Μ } {Y : X β π₯ Μ } β ({x : X} β is-prop (Y x)) β left-cancellable prβ prβ-lc-bis f {u} {v} r = embeddings-are-lc prβ (prβ-is-embedding (Ξ» x β f {x})) r prβ-is-embedding-converse : {X : π€ Μ } {Y : X β π₯ Μ } β is-embedding (prβ {π€} {π₯} {X} {Y}) β (x : X) β is-prop (Y x) prβ-is-embedding-converse {π€} {π₯} {X} {Y} ie x = h where e : Ξ£ Y β X e = prβ {π€} {π₯} {X} {Y} isp : is-prop (fiber e x) isp = ie x s : Y x β fiber e x s y = (x , y) , refl r : fiber e x β Y x r ((x , y) , refl) = y rs : (y : Y x) β r (s y) οΌ y rs y = refl h : is-prop (Y x) h = left-cancellable-reflects-is-prop s (section-lc s (r , rs)) isp embedding-closed-under-βΌ : {X : π€ Μ } {Y : π₯ Μ } (f g : X β Y) β is-embedding f β g βΌ f β is-embedding g embedding-closed-under-βΌ f g e H y = equiv-to-prop (βΌ-fiber-β H y) (e y) K-idtofun-lc : K-axiom (π€ βΊ) β {X : π€ Μ } (x y : X) (A : X β π€ Μ ) β left-cancellable (idtofun (Id x y) (A y)) K-idtofun-lc {π€} k {X} x y A {p} {q} r = k (π€ Μ ) p q lc-maps-into-sets-are-embeddings : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β left-cancellable f β is-set Y β is-embedding f lc-maps-into-sets-are-embeddings {π€} {π₯} {X} {Y} f f-lc iss y (x , p) (x' , p') = to-Ξ£-Id (r , q) where r : x οΌ x' r = f-lc (p β (p' β»ΒΉ)) q : yoneda-nat x (Ξ» x β f x οΌ y) p x' r οΌ p' q = iss (yoneda-nat x (Ξ» x β f x οΌ y) p x' r) p' lc-maps-are-embeddings-with-K : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β left-cancellable f β K-axiom π₯ β is-embedding f lc-maps-are-embeddings-with-K {π€} {π₯} {X} {Y} f f-lc k = lc-maps-into-sets-are-embeddings f f-lc (k Y) id-is-embedding : {X : π€ Μ } β is-embedding (id {π€} {X}) id-is-embedding = singleton-types'-are-props β-is-embedding : {X : π€ Μ } {Y : π₯ Μ } {Z : π¦ Μ } {f : X β Y} {g : Y β Z} β is-embedding f β is-embedding g β is-embedding (g β f) β-is-embedding {π€} {π₯} {π¦} {X} {Y} {Z} {f} {g} e d = h where T : (z : Z) β π€ β π₯ β π¦ Μ T z = Ξ£ (y , _) κ fiber g z , fiber f y T-is-prop : (z : Z) β is-prop (T z) T-is-prop z = subtype-of-prop-is-prop prβ (prβ-lc (Ξ» {t} β e (prβ t))) (d z) Ο : (z : Z) β fiber (g β f) z β T z Ο z (x , p) = (f x , p) , x , refl Ξ³ : (z : Z) β T z β fiber (g β f) z Ξ³ z ((.(f x) , p) , x , refl) = x , p Ξ³Ο : (z : Z) (t : fiber (g β f) z) β Ξ³ z (Ο z t) οΌ t Ξ³Ο .(g (f x)) (x , refl) = refl h : (z : Z) β is-prop (fiber (g β f) z) h z = subtype-of-prop-is-prop (Ο z) (sections-are-lc (Ο z) (Ξ³ z , (Ξ³Ο z))) (T-is-prop z) \end{code} TODO. Redo the above proof using the technique of the following proof. \begin{code} factor-is-embedding : {X : π€ Μ } {Y : π₯ Μ } {Z : π¦ Μ } (f : X β Y) (g : Y β Z) β is-embedding (g β f) β is-embedding g β is-embedding f factor-is-embedding {π€} {π₯} {π¦} {X} {Y} {Z} f g i j = Ξ³ where a : (x x' : X) β (x οΌ x') β (g (f x) οΌ g (f x')) a x x' = ap (g β f) {x} {x'} , embedding-embedding' (g β f) i x x' b : (y y' : Y) β (y οΌ y') β (g y οΌ g y') b y y' = ap g {y} {y'} , embedding-embedding' g j y y' c : (x x' : X) β (f x οΌ f x') β (x οΌ x') c x x' = (f x οΌ f x') ββ¨ b (f x) (f x') β© (g (f x) οΌ g (f x')) ββ¨ β-sym (a x x') β© (x οΌ x') β Ξ³ : is-embedding f Ξ³ = embedding-criterion' f c embedding-exponential : FunExt β {X : π€ Μ } {Y : π₯ Μ } {A : π¦ Μ } (f : X β Y) β is-embedding f β is-embedding (Ξ» (Ο : A β X) β f β Ο) embedding-exponential {π€} {π₯} {π¦} fe {X} {Y} {A} f i = Ξ³ where g : (Ο Ο' : A β X) (a : A) β (Ο a οΌ Ο' a) β (f (Ο a) οΌ f (Ο' a)) g Ο Ο' a = ap f {Ο a} {Ο' a} , embedding-embedding' f i (Ο a) (Ο' a) h : (Ο Ο' : A β X) β Ο βΌ Ο' β f β Ο βΌ f β Ο' h Ο Ο' = Ξ -cong (fe π¦ π€) (fe π¦ π₯) A (Ξ» a β Ο a οΌ Ο' a) (Ξ» a β f (Ο a) οΌ f (Ο' a)) (g Ο Ο') k : (Ο Ο' : A β X) β (f β Ο οΌ f β Ο') β (Ο οΌ Ο') k Ο Ο' = (f β Ο οΌ f β Ο') ββ¨ β-funext (fe π¦ π₯) (f β Ο) (f β Ο') β© (f β Ο βΌ f β Ο') ββ¨ β-sym (h Ο Ο') β© (Ο βΌ Ο') ββ¨ β-sym (β-funext (fe π¦ π€) Ο Ο') β© (Ο οΌ Ο') β Ξ³ : is-embedding (f β_) Ξ³ = embedding-criterion' (f β_) k disjoint-images : {X : π€ Μ } {Y : π₯ Μ } {A : π¦ Μ } β (X β A) β (Y β A) β π€ β π₯ β π¦ Μ disjoint-images f g = β x y β f x β g y disjoint-cases-embedding : {X : π€ Μ } {Y : π₯ Μ } {A : π¦ Μ } (f : X β A) (g : Y β A) β is-embedding f β is-embedding g β disjoint-images f g β is-embedding (cases f g) disjoint-cases-embedding {π€} {π₯} {π¦} {X} {Y} {A} f g ef eg d = Ξ³ where Ξ³ : (a : A) (Ο Ο : Ξ£ z κ X + Y , cases f g z οΌ a) β Ο οΌ Ο Ξ³ a (inl x , p) (inl x' , p') = r where q : x , p οΌ x' , p' q = ef a (x , p) (x' , p') h : fiber f a β fiber (cases f g) a h (x , p) = inl x , p r : inl x , p οΌ inl x' , p' r = ap h q Ξ³ a (inl x , p) (inr y , q) = π-elim (d x y (p β q β»ΒΉ)) Ξ³ a (inr y , q) (inl x , p) = π-elim (d x y (p β q β»ΒΉ)) Ξ³ a (inr y , q) (inr y' , q') = r where p : y , q οΌ y' , q' p = eg a (y , q) (y' , q') h : fiber g a β fiber (cases f g) a h (y , q) = inr y , q r : inr y , q οΌ inr y' , q' r = ap h p \end{code} TODO. (1) f : X β Y is an embedding iff fiber f (f x) is a singleton for every x : X. (2) f : X β Y is an embedding iff its corestriction to its image is an equivalence. This can be deduced directly from Yoneda. \begin{code} module _ {π€ π₯ π¦ π£} {X : π€ Μ } {A : X β π₯ Μ } {Y : π¦ Μ } {B : Y β π£ Μ } (f : X β Y) (g : (x : X) β A x β B (f x)) where inl-is-embedding : (X : π€ Μ ) (Y : π₯ Μ ) β is-embedding (inl {π€} {π₯} {X} {Y}) inl-is-embedding {π€} {π₯} X Y (inl a) (.a , refl) (.a , refl) = refl inl-is-embedding {π€} {π₯} X Y (inr b) (x , p) (x' , p') = π-elim (+disjoint p) inr-is-embedding : (X : π€ Μ ) (Y : π₯ Μ ) β is-embedding (inr {π€} {π₯} {X} {Y}) inr-is-embedding {π€} {π₯} X Y (inl b) (x , p) (x' , p') = π-elim (+disjoint' p) inr-is-embedding {π€} {π₯} X Y (inr a) (.a , refl) (.a , refl) = refl maps-of-props-into-sets-are-embeddings : {P : π€ Μ } {X : π₯ Μ } (f : P β X) β is-prop P β is-set X β is-embedding f maps-of-props-into-sets-are-embeddings f i j q (p , s) (p' , s') = to-Ξ£-οΌ (i p p' , j _ s') maps-of-props-are-embeddings : {P : π€ Μ } {Q : π₯ Μ } (f : P β Q) β is-prop P β is-prop Q β is-embedding f maps-of-props-are-embeddings f i j = maps-of-props-into-sets-are-embeddings f i (props-are-sets j) Γ-embedding : {X : π€ Μ } {Y : π₯Β Μ } {A : π¦ Μ } {B : π£ Μ } (f : X β A ) (g : Y β B) β is-embedding f β is-embedding g β is-embedding (Ξ» ((x , y) : X Γ Y) β (f x , g y)) Γ-embedding f g i j (a , b) = retract-of-prop (r , (s , rs)) (Γ-is-prop (i a) (j b)) where r : fiber f a Γ fiber g b β fiber (Ξ» (x , y) β f x , g y) (a , b) r ((x , s) , (y , t)) = (x , y) , to-Γ-οΌ s t s : fiber (Ξ» (x , y) β f x , g y) (a , b) β fiber f a Γ fiber g b s ((x , y) , p) = (x , ap prβ p) , (y , ap prβ p) rs : (Ο : fiber (Ξ» (x , y) β f x , g y) (a , b)) β r (s Ο) οΌ Ο rs ((x , y) , refl) = refl NatΞ£-is-embedding : {X : π€ Μ } (A : X β π₯ Μ ) (B : X β π¦ Μ ) (ΞΆ : Nat A B) β ((x : X) β is-embedding (ΞΆ x)) β is-embedding (NatΞ£ ΞΆ) NatΞ£-is-embedding A B ΞΆ i (x , b) = equiv-to-prop (β-sym (NatΞ£-fiber-equiv A B ΞΆ x b)) (i x b) NatΞ£-is-embedding-converse : {X : π€ Μ } (A : X β π₯ Μ ) (B : X β π¦ Μ ) (ΞΆ : Nat A B) β is-embedding (NatΞ£ ΞΆ) β ((x : X) β is-embedding (ΞΆ x)) NatΞ£-is-embedding-converse A B ΞΆ e x b = equiv-to-prop (NatΞ£-fiber-equiv A B ΞΆ x b) (e (x , b)) NatΞ -is-embedding : {X : π€ Μ } (A : X β π₯ Μ ) (B : X β π¦ Μ ) (ΞΆ : Nat A B) β funext π€ (π₯ β π¦) β ((x : X) β is-embedding (ΞΆ x)) β is-embedding (NatΞ ΞΆ) NatΞ -is-embedding {π€} {π₯} {π¦} A B ΞΆ fe i g = equiv-to-prop (β-sym (NatΞ -fiber-equiv A B ΞΆ (lower-funext π€ π₯ fe) g)) (Ξ -is-prop fe (Ξ» x β i x (g x))) \end{code} For any proposition P, the unique map P β π is an embedding: \begin{code} prop-embedding : (P : π€ Μ ) β is-prop P β β π₯ β is-embedding (unique-to-π {π€} {π₯}) prop-embedding P i π₯ * (p , r) (p' , r') = to-Γ-οΌ (i p p') (props-are-sets π-is-prop r r') \end{code} Added by Tom de Jong. If a type X embeds into a proposition, then X is itself a proposition. \begin{code} embedding-into-prop : {X : π€ Μ } {P : π₯ Μ } β is-prop P β X βͺ P β is-prop X embedding-into-prop i (f , e) x y = d where a : x οΌ y β f x οΌ f y a = ap f {x} {y} b : is-equiv a b = embedding-embedding' f e x y c : f x οΌ f y c = i (f x) (f y) d : x οΌ y d = inverse a b c \end{code} \begin{code} infix 0 _βͺ_ \end{code}