module Data.Fin.Properties where
Finite sets - properties🔗
Ordering🔗
As noted in Data.Fin.Base
, we’ve set up
the ordering on Fin
so that we can re-use all the proofs
about the ordering on Nat
.
However, there are still quite a few interesting things one can say
about skip
and squish
. In particular, we can prove the
simplicial identities, which characterize the interactions
between these two functions.
These lemmas might seem somewhat arbitrary and complicated, which is
true! However, they are enough to describe all the possible interactions
of skip
and squish
, which in turn are the building
blocks for every monotone function between Fin
, so it’s not that surprising that
they would be a bit of a mess!
: ∀ {n} (i j : Fin (suc n)) → i ≤ j
skip-comm → ∀ x → skip (weaken i) (skip j x) ≡ skip (fsuc j) (skip i x)
= refl
skip-comm fzero j le x (fsuc i) (fsuc j) le fzero = refl
skip-comm (fsuc i) (fsuc j) (Nat.s≤s le) (fsuc x) = ap fsuc (skip-comm i j le x)
skip-comm
: ∀ {n} (i j : Fin n) → i ≤ j
drop-comm → ∀ x → squish j (squish (weaken i) x) ≡ squish i (squish (fsuc j) x)
= refl
drop-comm fzero fzero le fzero (fsuc x) = refl
drop-comm fzero fzero le (fsuc j) le fzero = refl
drop-comm fzero (fsuc j) le (fsuc x) = refl
drop-comm fzero (fsuc i) (fsuc j) le fzero = refl
drop-comm (fsuc i) (fsuc j) (Nat.s≤s le) (fsuc x) = ap fsuc (drop-comm i j le x)
drop-comm
: ∀ {n} (i : Fin (suc n)) (j : Fin n) → i < fsuc j
squish-skip-comm → ∀ x → squish (fsuc j) (skip (weaken i) x) ≡ skip i (squish j x)
(Nat.s≤s p) x = refl
squish-skip-comm fzero j (fsuc i) (fsuc j) (Nat.s≤s p) fzero = refl
squish-skip-comm (fsuc i) (fsuc j) (Nat.s≤s p) (fsuc x) =
squish-skip-comm (squish-skip-comm i j p x)
ap fsuc
: ∀ {n} (i j : Fin n) → i ≡ j
squish-skip → ∀ x → squish j (skip (weaken j) x) ≡ x
= refl
squish-skip fzero fzero p x (fsuc j) p x = absurd (fzero≠fsuc p)
squish-skip fzero (fsuc i) fzero p x = refl
squish-skip (fsuc i) (fsuc j) p fzero = refl
squish-skip (fsuc i) (fsuc j) p (fsuc x) = ap fsuc (squish-skip i j (fsuc-inj p) x)
squish-skip
: ∀ {n} (i : Fin (suc n)) (j : Fin n) → i ≡ fsuc j
squish-skip-fsuc → ∀ x → squish j (skip i x) ≡ x
= refl
squish-skip-fsuc fzero fzero p x (fsuc j) p x = absurd (fzero≠fsuc p)
squish-skip-fsuc fzero (fsuc i) fzero p fzero = refl
squish-skip-fsuc (fsuc fzero) fzero p (fsuc x) = refl
squish-skip-fsuc (fsuc (fsuc i)) fzero p (fsuc x) =
squish-skip-fsuc (fzero≠fsuc (fsuc-inj (sym p)))
absurd (fsuc i) (fsuc j) p fzero = refl
squish-skip-fsuc (fsuc i) (fsuc j) p (fsuc x) =
squish-skip-fsuc (squish-skip-fsuc i j (fsuc-inj p) x)
ap fsuc
: ∀ {l k} → Fin (suc l) ≃ Fin (suc k) → Fin l ≃ Fin k
Fin-peel {l} {k} sl≃sk = (Iso→Equiv (l→k , (iso k→l b→a→b a→b→a))) where
Fin-peel : Fin (suc k) ≃ Fin (suc l)
sk≃sl = sl≃sk e⁻¹
sk≃sl module sl≃sk = Equiv sl≃sk
module sk≃sl = Equiv sk≃sl
: Fin l → Fin k
l→k with inspect (sl≃sk.to (fsuc x))
l→k x ... | fsuc y , _ = y
... | fzero , p with inspect (sl≃sk.to fzero)
... | fsuc y , _ = y
... | fzero , q = absurd (fzero≠fsuc (sl≃sk.injective₂ q p))
: Fin k → Fin l
k→l with inspect (sk≃sl.to (fsuc x))
k→l x ... | fsuc x , _ = x
... | fzero , p with inspect (sk≃sl.to fzero)
... | fsuc y , _ = y
... | fzero , q = absurd (fzero≠fsuc (sk≃sl.injective₂ q p))
: ∀ {ℓ} {A : Type ℓ} {y : A} .{x : ⊥} → absurd x ≡ y
absurd-path {x = ()}
absurd-path
: ∀ a → k→l (l→k a) ≡ a
a→b→a with inspect (sl≃sk.to (fsuc a))
a→b→a a | fsuc x , p' with inspect (sk≃sl.to (fsuc x))
a→b→a a | fsuc x , p' | fsuc y , q' = fsuc-inj (
a→b→a a (sk≃sl.to) (sym p') ∙ sl≃sk.η _)
sym q' ∙ ap | fsuc x , p' | fzero , q' = absurd contra where
a→b→a a = sl≃sk.injective₂ p' (sl≃sk.ε (fsuc x))
r = fzero≠fsuc (sym (r ∙ q'))
contra | fzero , p' with inspect (sl≃sk.to fzero)
a→b→a a | fzero , p' | fsuc x , q' with inspect (sk≃sl.to (fsuc x))
a→b→a a | fzero , p' | fsuc x , q' | fsuc y , r' = absurd do
a→b→a a (sym (sym r' ∙ ap sk≃sl.to (sym q') ∙ sl≃sk.η fzero))
fzero≠fsuc | fzero , p' | fsuc x , q' | fzero , r' with inspect (sk≃sl.to fzero)
a→b→a a | fzero , p' | fsuc x , q' | fzero , r' | fsuc z , s = fsuc-inj $
a→b→a a .to (sym p') ∙ sl≃sk.η (fsuc a)
sym s ∙ ap sk≃sl| fzero , p' | fsuc x , q' | fzero , r' | fzero , s = absurd-path
a→b→a a | fzero , p' | fzero , q' = absurd (fzero≠fsuc $
a→b→a a .injective₂ q' p')
sl≃sk
: ∀ b → l→k (k→l b) ≡ b
b→a→b with inspect (sk≃sl.to (fsuc b))
b→a→b b | fsuc x , p' with inspect (sl≃sk.to (fsuc x))
b→a→b b | fsuc x , p' | fsuc y , q' = fsuc-inj $
b→a→b b (sl≃sk.to) (sym p') ∙ sk≃sl.η _
sym q' ∙ ap | fsuc x , p' | fzero , q' = absurd contra where
b→a→b b = sk≃sl.injective₂ p' (sk≃sl.ε (fsuc x))
r = fzero≠fsuc (sym (r ∙ q'))
contra | fzero , p' with inspect (sk≃sl.to fzero)
b→a→b b | fzero , p' | fsuc x , q' with inspect (sl≃sk.to (fsuc x))
b→a→b b | fzero , p' | fsuc x , q' | fsuc y , r' = absurd (fzero≠fsuc $
b→a→b b (sym r' ∙ ap (sl≃sk.to) (sym q') ∙ sk≃sl.η _))
sym | fzero , p' | fsuc x , q' | fzero , r' with inspect (sl≃sk.to fzero)
b→a→b b | fzero , p' | fsuc x , q' | fzero , r' | fsuc z , s = fsuc-inj $
b→a→b a (sl≃sk.to) (sym p') ∙ sk≃sl.η (fsuc a)
sym s ∙ ap | fzero , p' | fsuc x , q' | fzero , r' | fzero , s = absurd-path
b→a→b a | fzero , p' | fzero , q' = absurd (fzero≠fsuc $
b→a→b b .injective₂ q' p')
sk≃sl
: ∀ {l k} → Fin l ≃ Fin k → l ≡ k
Fin-injective {zero} {zero} l≃k = refl
Fin-injective {zero} {suc k} l≃k with equiv→inverse (l≃k .snd) fzero
Fin-injective ... | ()
{suc l} {zero} l≃k with l≃k .fst fzero
Fin-injective ... | ()
{suc l} {suc k} sl≃sk = ap suc $ Fin-injective (Fin-peel sl≃sk)
Fin-injective
: ∀ {n} (x : ℕ< n) → to-ℕ< {n = n} (from-ℕ< x) ≡ x
to-from-ℕ< {n = suc n} x = Σ-prop-path! (to-from-ℕ {n = suc n} x) where
to-from-ℕ< : ∀ {n} x → to-nat {n = n} (from-ℕ< x) ≡ x .fst
to-from-ℕ {n = suc n} (zero , p) = refl
to-from-ℕ {n = suc n} (suc x , Nat.s≤s p) = ap suc (to-from-ℕ {n = n} (x , p))
to-from-ℕ
: ∀ {n} (x : Fin n) → from-ℕ< (to-ℕ< x) ≡ x
from-to-ℕ< = refl
from-to-ℕ< fzero (fsuc x) = ap fsuc (from-to-ℕ< x)
from-to-ℕ<
: ∀ {n} → Fin n ≃ ℕ< n
Fin≃ℕ< {n} = to-ℕ< , is-iso→is-equiv (iso from-ℕ< (to-from-ℕ< {n}) from-to-ℕ<)
Fin≃ℕ<
avoid-injective: ∀ {n} (i : Fin (suc n)) {j k : Fin (suc n)} {i≠j : i ≠ j} {i≠k : i ≠ k}
→ avoid i j i≠j ≡ avoid i k i≠k → j ≡ k
{fzero} {k} {i≠j} p = absurd (i≠j refl)
avoid-injective fzero {fsuc j} {fzero} {i≠k = i≠k} p = absurd (i≠k refl)
avoid-injective fzero {suc n} fzero {fsuc j} {fsuc k} p = ap fsuc p
avoid-injective {suc n} (fsuc i) {fzero} {fzero} p = refl
avoid-injective {suc n} (fsuc i) {fzero} {fsuc k} p = absurd (fzero≠fsuc p)
avoid-injective {suc n} (fsuc i) {fsuc j} {fzero} p = absurd (fzero≠fsuc (sym p))
avoid-injective {suc n} (fsuc i) {fsuc j} {fsuc k} p =
avoid-injective (avoid-injective {n} i {j} {k} (fsuc-inj p))
ap fsuc
skip-injective: ∀ {n} (i : Fin (suc n)) (j k : Fin n)
→ skip i j ≡ skip i k → j ≡ k
= fsuc-inj p
skip-injective fzero j k p (fsuc i) fzero fzero p = refl
skip-injective (fsuc i) fzero (fsuc k) p = absurd (fzero≠fsuc p)
skip-injective (fsuc i) (fsuc j) fzero p = absurd (fzero≠fsuc (sym p))
skip-injective (fsuc i) (fsuc j) (fsuc k) p = ap fsuc (skip-injective i j k (fsuc-inj p))
skip-injective
skip-skips: ∀ {n} (i : Fin (suc n)) (j : Fin n)
→ skip i j ≠ i
= fzero≠fsuc (sym p)
skip-skips fzero j p (fsuc i) fzero p = fzero≠fsuc p
skip-skips (fsuc i) (fsuc j) p = skip-skips i j (fsuc-inj p)
skip-skips
avoid-skip: ∀ {n} (i : Fin (suc n)) (j : Fin n) {neq : i ≠ skip i j}
→ avoid i (skip i j) neq ≡ j
= refl
avoid-skip fzero fzero (fsuc j) = refl
avoid-skip fzero (fsuc i) fzero = refl
avoid-skip (fsuc i) (fsuc j) = ap fsuc (avoid-skip i j)
avoid-skip
skip-avoid: ∀ {n} (i : Fin (suc n)) (j : Fin (suc n)) {i≠j : i ≠ j}
→ skip i (avoid i j i≠j) ≡ j
{i≠j} = absurd (i≠j refl)
skip-avoid fzero fzero {suc n} fzero (fsuc j) = refl
skip-avoid {suc n} (fsuc i) fzero = refl
skip-avoid {suc n} (fsuc i) (fsuc j) = ap fsuc (skip-avoid i j) skip-avoid
Iterated products and sums🔗
We can break down and over finite sets as iterated products and sums, respectively.
Fin-suc-Π: ∀ {ℓ} {n} {A : Fin (suc n) → Type ℓ}
→ (∀ x → A x) ≃ (A fzero × (∀ x → A (fsuc x)))
= Iso→Equiv λ where
Fin-suc-Π .fst f → f fzero , (λ x → f (fsuc x))
.snd .is-iso.inv (z , s) → fin-cons z s
.snd .is-iso.rinv x → refl
.snd .is-iso.linv k i fzero → k fzero
.snd .is-iso.linv k i (fsuc n) → k (fsuc n)
Fin-suc-Σ: ∀ {ℓ} {n} {A : Fin (suc n) → Type ℓ}
→ Σ (Fin (suc n)) A ≃ (A fzero ⊎ Σ (Fin n) (A ∘ fsuc))
= Iso→Equiv λ where
Fin-suc-Σ .fst (fzero , a) → inl a
.fst (fsuc x , a) → inr (x , a)
.snd .is-iso.inv (inl a) → fzero , a
.snd .is-iso.inv (inr (x , a)) → fsuc x , a
.snd .is-iso.rinv (inl _) → refl
.snd .is-iso.rinv (inr _) → refl
.snd .is-iso.linv (fzero , a) → refl
.snd .is-iso.linv (fsuc x , a) → refl
Finite choice🔗
An important fact about the (standard) finite sets in constructive
mathematics is that they always support choice, which we phrase
below as a “search” operator: if
is any Monoidal
functor on types,
then it commutes with products. Since
over
are
iterated products, we have
that
commutes with
Fin-Monoidal: ∀ {ℓ} n {A : Fin n → Type ℓ} {M}
(let module M = Effect M)
→ ⦃ Monoidal M ⦄
→ (∀ x → M.₀ (A x)) → M.₀ (∀ x → A x)
_ = invmap (λ _ ()) _ munit
Fin-Monoidal zero (suc n) k =
Fin-Monoidal (k 0 <,> Fin-Monoidal n (k ∘ fsuc)) Fin-suc-Π e⁻¹ <≃>
_ = Idiom
In particular, instantiating
with the propositional truncation
(which is an Idiom
and hence Monoidal
), we get a version of the axiom of choice for finite
sets.
finite-choice: ∀ {ℓ} n {A : Fin n → Type ℓ}
→ (∀ x → ∥ A x ∥) → ∥ (∀ x → A x) ∥
= Fin-Monoidal n finite-choice n
An immediate consequence is that surjections into a finite set (thus, between finite sets) merely split:
finite-surjection-split: ∀ {ℓ} {n} {B : Type ℓ}
→ (f : B → Fin n) → is-surjective f
→ ∥ (∀ x → fibre f x) ∥
= finite-choice _ finite-surjection-split f
Dually, we have that any Alternative
functor
commutes with
on finite sets, since those are iterated sums.
Fin-Alternative: ∀ {ℓ} n {A : Fin n → Type ℓ} {M}
(let module M = Effect M)
→ ⦃ Alternative M ⦄
→ (∀ x → M.₀ (A x)) → M.₀ (Σ (Fin n) A)
_ = invmap (λ ()) (λ ()) empty
Fin-Alternative zero (suc n) k =
Fin-Alternative (k 0 <+> Fin-Alternative n (k ∘ fsuc)) Fin-suc-Σ e⁻¹ <≃>
As a consequence, instantiating
with Dec
, we get that finite sets
are exhaustible and omniscient, which
means that any family of decidable types indexed by a finite sets yields
decidable
and
respectively.
instance
Dec-Fin-∀: ∀ {n ℓ} {A : Fin n → Type ℓ}
→ ⦃ ∀ {x} → Dec (A x) ⦄ → Dec (∀ x → A x)
{n} ⦃ d ⦄ = Fin-Monoidal n (λ _ → d)
Dec-Fin-∀
Dec-Fin-Σ: ∀ {n ℓ} {A : Fin n → Type ℓ}
→ ⦃ ∀ {x} → Dec (A x) ⦄ → Dec (Σ (Fin n) A)
{n} ⦃ d ⦄ = Fin-Alternative n λ _ → d Dec-Fin-Σ
Fin-omniscience: ∀ {n ℓ} (P : Fin n → Type ℓ) ⦃ _ : ∀ {x} → Dec (P x) ⦄
→ (Σ[ j ∈ Fin n ] P j × ∀ k → P k → j ≤ k) ⊎ (∀ x → ¬ P x)
{zero} P = inr λ ()
Fin-omniscience {suc n} P with holds? (P 0)
Fin-omniscience ... | yes here = inl (0 , here , λ _ _ → 0≤x)
... | no ¬here with Fin-omniscience (P ∘ fsuc)
... | inl (ix , pix , least) = inl (fsuc ix , pix , fin-cons (λ here → absurd (¬here here)) λ i pi → Nat.s≤s (least i pi))
... | inr nowhere = inr (fin-cons ¬here nowhere)
Fin-omniscience-neg: ∀ {n ℓ} (P : Fin n → Type ℓ) ⦃ _ : ∀ {x} → Dec (P x) ⦄
→ (∀ x → P x) ⊎ (Σ[ j ∈ Fin n ] ¬ P j × ∀ k → ¬ P k → j ≤ k)
with Fin-omniscience (¬_ ∘ P)
Fin-omniscience-neg P ... | inr p = inl λ i → dec→dne (p i)
... | inl (j , ¬pj , least) = inr (j , ¬pj , least)
Fin-find: ∀ {n ℓ} {P : Fin n → Type ℓ} ⦃ _ : ∀ {x} → Dec (P x) ⦄
→ ¬ (∀ x → P x)
→ Σ[ x ∈ Fin n ] ¬ P x × ∀ y → ¬ P y → x ≤ y
{P = P} ¬p with Fin-omniscience-neg P
Fin-find ... | inl p = absurd (¬p p)
... | inr p = p
Injections and surjections🔗
The standard finite sets are Dedekind-finite, which means that every injection is a bijection. We prove this by a straightforward but annoying induction on
Fin-injection→equiv: ∀ {n} (f : Fin n → Fin n)
→ injective f → is-equiv f
{zero} f inj .is-eqv ()
Fin-injection→equiv {suc n} f inj .is-eqv i with f 0 ≡? i
Fin-injection→equiv ... | yes p = contr (0 , p) λ (j , p') → Σ-prop-path! (inj (p ∙ sym p'))
... | no ¬p = contr
(fsuc (rec .centre .fst) , avoid-injective (f 0) (rec .centre .snd))
λ where
(fzero , p) → absurd (¬p p)
(fsuc j , p) → Σ-prop-path! (ap (fsuc ∘ fst)
(rec .paths (j , ap₂ (avoid (f 0)) p prop!)))
where
= Fin-injection→equiv {n}
rec (λ x → avoid (f 0) (f (fsuc x)) (fzero≠fsuc ∘ inj))
(λ p → fsuc-inj (inj (avoid-injective (f 0) p)))
.is-eqv (avoid (f 0) i ¬p)
Since every surjection between finite sets splits, any surjection has an injective right inverse, which is thus a bijection; by general properties of equivalences, this implies that is also a bijection.
Fin-surjection→equiv: ∀ {n} (f : Fin n → Fin n)
→ is-surjective f → is-equiv f
= case finite-surjection-split f surj of λ split →
Fin-surjection→equiv f surj (snd ∘ split)
left-inverse→equiv (Fin-injection→equiv (fst ∘ split)
(right-inverse→injective f (snd ∘ split)))
Vector operations🔗
avoid-insert: ∀ {n} {ℓ} {A : Type ℓ}
→ (ρ : Fin n → A)
→ (i : Fin (suc n)) (a : A)
→ (j : Fin (suc n))
→ (i≠j : i ≠ j)
→ (ρ [ i ≔ a ]) j ≡ ρ (avoid i j i≠j)
{n = n} ρ fzero a fzero i≠j = absurd (i≠j refl)
avoid-insert {n = suc n} ρ fzero a (fsuc j) i≠j = refl
avoid-insert {n = suc n} ρ (fsuc i) a fzero i≠j = refl
avoid-insert {n = suc n} ρ (fsuc i) a (fsuc j) i≠j =
avoid-insert (ρ ∘ fsuc) i a j (i≠j ∘ ap fsuc)
avoid-insert
insert-lookup: ∀ {n} {ℓ} {A : Type ℓ}
→ (ρ : Fin n → A)
→ (i : Fin (suc n)) (a : A)
→ (ρ [ i ≔ a ]) i ≡ a
{n = n} ρ fzero a = refl
insert-lookup {n = suc n} ρ (fsuc i) a = insert-lookup (ρ ∘ fsuc) i a
insert-lookup
delete-insert: ∀ {n} {ℓ} {A : Type ℓ}
→ (ρ : Fin n → A)
→ (i : Fin (suc n)) (a : A)
→ ∀ j → delete (ρ [ i ≔ a ]) i j ≡ ρ j
= refl
delete-insert ρ fzero a j (fsuc i) a fzero = refl
delete-insert ρ (fsuc i) a (fsuc j) = delete-insert (ρ ∘ fsuc) i a j
delete-insert ρ
insert-delete: ∀ {n} {ℓ} {A : Type ℓ}
→ (ρ : Fin (suc n) → A)
→ (i : Fin (suc n)) (a : A)
→ ρ i ≡ a
→ ∀ j → ((delete ρ i) [ i ≔ a ]) j ≡ ρ j
{n = n} ρ fzero a p fzero = sym p
insert-delete {n = n} ρ fzero a p (fsuc j) = refl
insert-delete {n = suc n} ρ (fsuc i) a p fzero = refl
insert-delete {n = suc n} ρ (fsuc i) a p (fsuc j) = insert-delete (ρ ∘ fsuc) i a p j insert-delete