Documentation

Mathlib.Data.Fintype.Defs

Finite types #

This file defines a typeclass to state that a type is finite.

Main declarations #

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

These files also contain appropriate Infinite instances for these types.

Infinite instances for , , Multiset α, and List α are in Data.Fintype.Lattice.

class Fintype (α : Type u_4) :
Type u_4

Fintype α means that α is finite, i.e. there are only finitely many distinct elements of type α. The evidence of this is a finset elems (a list up to permutation without duplicates), together with a proof that everything of type α is in the list.

Instances
    def Finset.univ {α : Type u_1} [Fintype α] :

    univ is the universal finite set of type Finset α implied from the assumption Fintype α.

    Equations
    Instances For
      @[simp]
      theorem Finset.mem_univ {α : Type u_1} [Fintype α] (x : α) :
      theorem Finset.mem_univ_val {α : Type u_1} [Fintype α] (x : α) :
      theorem Finset.eq_univ_iff_forall {α : Type u_1} [Fintype α] {s : Finset α} :
      s = univ ∀ (x : α), x s
      theorem Finset.eq_univ_of_forall {α : Type u_1} [Fintype α] {s : Finset α} :
      (∀ (x : α), x s)s = univ
      @[simp]
      theorem Finset.coe_univ {α : Type u_1} [Fintype α] :
      @[simp]
      theorem Finset.coe_eq_univ {α : Type u_1} [Fintype α] {s : Finset α} :
      @[simp]
      theorem Finset.subset_univ {α : Type u_1} [Fintype α] (s : Finset α) :

      Elaborate set builder notation for Finset.

      See also

      • Data.Set.Defs for the Set builder notation elaborator that this elaborator partly overrides.
      • Data.Finset.Basic for the Finset builder notation elaborator partly overriding this one for syntax of the form {x ∈ s | p x}.
      • Data.Fintype.Basic for the Finset 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 the Finset 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
        instance Fintype.decidableForallFintype {α : Type u_1} {p : αProp} [DecidablePred p] [Fintype α] :
        Decidable (∀ (a : α), p a)
        Equations
        instance Fintype.decidableExistsFintype {α : Type u_1} {p : αProp} [DecidablePred p] [Fintype α] :
        Decidable (∃ (a : α), p a)
        Equations
        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
        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
        def Fintype.subtype {α : Type u_1} {p : αProp} (s : Finset α) (H : ∀ (x : α), x s p x) :
        Fintype { x : α // p x }

        Given a predicate that can be represented by a finset, the subtype associated to the predicate is a fintype.

        Equations
        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
          Instances For
            Equations
            instance OrderDual.fintype (α : Type u_4) [Fintype α] :
            Equations
            instance OrderDual.finite (α : Type u_4) [Finite α] :
            instance Lex.fintype (α : Type u_4) [Fintype α] :
            Equations