Encodable types #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file defines encodable (constructively countable) types as a typeclass.
This is used to provide explicit encode/decode functions from and to ℕ
, with the information that
those functions are inverses of each other.
The difference with denumerable
is that finite types are encodable. For infinite types,
encodable
and denumerable
agree.
Main declarations #
encodable α
: States that there exists an explicit encoding functionencode : α → ℕ
with a partial inversedecode : ℕ → option α
.decode₂
: Version ofdecode
that is equal tonone
outside of the range ofencode
. Useful as we do not require this in the definition ofdecode
.ulower α
: Any encodable type has an equivalent type living in the lowest universe, namely a subtype ofℕ
.ulower α
finds it.
Implementation notes #
The point of asking for an explicit partial inverse decode : ℕ → option α
to encode : α → ℕ
is
to make the range of encode
decidable even when the finiteness of α
is not.
- encode : α → ℕ
- decode : ℕ → option α
- encodek : ∀ (a : α), encodable.decode α (encodable.encode a) = option.some a
Constructively countable type. Made from an explicit injection encode : α → ℕ
and a partial
inverse decode : ℕ → option α
. Note that finite types are countable. See denumerable
if you
wish to enforce infiniteness.
Instances of this typeclass
- is_empty.to_encodable
- denumerable.to_encodable
- primcodable.to_encodable
- nat.encodable
- punit.encodable
- option.encodable
- sum.encodable
- bool.encodable
- Prop.encodable
- sigma.encodable
- prod.encodable
- subtype.encodable
- fin.encodable
- int.encodable
- pnat.encodable
- ulift.encodable
- plift.encodable
- ulower.encodable
- subgroup.opposite.encodable
- add_subgroup.opposite.encodable
- list.encodable
- multiset.encodable
- vector.encodable
- encodable.fin_arrow
- encodable.fin_pi
- finset.encodable
- encodable.fintype_arrow_of_encodable
- topological_space.encodable_countable_basis
- rat.encodable
- array.encodable
- W_type.encodable
- first_order.language.term.encodable
- prop_encodable.prop_form.encodable
Instances of other typeclasses for encodable
- encodable.has_sizeof_inst
An encodable type has decidable equality. Not set as an instance because this is usually not the best way to infer decidability.
Equations
If α
is encodable and there is an injection f : β → α
, then β
is encodable as well.
Equations
- encodable.of_left_injection f finv linv = {encode := λ (b : β), encodable.encode (f b), decode := λ (n : ℕ), (encodable.decode α n).bind finv, encodek := _}
If α
is encodable and f : β → α
is invertible, then β
is encodable as well.
Equations
- encodable.of_left_inverse f finv linv = encodable.of_left_injection f (option.some ∘ finv) _
Encodability is preserved by equivalence.
Equations
- encodable.of_equiv α e = encodable.of_left_inverse ⇑e ⇑(e.symm) _
Equations
- nat.encodable = {encode := id ℕ, decode := option.some ℕ, encodek := nat.encodable._proof_1}
Equations
- is_empty.to_encodable = {encode := is_empty_elim (λ (a : α), ℕ), decode := λ (n : ℕ), option.none, encodek := _}
Equations
- punit.encodable = {encode := λ (_x : punit), 0, decode := λ (n : ℕ), n.cases_on (option.some punit.star) (λ (_x : ℕ), option.none), encodek := punit.encodable._proof_1}
If α
is encodable, then so is option α
.
Equations
- option.encodable = {encode := λ (o : option α), o.cases_on 0 (λ (a : α), (encodable.encode a).succ), decode := λ (n : ℕ), n.cases_on (option.some option.none) (λ (m : ℕ), option.map option.some (encodable.decode α m)), encodek := _}
Failsafe variant of decode
. decode₂ α n
returns the preimage of n
under encode
if it
exists, and returns none
if it doesn't. This requirement could be imposed directly on decode
but
is not to help make the definition easier to use.
Equations
- encodable.decode₂ α n = (encodable.decode α n).bind (option.guard (λ (a : α), encodable.encode a = n))
The encoding function has decidable range.
Equations
- encodable.decidable_range_encode α = λ (x : ℕ), decidable_of_iff ↥((encodable.decode₂ α x).is_some) _
An encodable type is equivalent to the range of its encoding function.
Equations
- encodable.equiv_range_encode α = {to_fun := λ (a : α), ⟨encodable.encode a, _⟩, inv_fun := λ (n : ↥(set.range encodable.encode)), option.get _, left_inv := _, right_inv := _}
A type with unique element is encodable. This is not an instance to avoid diamonds.
Equations
- unique.encodable = {encode := λ (_x : α), 0, decode := λ (_x : ℕ), option.some inhabited.default, encodek := _}
Explicit encoding function for the sum of two encodable types.
Equations
- encodable.encode_sum (sum.inr b) = bit1 (encodable.encode b)
- encodable.encode_sum (sum.inl a) = bit0 (encodable.encode a)
Explicit decoding function for the sum of two encodable types.
Equations
- encodable.decode_sum n = encodable.decode_sum._match_1 n.bodd_div2
- encodable.decode_sum._match_1 (bool.tt, m) = option.map sum.inr (encodable.decode β m)
- encodable.decode_sum._match_1 (bool.ff, m) = option.map sum.inl (encodable.decode α m)
If α
and β
are encodable, then so is their sum.
Equations
- sum.encodable = {encode := encodable.encode_sum _inst_2, decode := encodable.decode_sum _inst_2, encodek := _}
Explicit encoding function for sigma γ
Equations
- encodable.encode_sigma ⟨a, b⟩ = nat.mkpair (encodable.encode a) (encodable.encode b)
Explicit decoding function for sigma γ
Equations
- encodable.decode_sigma n = encodable.decode_sigma._match_1 (nat.unpair n)
- encodable.decode_sigma._match_1 (n₁, n₂) = (encodable.decode α n₁).bind (λ (a : α), option.map (sigma.mk a) (encodable.decode (γ a) n₂))
Equations
- sigma.encodable = {encode := encodable.encode_sigma (λ (a : α), _inst_2 a), decode := encodable.decode_sigma (λ (a : α), _inst_2 a), encodek := _}
If α
and β
are encodable, then so is their product.
Equations
- prod.encodable = encodable.of_equiv (Σ (_x : α), β) (equiv.sigma_equiv_prod α β).symm
Explicit encoding function for a decidable subtype of an encodable type
Equations
Explicit decoding function for a decidable subtype of an encodable type
Equations
- encodable.decode_subtype v = (encodable.decode α v).bind (λ (a : α), dite (P a) (λ (h : P a), option.some ⟨a, h⟩) (λ (h : ¬P a), option.none))
A decidable subtype of an encodable type is encodable.
Equations
- subtype.encodable = {encode := encodable.encode_subtype encA, decode := encodable.decode_subtype (λ (a : α), decP a), encodek := _}
Equations
- fin.encodable n = encodable.of_equiv {i // i < n} fin.equiv_subtype
Equations
Equations
The lift of an encodable type is encodable.
Equations
The lift of an encodable type is encodable.
Equations
If β
is encodable and there is an injection f : α → β
, then α
is encodable as well.
Equations
- encodable.of_inj f hf = encodable.of_left_injection f (function.partial_inv f) _
If α
is countable, then it has a (non-canonical) encodable
structure.
Equations
See also nonempty_fintype
, nonempty_denumerable
.
The equivalence between the encodable type α
and ulower α : Type
.
Equations
Lowers an a : α
into ulower α
.
Equations
- ulower.down a = ⇑(ulower.equiv α) a
Equations
Constructive choice function for a decidable subtype of an encodable type.
Equations
- encodable.choose_x h = have this : ∃ (n : ℕ), good p (encodable.decode α n), from _, encodable.choose_x._match_4 (encodable.decode α (nat.find this)) _
- encodable.choose_x._match_4 (option.some a) h = ⟨a, h⟩
- _ = _
Constructive choice function for a decidable predicate over an encodable type.
Equations
The encode
function, viewed as an embedding.
Equations
- encodable.encode' α = {to_fun := encodable.encode _inst_1, inj' := _}
Given a directed r
function f : α → β
defined on an encodable inhabited type,
construct a noncomputable sequence such that r (f (x n)) (f (x (n + 1)))
and r (f a) (f (x (encode a + 1))
.
Equations
- directed.sequence f hf (n + 1) = let p : α := directed.sequence f hf n in directed.sequence._match_1 f hf p (encodable.decode α n)
- directed.sequence f hf 0 = inhabited.default
- directed.sequence._match_1 f hf p (option.some a) = classical.some _
- directed.sequence._match_1 f hf p option.none = classical.some _
Representative of an equivalence class. This is a computable version of quot.out
for a setoid
on an encodable type.
Equations
- q.rep = encodable.choose _
The quotient of an encodable space by a decidable equivalence relation is encodable.
Equations
- encodable_quotient = {encode := λ (q : quotient s), encodable.encode q.rep, decode := λ (n : ℕ), quotient.mk <$> encodable.decode α n, encodek := _}