Finite sets #
Terms of type Finset α
are one way of talking about finite subsets of α
in Mathlib.
Below, Finset α
is defined as a structure with 2 fields:
Finsets in Lean are constructive in that they have an underlying List
that enumerates their
elements. In particular, any function that uses the data of the underlying list cannot depend on its
ordering. This is handled on the Multiset
level by multiset API, so in most cases one needn't
worry about it explicitly.
Finsets give a basic foundation for defining finite sums and products over types:
Lean refers to these operations as big operators.
More information can be found in Mathlib/Algebra/BigOperators/Group/Finset.lean
.
Finsets are directly used to define fintypes in Lean.
A Fintype α
instance for a type α
consists of a universal Finset α
containing every term of
α
, called univ
. See Mathlib/Data/Fintype/Basic.lean
.
There is also univ'
, the noncomputable partner to univ
,
which is defined to be α
as a finset if α
is finite,
and the empty finset otherwise. See Mathlib/Data/Fintype/Basic.lean
.
Finset.card
, the size of a finset is defined in Mathlib/Data/Finset/Card.lean
.
This is then used to define Fintype.card
, the size of a type.
File structure #
This file defines the Finset
type and the membership and subset relations between finsets.
Most constructions involving Finset
s have been split off to their own files.
Main definitions #
Finset
: Defines a type for the finite subsets ofα
. Constructing aFinset
requires two pieces of data:val
, aMultiset α
of elements, andnodup
, a proof thatval
has no duplicates.Finset.instMembershipFinset
: Defines membershipa ∈ (s : Finset α)
.Finset.instCoeTCFinsetSet
: Provides a coercions : Finset α
tos : Set α
.Finset.instCoeSortFinsetType
: Coerces : Finset α
to the type of allx ∈ s
.
Tags #
finite sets, finset
val
contains no duplicates
Equations
- x✝.decidableEq x = decidable_of_iff (x✝.val = x.val) ⋯
membership #
Equations
- Finset.decidableMem a s = Multiset.decidableMem a s.val
set coercion #
Equations
- Finset.decidableMem' a s = Finset.decidableMem a s
extensionality #
type coercion #
Subset and strict subset relations #
Equations
- Finset.partialOrder = PartialOrder.mk ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Alias of the reverse direction of Finset.coe_subset
.
Equations
- ⋯ = ⋯
Coercion to Set α
as an OrderEmbedding
.
Equations
- Finset.coeEmb = { toFun := Finset.toSet, inj' := ⋯, map_rel_iff' := ⋯ }
Instances For
Assorted results #
These results can be defined using the current imports, but deserve to be given a nicer home.
Equations
- x✝.instDecidableRelSubset x = Finset.decidableDforallFinset
Equations
- x✝.instDecidableRelSSubset x = instDecidableAnd
Equations
- Finset.instDecidableLE = Finset.instDecidableRelSubset
Equations
- Finset.instDecidableLT = Finset.instDecidableRelSSubset
Equations
- Finset.decidableExistsAndFinset = decidable_of_iff (∃ (a : α) (_ : a ∈ s), p a) ⋯
Equations
- Finset.decidableExistsAndFinsetCoe = Finset.decidableExistsAndFinset
decidable equality for functions whose domain is bounded by finsets
Equations
- Finset.decidableEqPiFinset = Multiset.decidableEqPiMultiset
Equations
- List.instDecidableInjOnToSetOfDecidableEq = inferInstanceAs (Decidable (∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y))
Equations
- List.instDecidableMapsToToSetOfDecidablePredMemSet = inferInstanceAs (Decidable (∀ x ∈ s, f x ∈ t))
Equations
- List.instDecidableSurjOnToSetOfDecidableEq = inferInstanceAs (Decidable (∀ x ∈ t', ∃ y ∈ s, f y = x))
Equations
- List.instDecidableBijOnToSetOfDecidableEq = inferInstanceAs (Decidable (Set.MapsTo f ↑s ↑t' ∧ Set.InjOn f ↑s ∧ Set.SurjOn f ↑s ↑t'))