Skip to content

Commit 5643247

Browse files
committed
defn: embeddings of sets are stable under pushout
1 parent c695af7 commit 5643247

File tree

4 files changed

+311
-1
lines changed

4 files changed

+311
-1
lines changed

src/1Lab/Extensionality.agda

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -363,7 +363,12 @@ instance
363363
Extensional-equiv
364364
: {ℓ ℓ' ℓr} {A : Type ℓ} {B : Type ℓ'}
365365
→ ⦃ ea : Extensional (A B) ℓr ⦄ Extensional (A ≃ B) ℓr
366-
Extensional-equiv ⦃ ea ⦄ = Σ-prop-extensional (λ x is-equiv-is-prop _) ea
366+
Extensional-equiv ⦃ ea ⦄ = Σ-prop-extensional (λ x hlevel 1) ea
367+
368+
Extensional-emb
369+
: {ℓ ℓ' ℓr} {A : Type ℓ} {B : Type ℓ'}
370+
→ ⦃ ea : Extensional (A B) ℓr ⦄ Extensional (A ↪ B) ℓr
371+
Extensional-emb ⦃ ea ⦄ = Σ-prop-extensional (λ x hlevel 1) ea
367372

368373
Extensional-tr-map
369374
: {ℓ ℓ' ℓr} {A : Type ℓ} {B : Type ℓ'}

src/Data/Fin/Finite.lagda.md

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ open import 1Lab.Prelude
44
55
open import Algebra.Group.Homotopy.BAut
66
7+
open import Data.Set.Coequaliser
78
open import Data.Fin.Properties
89
open import Data.Fin.Closure
910
open import Data.Fin.Base
@@ -263,6 +264,27 @@ Finite-Lift = Finite-≃ (Lift-≃ e⁻¹)
263264
```
264265
-->
265266
267+
```agda
268+
abstract instance
269+
Finite-Coeq : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f g : A → B} ⦃ _ : Finite A ⦄ ⦃ _ : Finite B ⦄ → Finite (Coeq f g)
270+
```
271+
272+
<!--
273+
```agda
274+
Finite-Coeq {A = A} {B} {f} {g} ⦃ fin ae ⦄ ⦃ fin be ⦄ = ∥-∥-out! do
275+
ae ← ae
276+
be ← be
277+
let
278+
f' = Equiv.to be ∘ f ∘ Equiv.from ae
279+
g' = Equiv.to be ∘ g ∘ Equiv.from ae
280+
281+
fn : Σ[ n ∈ Nat ] Fin n ≃ Coeq f' g'
282+
fn = Finite-coequaliser f' g'
283+
284+
pure (fin {cardinality = fn .fst} (inc (Coeq-ap ae be refl refl ∙e Equiv.inverse (fn .snd))))
285+
```
286+
-->
287+
266288
<!--
267289
```agda
268290
card-zero→empty : ∥ A ≃ Fin 0 ∥ → ¬ A

src/Data/Set/Coequaliser.lagda.md

Lines changed: 33 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -690,5 +690,38 @@ instance
690690
: ∀ {ℓ ℓ'} {A : Type ℓ} {R : A → A → Type ℓ'} {x y} {n}
691691
→ H-Level (Closure R x y) (suc n)
692692
Closure-H-Level = prop-instance squash
693+
694+
Coeq-ap
695+
: ∀ {ℓa ℓa' ℓb ℓb'} {A : Type ℓa} {A' : Type ℓa'} {B : Type ℓb} {B' : Type ℓb'}
696+
{f g : A → B} {f' g' : A' → B'} (ea : A ≃ A') (eb : B ≃ B')
697+
→ (p : f' ≡ Equiv.to eb ∘ f ∘ Equiv.from ea) (q : g' ≡ Equiv.to eb ∘ g ∘ Equiv.from ea)
698+
→ Coeq f g ≃ Coeq f' g'
699+
Coeq-ap {f = f} {g} {f'} {g'} ea eb p q = Iso→Equiv (to , iso from (happly ri) (happly li)) where
700+
module ea = Equiv ea
701+
module eb = Equiv eb
702+
703+
to : Coeq f g → Coeq f' g'
704+
to (inc x) = inc (eb.to x)
705+
to (glue x i) = along i $
706+
inc (eb.to (f x)) ≡˘⟨ ap Coeq.inc (happly p (ea.to x) ∙ ap eb.to (ap f (ea.η x))) ⟩
707+
inc (f' (ea.to x)) ≡⟨ Coeq.glue {f = f'} {g'} (ea.to x) ⟩
708+
inc (g' (ea.to x)) ≡⟨ ap Coeq.inc (happly q (ea.to x) ∙ ap eb.to (ap g (ea.η x))) ⟩
709+
inc (eb.to (g x)) ∎
710+
to (squash x y p q i j) = squash (to x) (to y) (λ i → to (p i)) (λ i → to (q i)) i j
711+
712+
from : Coeq f' g' → Coeq f g
713+
from (inc x) = inc (eb.from x)
714+
from (glue x i) = along i $
715+
inc (eb.from (f' x)) ≡⟨ ap Coeq.inc (eb.injective (eb.ε _ ∙ happly p x)) ⟩
716+
inc (f (ea.from x)) ≡⟨ Coeq.glue (ea.from x) ⟩
717+
inc (g (ea.from x)) ≡⟨ ap Coeq.inc (eb.injective (sym (happly q x) ∙ sym (eb.ε _))) ⟩
718+
inc (eb.from (g' x)) ∎
719+
from (squash x y p q i j) = squash (from x) (from y) (λ i → from (p i)) (λ i → from (q i)) i j
720+
721+
li : from ∘ to ≡ λ x → x
722+
li = ext λ x → ap inc (eb.η _)
723+
724+
ri : to ∘ from ≡ λ x → x
725+
ri = ext λ x → ap inc (eb.ε _)
693726
```
694727
-->

src/Data/Set/Pushout.lagda.md

Lines changed: 250 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,250 @@
1+
<!--
2+
```agda
3+
open import 1Lab.Prelude
4+
5+
open import Data.Set.Coequaliser
6+
open import Data.Sum
7+
```
8+
-->
9+
10+
```agda
11+
module Data.Set.Pushout where
12+
```
13+
14+
# Pushouts of sets
15+
16+
<!--
17+
```agda
18+
private variable
19+
ℓ ℓ' ℓ'' : Level
20+
A B C : Type ℓ
21+
f g : C → A
22+
23+
pattern incl x = inc (inl x)
24+
pattern incr x = inc (inr x)
25+
```
26+
-->
27+
28+
The [[pushout]] of two [[sets]] can be, as usual, constructed as a
29+
[[coequaliser]] of maps into a [[coproduct]]. However, pushouts in the
30+
category $\Sets$ have a nice property: they preserve [[embeddings]].
31+
This module is dedicated to proving this.
32+
33+
```agda
34+
Pushout
35+
: ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''}
36+
→ (f : C → A) (g : C → B)
37+
→ Type (ℓ ⊔ ℓ' ⊔ ℓ'')
38+
Pushout {A = A} {B} {C} f g = Coeq {B = A ⊎ B} (λ z → inl (f z)) (λ z → inr (g z))
39+
```
40+
41+
We'll also need the isomorphism between $A \sqcup_C B$ and $B \sqcup_C
42+
A$, constructed by swapping the constructors.
43+
44+
```agda
45+
swap-pushout : Pushout f g → Pushout g f
46+
swap-pushout (incl x) = incr x
47+
swap-pushout (incr x) = incl x
48+
swap-pushout (glue x i) = glue x (~ i)
49+
swap-pushout (squash x y p q i j) =
50+
let f = swap-pushout in squash (f x) (f y) (λ i → f (p i)) (λ i → f (q i)) i j
51+
```
52+
53+
<!--
54+
```agda
55+
swap-swap : swap-pushout {f = f} {g = g} ∘ swap-pushout ≡ λ x → x
56+
swap-swap = ext λ where
57+
(inl x) → refl
58+
(inr x) → refl
59+
60+
swap-pushout-is-equiv : is-equiv (swap-pushout {f = f} {g = g})
61+
swap-pushout-is-equiv = is-iso→is-equiv $
62+
iso swap-pushout (happly swap-swap) (happly swap-swap)
63+
```
64+
-->
65+
66+
## Pushouts of injections
67+
68+
<!--
69+
```agda
70+
module
71+
_ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''}
72+
⦃ bset : H-Level B 2 ⦄
73+
(f : C → A) (g : C → B) (f-emb : is-embedding f)
74+
where
75+
```
76+
-->
77+
78+
Let us now get to the proof of our key result. Suppose $f : C \mono A$
79+
is an embedding and $g : C \to B$ is some other map. Our objective is to
80+
prove that, in the square
81+
82+
~~~{.quiver}
83+
\[\begin{tikzcd}[ampersand replacement=\&]
84+
C \&\& B \\
85+
\\
86+
A \&\& {A \sqcup_C B}
87+
\arrow["g", from=1-1, to=1-3]
88+
\arrow["f"', hook, from=1-1, to=3-1]
89+
\arrow["{\rm{incr}}", from=1-3, to=3-3]
90+
\arrow["{\rm{incl}}"', from=3-1, to=3-3]
91+
\end{tikzcd}\]
92+
~~~
93+
94+
the constructor $\rm{incl} : B \to A \sqcup_C B$ is an embedding.^[In
95+
this situation, the map $\rm{incr}$ is often referred to as the **cobase
96+
change** of $f$ along $g$.]. As usual when working with higher inductive
97+
types, we'll employ encode-decode, characterising the path spaces based
98+
at some $b : B$ by means of a family $F$ over $A \sqcup_C B$. We know
99+
already we want $F(\rm{incr}\, b')$ to be $b = b'$, since this is
100+
exactly injectivity of $\rm{incr}$. But what should the fibre over
101+
$\rm{incl}\, a$ be?
102+
103+
A possibility comes from generalising the generating equality
104+
$\rm{glue}$ from identifying $\rm{incl} (f x) = \rm{incr} (g x)$ to
105+
identifying any $a$, $b$ provided we can cough up $c : C$ with $f c = a$
106+
and $g c = b$. While this is a trivial rephrasing, it does inspire
107+
confidence: can we expect $\rm{incr}\, b = \rm{incl}\, a$ to be given by
108+
fibres of $\langle f, g \rangle$ over $(a, b)$?
109+
110+
Well, if this were the case, this these fibres would necessarily need to
111+
be propositions, so we can start by showing that. The proof turns out
112+
very simple: the type $$\sum_{c\, :\, C} (f c = a) \times (g c = b)$$ is
113+
equivalent to $$\sum_{(c,-)\, :\, f^*(a)} g^*(c)$$, and since $f$ is an
114+
embedding, this is a subtype of a proposition.
115+
116+
<!--
117+
```agda
118+
private module _ (b : B) where
119+
```
120+
-->
121+
122+
```agda
123+
rem₁ : ∀ a → is-prop (Σ[ c ∈ C ] (f c ≡ a) × (g c ≡ b))
124+
rem₁ a (c , α , β) (c' , α' , β') = ap fst it ,ₚ ap snd it ,ₚ prop! where
125+
it = f-emb a (c , α) (c' , α')
126+
```
127+
128+
<!--
129+
```agda
130+
instance
131+
rem₁' : ∀ {a n} → H-Level (Σ[ c ∈ C ] (f c ≡ a) × (g c ≡ b)) (suc n)
132+
rem₁' {a} = prop-instance (rem₁ a)
133+
{-# OVERLAPPING rem₁' #-}
134+
```
135+
-->
136+
137+
We can then turn back to defining our family $F$. Both of the point
138+
cases are handled, and since we're mapping into $\Omega$, that also
139+
handles the truncation.
140+
141+
```agda
142+
code : Pushout f g → Prop (ℓ ⊔ ℓ' ⊔ ℓ'')
143+
code (incr b') = el! (Lift (ℓ ⊔ ℓ'') (b ≡ b'))
144+
code (incl a) = el! (Σ[ c ∈ C ] (f c ≡ a) × (g c ≡ b))
145+
```
146+
147+
It remains to handle the paths. Since we're mapping into a universe, it
148+
will suffice to show that the types
149+
$$\sum_{c : C} (f c = f x) \times (g c = b)$$
150+
and $b = g x$ are equivalent, and, since we're mapping into
151+
propositions, it will suffice to provide functions in either direction.
152+
Going backwards, this is trivial; In the other direction, we show $b =
153+
g x$ since $f c = f x$ implies $c = x$ by injectivity of $f$, and
154+
we have $g c = b$ by assumption.
155+
156+
```agda
157+
code (glue x i) =
158+
let
159+
from : Lift _ (b ≡ g x) → Σ[ c ∈ C ] (f c ≡ f x) × (g c ≡ b)
160+
from (lift p) = x , refl , sym p
161+
162+
to : Σ[ c ∈ C ] (f c ≡ f x) × (g c ≡ b) → Lift _ (b ≡ g x)
163+
to (x , α , β) = lift (sym β ∙ ap g (ap fst (f-emb _ (_ , α) (_ , refl))))
164+
```
165+
166+
<!--
167+
```agda
168+
in n-ua
169+
{X = el! (Σ[ c ∈ C ] (f c ≡ f x) × (g c ≡ b))}
170+
{Y = el! (Lift (ℓ ⊔ ℓ'') (b ≡ g x))}
171+
(prop-ext! to from) i
172+
173+
code (squash x y p q i j) = n-Type-is-hlevel 1
174+
(code x) (code y) (λ i → code (p i)) (λ i → code (q i)) i j
175+
```
176+
-->
177+
178+
We then have the two decoding functions. In the case with matched
179+
endpoints, we have exactly our original goal: injectivity of
180+
`incr`{.Agda}.
181+
182+
```agda
183+
decodeᵣ : injective incr
184+
decodeᵣ {a} p = transport (λ i → ⌞ code a (p i) ⌟) (lift refl) .lower
185+
```
186+
187+
<!--
188+
```agda
189+
f-inj→incr-inj : is-embedding {A = B} {B = Pushout f g} incr
190+
f-inj→incr-inj = injective→is-embedding! λ {x} r →
191+
transport (λ i → ⌞ code x (r i) ⌟) (lift refl) .lower
192+
```
193+
-->
194+
195+
In the mismatched case, we have a function turning paths $\rm{incr}\, b
196+
= \rm{incl}\, a$ into a fibre of $\langle f, g \rangle$ over $(a, b)$.
197+
It's easy to show that projecting the point is the inverse to $g$,
198+
considered as a function $f^*(x) \to \rm{incr}^*(\rm{incl}\, x)$.
199+
200+
```agda
201+
decodeₗ
202+
: ∀ {a b} (p : incr b ≡ incl a)
203+
→ Σ[ c ∈ C ] (f c ≡ a × g c ≡ b)
204+
decodeₗ {b = b} p = transport (λ i → ⌞ code b (p i) ⌟) (lift refl)
205+
206+
pushout-is-pullback' : ∀ x → fibre f x ≃ fibre {B = Pushout f g} incr (incl x)
207+
pushout-is-pullback' x .fst (y , p) = g y , sym (glue y) ∙ ap incl p
208+
pushout-is-pullback' x .snd = is-iso→is-equiv $ iso from ri λ f → f-emb x _ _ where
209+
from : fibre {B = Pushout f g} incr (incl x) → fibre f x
210+
from (y , p) with (c , α , β) ← decodeₗ p = c , α
211+
212+
ri : is-right-inverse from (pushout-is-pullback' x .fst)
213+
ri (y , p) with (c , α , β) ← decodeₗ p = Σ-prop-path! β
214+
```
215+
216+
Finally, we can extend this to an equivalence between $C$ and the
217+
pullback of $\rm{incl}$ along $\rm{incr}$ by general
218+
[[family-fibration|object classifiers]] considerations.
219+
220+
```agda
221+
_ : C ≃ (Σ[ a ∈ A ] Σ[ b ∈ B ] incl a ≡ incr b)
222+
_ = C ≃⟨ Total-equiv f ⟩
223+
Σ[ a ∈ A ] fibre f a ≃⟨ Σ-ap-snd pushout-is-pullback' ⟩
224+
Σ[ a ∈ A ] fibre incr (incl a) ≃⟨⟩
225+
Σ[ a ∈ A ] Σ[ b ∈ B ] (incr b ≡ incl a) ≃⟨ Σ-ap-snd (λ a → Σ-ap-snd λ b → sym-equiv) ⟩
226+
Σ[ a ∈ A ] Σ[ b ∈ B ] (incl a ≡ incr b) ≃∎
227+
```
228+
229+
<!--
230+
```agda
231+
-- That equivalence computes like garbage on the third component, and
232+
-- we can do better:
233+
234+
pushout-is-pullback : C ≃ (Σ[ a ∈ A ] Σ[ b ∈ B ] incl a ≡ incr b)
235+
pushout-is-pullback .fst x = f x , g x , glue x
236+
pushout-is-pullback .snd = is-iso→is-equiv (iso from ri λ x → refl) where
237+
module _ (x : _) (let β = x .snd .snd) where
238+
from : C
239+
from = decodeₗ (sym β) .fst
240+
241+
ri : (f from , g from , glue from) ≡ (_ , _ , β)
242+
ri = decodeₗ (sym β) .snd .fst ,ₚ decodeₗ (sym β) .snd .snd ,ₚ prop!
243+
244+
module _ ⦃ aset : H-Level A 2 ⦄ (f : C → A) (g : C → B) (g-emb : is-embedding g) where
245+
g-inj→incl-inj : is-embedding {A = A} {B = Pushout f g} incl
246+
g-inj→incl-inj = ∙-is-embedding
247+
(f-inj→incr-inj g f g-emb)
248+
(is-equiv→is-embedding swap-pushout-is-equiv)
249+
```
250+
-->

0 commit comments

Comments
 (0)