Equivalences involving List
-like types #
This file defines some additional constructive equivalences using Encodable
and the pairing
function on ℕ
.
Explicit encoding function for List α
Equations
- Encodable.encodeList [] = 0
- Encodable.encodeList (a :: l) = (Nat.pair (Encodable.encode a) (Encodable.encodeList l)).succ
Instances For
Explicit decoding function for List α
Equations
- Encodable.decodeList 0 = some []
- Encodable.decodeList v.succ = match Nat.unpair v, ⋯ with | (v₁, v₂), h => let_fun this := ⋯; (fun (x1 : α) (x2 : List α) => x1 :: x2) <$> Encodable.decode v₁ <*> Encodable.decodeList v₂
Instances For
If α
is encodable, then so is List α
. This uses the pair
and unpair
functions from
Data.Nat.Pairing
.
Equations
- List.encodable = { encode := Encodable.encodeList, decode := Encodable.decodeList, encodek := ⋯ }
These two lemmas are not about lists, but are convenient to keep here and don't
require Finset.sort
.
A listable type with decidable equality is encodable.
Equations
- Encodable.encodableOfList l H = { encode := fun (a : α) => List.idxOf a l, decode := fun (x : ℕ) => l[x]?, encodek := ⋯ }
Instances For
A finite type is encodable. Because the encoding is not unique, we wrap it in Trunc
to
preserve computability.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A noncomputable way to arbitrarily choose an ordering on a finite type.
It is not made into a global instance, since it involves an arbitrary choice.
This can be locally made into an instance with attribute [local instance] Fintype.toEncodable
.
Equations
Instances For
If α
is denumerable, then so is List α
.
Equations
- Denumerable.denumerableList = { toEncodable := List.encodable, decode_inv := ⋯ }
A list on a unique type is equivalent to ℕ by sending each list to its length.
Equations
- Equiv.listUniqueEquiv α = { toFun := List.length, invFun := fun (n : ℕ) => List.replicate n default, left_inv := ⋯, right_inv := ⋯ }
Instances For
The type lists on unit is canonically equivalent to the natural numbers.