Subsets of finite types #
In a Fintype
, all Set
s are automatically Finset
s, and there are only finitely many of them.
Main results #
Set.toFinset
: convert a subset of a finite type to aFinset
Finset.fintypeCoeSort
:((s : Finset α) : Type*)
is a finite typeFintype.finsetEquivSet
:Finset α
andSet α
are equivalent ifα
is aFintype
Construct a finset enumerating a set s
, given a Fintype
instance.
Equations
- s.toFinset = Finset.map (Function.Embedding.subtype fun (x : α) => x ∈ s) Finset.univ
Instances For
Membership of a set with a Fintype
instance is decidable.
Using this as an instance leads to potential loops with Subtype.fintype
under certain decidability
assumptions, so it should only be declared a local instance.
Equations
- s.decidableMemOfFintype a = decidable_of_iff (a ∈ s.toFinset) ⋯
Instances For
Alias of the reverse direction of Set.toFinset_nonempty
.
Equations
Equations
Equations
Equations
- «Prop».fintype = { elems := { val := {True, False}, nodup := «Prop».fintype.proof_1 }, complete := «Prop».fintype.proof_2 }
Equations
A set on a fintype, when coerced to a type, is a fintype.
Equations
- setFintype s = Subtype.fintype fun (x : α) => x ∈ s
Instances For
Given Fintype α
, finsetEquivSet
is the equiv between Finset α
and Set α
. (All
sets on a finite type are finite.)
Equations
- Fintype.finsetEquivSet = { toFun := Finset.toSet, invFun := fun (s : Set α) => s.toFinset, left_inv := ⋯, right_inv := ⋯ }
Instances For
Given a fintype α
, finsetOrderIsoSet
is the order isomorphism between Finset α
and Set α
(all sets on a finite type are finite).
Equations
- Fintype.finsetOrderIsoSet = { toEquiv := Fintype.finsetEquivSet, map_rel_iff' := ⋯ }
Instances For
finset% t
elaborates t
as a Finset
.
If t
is a Set
, then inserts Set.toFinset
.
Does not make use of the expected type; useful for big operators over finsets.
#check finset% Finset.range 2 -- Finset Nat
#check finset% (Set.univ : Set Bool) -- Finset Bool
Equations
- finsetStx = Lean.ParserDescr.node `finsetStx 1022 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "finset% ") (Lean.ParserDescr.cat `term 0))
Instances For
quot_precheck
for the finset%
syntax.
Equations
- One or more equations did not get rendered due to their size.