Finite types #
This file defines a typeclass to state that a type is finite.
Main declarations #
Fintype α
: Typeclass saying that a type is finite. It takes as fields aFinset
and a proof that all terms of typeα
are in it.Finset.univ
: The finset of all elements of a fintype.
See Data.Fintype.Basic
for elementary results,
Data.Fintype.Card
for the cardinality of a fintype,
the equivalence with Fin (Fintype.card α)
, and pigeonhole principles.
Instances #
Instances for Fintype
for
{x // p x}
are in this file asFintype.subtype
Option α
are inData.Fintype.Option
α × β
are inData.Fintype.Prod
α ⊕ β
are inData.Fintype.Sum
Σ (a : α), β a
are inData.Fintype.Sigma
These files also contain appropriate Infinite
instances for these types.
Infinite
instances for ℕ
, ℤ
, Multiset α
, and List α
are in Data.Fintype.Lattice
.
Elaborate set builder notation for Finset
.
{x | p x}
is elaborated asFinset.filter (fun x ↦ p x) Finset.univ
if the expected type isFinset ?α
.{x : α | p x}
is elaborated asFinset.filter (fun x : α ↦ p x) Finset.univ
if the expected type isFinset ?α
.{x ∉ s | p x}
is elaborated asFinset.filter (fun x ↦ p x) sᶜ
if either the expected type isFinset ?α
or the expected type is notSet ?α
ands
has expected typeFinset ?α
.{x ≠ a | p x}
is elaborated asFinset.filter (fun x ↦ p x) {a}ᶜ
if the expected type isFinset ?α
.
See also
Data.Set.Defs
for theSet
builder notation elaborator that this elaborator partly overrides.Data.Finset.Basic
for theFinset
builder notation elaborator partly overriding this one for syntax of the form{x ∈ s | p x}
.Data.Fintype.Basic
for theFinset
builder notation elaborator handling syntax of the form{x | p x}
,{x : α | p x}
,{x ∉ s | p x}
,{x ≠ a | p x}
.Order.LocallyFinite.Basic
for theFinset
builder notation elaborator handling syntax of the form{x ≤ a | p x}
,{x ≥ a | p x}
,{x < a | p x}
,{x > a | p x}
.
TODO: Write a delaborator
Equations
- One or more equations did not get rendered due to their size.
Instances For
instance
Fintype.decidablePiFintype
{α : Type u_5}
{β : α → Type u_4}
[(a : α) → DecidableEq (β a)]
[Fintype α]
:
DecidableEq ((a : α) → β a)
Equations
- Fintype.decidablePiFintype f g = decidable_of_iff (∀ a ∈ Fintype.elems, f a = g a) ⋯
instance
Fintype.decidableForallFintype
{α : Type u_1}
{p : α → Prop}
[DecidablePred p]
[Fintype α]
:
Decidable (∀ (a : α), p a)
Equations
- Fintype.decidableForallFintype = decidable_of_iff (∀ a ∈ Finset.univ, p a) ⋯
instance
Fintype.decidableExistsFintype
{α : Type u_1}
{p : α → Prop}
[DecidablePred p]
[Fintype α]
:
Decidable (∃ (a : α), p a)
Equations
- Fintype.decidableExistsFintype = decidable_of_iff (∃ a ∈ Finset.univ, p a) ⋯
instance
Fintype.decidableMemRangeFintype
{α : Type u_1}
{β : Type u_2}
[Fintype α]
[DecidableEq β]
(f : α → β)
:
DecidablePred fun (x : β) => x ∈ Set.range f
Equations
instance
Fintype.decidableSubsingleton
{α : Type u_1}
[Fintype α]
[DecidableEq α]
{s : Set α}
[DecidablePred fun (x : α) => x ∈ s]
:
Equations
- Fintype.decidableSubsingleton = decidable_of_iff (∀ a ∈ s, ∀ b ∈ s, a = b) ⋯
instance
Fintype.decidableEqEquivFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype α]
:
DecidableEq (α ≃ β)
Equations
- Fintype.decidableEqEquivFintype a b = decidable_of_iff (a.toFun = b.toFun) ⋯
instance
Fintype.decidableEqEmbeddingFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype α]
:
DecidableEq (α ↪ β)
Equations
- Fintype.decidableEqEmbeddingFintype a b = decidable_of_iff (⇑a = ⇑b) ⋯
theorem
Fintype.nodup_map_univ_iff_injective
{α : Type u_1}
{β : Type u_2}
[Fintype α]
{f : α → β}
:
instance
Fintype.decidableInjectiveFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype α]
:
Equations
instance
Fintype.decidableSurjectiveFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype α]
[Fintype β]
:
Equations
instance
Fintype.decidableBijectiveFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype α]
[Fintype β]
:
Equations
instance
Fintype.decidableRightInverseFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq α]
[Fintype α]
(f : α → β)
(g : β → α)
:
Equations
instance
Fintype.decidableLeftInverseFintype
{α : Type u_1}
{β : Type u_2}
[DecidableEq β]
[Fintype β]
(f : α → β)
(g : β → α)
:
Equations
Given a predicate that can be represented by a finset, the subtype associated to the predicate is a fintype.
Equations
- Fintype.subtype s H = { elems := { val := Multiset.pmap Subtype.mk s.val ⋯, nodup := ⋯ }, complete := ⋯ }
Instances For
def
Fintype.ofFinset
{α : Type u_1}
{p : Set α}
(s : Finset α)
(H : ∀ (x : α), x ∈ s ↔ x ∈ p)
:
Fintype ↑p
Construct a fintype from a finset with the same elements.
Equations
- Fintype.ofFinset s H = Fintype.subtype s H
Instances For
Equations
- Bool.fintype = { elems := { val := {true, false}, nodup := Bool.fintype.proof_1 }, complete := Bool.fintype.proof_2 }
Equations
- OrderDual.fintype α = inst✝
Equations
- Lex.fintype α = inst✝