Cardinalities of finite types #
Main declarations #
- Fintype.card α: Cardinality of a fintype. Equal to- Finset.univ.card.
- Fintype.truncEquivFin: A fintype- αis computably equivalent to- Fin (card α). The- Trunc-free, noncomputable version is- Fintype.equivFin.
- Fintype.truncEquivOfCardEq- Fintype.equivOfCardEq: Two fintypes of same cardinality are equivalent. See above.
- Fin.equiv_iff_eq:- Fin m ≃ Fin niff- m = n.
- Infinite.natEmbedding: An embedding of- ℕinto an infinite type.
We also provide the following versions of the pigeonholes principle.
- Fintype.exists_ne_map_eq_of_card_ltand- isEmpty_of_card_lt: Finitely many pigeons and pigeonholes. Weak formulation.
- Finite.exists_ne_map_eq_of_infinite: Infinitely many pigeons in finitely many pigeonholes. Weak formulation.
- Finite.exists_infinite_fiber: Infinitely many pigeons in finitely many pigeonholes. Strong formulation.
Some more pigeonhole-like statements can be found in Data.Fintype.CardEmbedding.
Types which have an injection from/a surjection to an Infinite type are themselves Infinite.
See Infinite.of_injective and Infinite.of_surjective.
Instances #
We provide Infinite instances for
card α is the number of elements in α, defined when α is a fintype.
Equations
- Fintype.card α = Finset.univ.card
Instances For
There is (computably) an equivalence between α and Fin (card α).
Since it is not unique and depends on which permutation
of the universe list is used, the equivalence is wrapped in Trunc to
preserve computability.
See Fintype.equivFin for the noncomputable version,
and Fintype.truncEquivFinOfCardEq and Fintype.equivFinOfCardEq
for an equiv α ≃ Fin n given Fintype.card α = n.
See Fintype.truncFinBijection for a version without [DecidableEq α].
Equations
- One or more equations did not get rendered due to their size.
Instances For
There is (noncomputably) an equivalence between α and Fin (card α).
See Fintype.truncEquivFin for the computable version,
and Fintype.truncEquivFinOfCardEq and Fintype.equivFinOfCardEq
for an equiv α ≃ Fin n given Fintype.card α = n.
Equations
- Fintype.equivFin α = (Fintype.truncEquivFin α).out
Instances For
There is (computably) a bijection between Fin (card α) and α.
Since it is not unique and depends on which permutation
of the universe list is used, the bijection is wrapped in Trunc to
preserve computability.
See Fintype.truncEquivFin for a version that gives an equivalence
given [DecidableEq α].
Equations
- One or more equations did not get rendered due to their size.
Instances For
If the cardinality of α is n, there is computably a bijection between α and Fin n.
See Fintype.equivFinOfCardEq for the noncomputable definition,
and Fintype.truncEquivFin and Fintype.equivFin for the bijection α ≃ Fin (card α).
Equations
- Fintype.truncEquivFinOfCardEq h = Trunc.map (fun (e : α ≃ Fin (Fintype.card α)) => e.trans (finCongr h)) (Fintype.truncEquivFin α)
Instances For
If the cardinality of α is n, there is noncomputably a bijection between α and Fin n.
See Fintype.truncEquivFinOfCardEq for the computable definition,
and Fintype.truncEquivFin and Fintype.equivFin for the bijection α ≃ Fin (card α).
Equations
Instances For
Two Fintypes with the same cardinality are (computably) in bijection.
See Fintype.equivOfCardEq for the noncomputable version,
and Fintype.truncEquivFinOfCardEq and Fintype.equivFinOfCardEq for
the specialization to Fin.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Two Fintypes with the same cardinality are (noncomputably) in bijection.
See Fintype.truncEquivOfCardEq for the computable version,
and Fintype.truncEquivFinOfCardEq and Fintype.equivFinOfCardEq for
the specialization to Fin.
Equations
- Fintype.equivOfCardEq h = (Fintype.truncEquivOfCardEq h).out
Instances For
Note: this lemma is specifically about Fintype.ofSubsingleton. For a statement about
arbitrary Fintype instances, use either Fintype.card_le_one_iff_subsingleton or
Fintype.card_unique.
Note: this lemma is specifically about Fintype.ofIsEmpty. For a statement about
arbitrary Fintype instances, use Fintype.card_eq_zero.
Fin as a map from ℕ to Type is injective. Note that since this is a statement about
equality of types, using it should be avoided if possible.
Given that α ⊕ β is a fintype, α is also a fintype. This is non-computable as it uses
that Sum.inl is an injection, but there's no clear inverse if α is empty.
Equations
- Fintype.sumLeft = Fintype.ofInjective Sum.inl ⋯
Instances For
Given that α ⊕ β is a fintype, β is also a fintype. This is non-computable as it uses
that Sum.inr is an injection, but there's no clear inverse if β is empty.
Equations
- Fintype.sumRight = Fintype.ofInjective Sum.inr ⋯
Instances For
Relation to Finite #
In this section we prove that α : Type* is Finite if and only if Fintype α is nonempty.
Noncomputably get a Fintype instance from a Finite instance. This is not an
instance because we want Fintype instances to be useful for computations.
Equations
- Fintype.ofFinite α = ⋯.some
Instances For
The pigeonhole principle for finitely many pigeons and pigeonholes.
This is the Fintype version of Finset.exists_ne_map_eq_of_card_lt_of_maps_to.
Alias of Fintype.card_eq_zero.
A Fintype with cardinality zero is equivalent to Empty.
Equations
- Fintype.cardEqZeroEquivEquivEmpty = (Equiv.ofIff ⋯).trans (Equiv.equivEmptyEquiv α).symm
Instances For
Equations
- ⋯ = ⋯
Alias of the forward direction of Finite.injective_iff_bijective.
Alias of the forward direction of Finite.surjective_iff_bijective.
Alias of the forward direction of Finite.injective_iff_surjective_of_equiv.
Alias of the reverse direction of Finite.injective_iff_surjective_of_equiv.
Construct an equivalence from functions that are inverse to each other.
Equations
- Equiv.ofLeftInverseOfCardLE hβα f g h = { toFun := f, invFun := g, left_inv := h, right_inv := ⋯ }
Instances For
Construct an equivalence from functions that are inverse to each other.
Equations
- Equiv.ofRightInverseOfCardLE hαβ f g h = { toFun := f, invFun := g, left_inv := ⋯, right_inv := h }
Instances For
An embedding from a Fintype to itself can be promoted to an equivalence.
Equations
- e.equivOfFintypeSelfEmbedding = Equiv.ofBijective ⇑e ⋯
Instances For
If ‖β‖ < ‖α‖ there are no embeddings α ↪ β.
This is a formulation of the pigeonhole principle.
Note this cannot be an instance as it needs h.
A constructive embedding of a fintype α in another fintype β when card α ≤ card β.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If two subtypes of a fintype have equal cardinality, so do their complements.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A non-infinite type is a fintype.
Equations
Instances For
Any type is (classically) either a Fintype, or Infinite.
One can obtain the relevant typeclasses via cases fintypeOrInfinite α.
Equations
- fintypeOrInfinite α = if h : Infinite α then PSum.inr h else PSum.inl (fintypeOfNotInfinite h)
Instances For
If s : Set α is a proper subset of α and f : α → s is injective, then α is infinite.
If s : Set α is a proper subset of α and f : s → α is surjective, then α is infinite.
Equations
- ⋯ = ⋯
Embedding of ℕ into an infinite type.
Equations
- Infinite.natEmbedding α = { toFun := Infinite.natEmbeddingAux α, inj' := ⋯ }
Instances For
The pigeonhole principle for infinitely many pigeons in finitely many pigeonholes. If there are infinitely many pigeons in finitely many pigeonholes, then there are at least two pigeons in the same pigeonhole.
See also: Fintype.exists_ne_map_eq_of_card_lt, Finite.exists_infinite_fiber.
The strong pigeonhole principle for infinitely many pigeons in finitely many pigeonholes. If there are infinitely many pigeons in finitely many pigeonholes, then there is a pigeonhole with infinitely many pigeons.
See also: Finite.exists_ne_map_eq_of_infinite
For any c : List ℕ whose sum is at most Fintype.card α,
we can find o : List (List α) whose members have no duplicate,
whose lengths given by c, and which are pairwise disjoint
A custom induction principle for fintypes. The base case is a subsingleton type,
and the induction step is for non-trivial types, and one can assume the hypothesis for
smaller types (via Fintype.card).
The major premise is Fintype α, so to use this with the induction tactic you have to give a name
to that instance and use that name.