module Data.Set.Pushout where
Pushouts of setsð
private variable â â' â'' : Level A B C : Type â f g : C â A pattern incl x = inc (inl x) pattern incr x = inc (inr x)
The pushout of two sets can be, as usual, constructed as a coequaliser of maps into a coproduct. However, pushouts in the category have a nice property: they preserve embeddings. This module is dedicated to proving this.
Pushout : â {â â' â''} {A : Type â} {B : Type â'} {C : Type â''} â (f : C â A) (g : C â B) â Type (â â â' â â'') Pushout {A = A} {B} {C} f g = Coeq {B = A â B} (λ z â inl (f z)) (λ z â inr (g z))
swap-pushout : Pushout f g â Pushout g f swap-pushout (incl x) = incr x swap-pushout (incr x) = incl x swap-pushout (glue x i) = glue x (~ i) swap-pushout (squash x y p q i j) = let f = swap-pushout in squash (f x) (f y) (λ i â f (p i)) (λ i â f (q i)) i j swap-swap : swap-pushout {f = f} {g = g} â swap-pushout ⡠λ x â x swap-swap = ext λ where (inl x) â refl (inr x) â refl swap-pushout-is-equiv : is-equiv (swap-pushout {f = f} {g = g}) swap-pushout-is-equiv = is-isoâis-equiv $ iso swap-pushout (happly swap-swap) (happly swap-swap)
Pushouts of injectionsð
module _ {â â' â''} {A : Type â} {B : Type â'} {C : Type â''} ⊠bset : H-Level B 2 ⊠(f : C â A) (g : C â B) (f-emb : is-embedding f) where
Let us now get to the proof of our key result. Suppose is an embedding and is some other map. Our objective is to prove that, in the square
the constructor is an embedding.1. As usual when working with higher inductive types, weâll employ encode-decode, characterising the path spaces based at some by means of a family over We know already we want to be since this is exactly injectivity of But what should the fibre over be?
A possibility comes from generalising the generating equality from identifying to identifying any provided we can cough up with and While this is a trivial rephrasing, it does inspire confidence: can we expect to be given by fibres of over
Well, if this were the case, these fibres would necessarily need to be propositions, so we can start by showing that. The proof turns out very simple: the type is equivalent to and since is an embedding, this is a subtype of a proposition.
remâ : â a â is-prop (Σ[ c â C ] (f c â¡ a) à (g c â¡ b)) remâ a (c , α , β) (c' , α' , β') = ap fst it ,â ap snd it ,â prop! where it = f-emb a (c , α) (c' , α')
instance remâ' : â {a n} â H-Level (Σ[ c â C ] (f c â¡ a) à (g c â¡ b)) (suc n) remâ' {a} = prop-instance (remâ a) {-# OVERLAPPING remâ' #-}
We can then turn back to defining our family Both of the point cases are handled, and since weâre mapping into that also handles the truncation.
code : Pushout f g â Prop (â â â' â â'') code (incr b') = el! (Lift (â â â'') (b â¡ b')) code (incl a) = el! (Σ[ c â C ] (f c â¡ a) à (g c â¡ b))
It remains to handle the paths. Since weâre mapping into a universe, it will suffice to show that the types and are equivalent, and, since weâre mapping into propositions, it will suffice to provide functions in either direction. Going backwards, this is trivial; In the other direction, we show since implies by injectivity of and we have by assumption.
code (glue x i) = let from : Lift _ (b â¡ g x) â Σ[ c â C ] (f c â¡ f x) à (g c â¡ b) from (lift p) = x , refl , sym p to : Σ[ c â C ] (f c â¡ f x) à (g c â¡ b) â Lift _ (b â¡ g x) to (x , α , β) = lift (sym β â ap g (ap fst (f-emb _ (_ , α) (_ , refl))))
in n-ua {X = el! (Σ[ c â C ] (f c â¡ f x) à (g c â¡ b))} {Y = el! (Lift (â â â'') (b â¡ g x))} (prop-ext! to from) i code (squash x y p q i j) = n-Type-is-hlevel 1 (code x) (code y) (λ i â code (p i)) (λ i â code (q i)) i j
We then have the two encoding functions. In the case with matched
endpoints, we have exactly our original goal: injectivity of incr
.
encodeáµ£ : injective incr encodeáµ£ {a} p = transport (λ i â â code a (p i) â) (lift refl) .lower
f-injâincr-inj : is-embedding {A = B} {B = Pushout f g} incr f-injâincr-inj = injectiveâis-embedding! encodeáµ£
In the mismatched case, we have a function turning paths into a fibre of over Itâs easy to show that projecting the point is the inverse to considered as a function
encodeâ : â {a b} (p : incr b â¡ incl a) â Σ[ c â C ] (f c â¡ a à g c â¡ b) encodeâ {b = b} p = transport (λ i â â code b (p i) â) (lift refl) pushout-is-pullback' : â x â fibre f x â fibre {B = Pushout f g} incr (incl x) pushout-is-pullback' x .fst (y , p) = g y , sym (glue y) â ap incl p pushout-is-pullback' x .snd = is-isoâis-equiv $ iso from ri λ f â f-emb x _ _ where from : fibre {B = Pushout f g} incr (incl x) â fibre f x from (y , p) with (c , α , β) â encodeâ p = c , α ri : is-right-inverse from (pushout-is-pullback' x .fst) ri (y , p) with (c , α , β) â encodeâ p = Σ-prop-path! β
Finally, we can extend this to an equivalence between and the pullback of along by general family-fibration considerations.
_ : C â (Σ[ a â A ] Σ[ b â B ] incl a â¡ incr b) _ = C ââš Total-equiv f â©â Σ[ a â A ] fibre f a â⚠Σ-ap-snd pushout-is-pullback' â©â Σ[ a â A ] fibre incr (incl a) ââšâ© Σ[ a â A ] Σ[ b â B ] (incr b â¡ incl a) â⚠Σ-ap-snd (λ a â Σ-ap-snd λ b â sym-equiv) â©â Σ[ a â A ] Σ[ b â B ] (incl a â¡ incr b) ââ
-- That equivalence computes like garbage on the third component, and -- we can do better: pushout-is-pullback : C â (Σ[ a â A ] Σ[ b â B ] incl a â¡ incr b) pushout-is-pullback .fst x = f x , g x , glue x pushout-is-pullback .snd = is-isoâis-equiv (iso from ri λ x â refl) where module _ (x : _) (let β = x .snd .snd) where from : C from = encodeâ (sym β) .fst ri : (f from , g from , glue from) â¡ (_ , _ , β) ri = encodeâ (sym β) .snd .fst ,â encodeâ (sym β) .snd .snd ,â prop! module _ ⊠aset : H-Level A 2 ⊠(f : C â A) (g : C â B) (g-emb : is-embedding g) where g-injâincl-inj : is-embedding {A = A} {B = Pushout f g} incl g-injâincl-inj = â-is-embedding (f-injâincr-inj g f g-emb) (is-equivâis-embedding swap-pushout-is-equiv)
In this situation, the map is often referred to as the cobase change of along â©ïž