Constructing finite sets by adding one element #
This file contains the definitions of {a} : Finset α
, insert a s : Finset α
and Finset.cons
,
all ways to construct a Finset
by adding one element.
Main declarations #
Finset.induction_on
: Induction on finsets. To prove a proposition about an arbitraryFinset α
, it suffices to prove it for the empty finset, and to show that if it holds for someFinset α
, then it holds for the finset obtained by inserting a new element.Finset.instSingletonFinset
: Denoted by{a}
; the finset consisting of one element.insert
andFinset.cons
: For anya : α
,insert s a
returnss ∪ {a}
.cons s a h
returns the same except that it requires a hypothesis stating thata
is not already ins
. This does not require decidable equality on the typeα
.
Tags #
finite sets, finset
Subset and strict subset relations #
singleton #
{a} : Finset a
is the set {a}
containing a
and nothing else.
This differs from insert a ∅
in that it does not require a DecidableEq
instance for α
.
Equations
- Finset.instSingleton = { singleton := fun (a : α) => { val := {a}, nodup := ⋯ } }
Alias of the forward direction of Finset.nonempty_iff_eq_singleton_default
.
Equations
- ⋯ = ⋯
Equations
- Finset.instUniqueSubtypeMemSingleton i = { default := ⟨i, ⋯⟩, uniq := ⋯ }
cons #
cons a s h
is the set {a} ∪ s
containing a
and the elements of s
. It is the same as
insert a s
when it is defined, but unlike insert a s
it does not require DecidableEq α
,
and the union is guaranteed to be disjoint.
Equations
- Finset.cons a s h = { val := a ::ₘ s.val, nodup := ⋯ }
Instances For
Useful in proofs by induction.
Alias of Finset.cons_nonempty
.
Split the added element of cons off a Pi type.
Equations
- Finset.consPiProd f has x = (x a ⋯, fun (i : α) (hi : i ∈ s) => x i ⋯)
Instances For
Combine a product with a pi type to pi of cons.
Equations
- Finset.prodPiCons f has x i hi = if h : i = a then cast ⋯ x.1 else x.2 i ⋯
Instances For
The equivalence between pi types on cons and the product.
Equations
- Finset.consPiProdEquiv f has = { toFun := Finset.consPiProd f has, invFun := Finset.prodPiCons f has, left_inv := ⋯, right_inv := ⋯ }
Instances For
insert #
insert a s
is the set {a} ∪ s
containing a
and the elements of s
.
Equations
- Finset.instInsert = { insert := fun (a : α) (s : Finset α) => { val := Multiset.ndinsert a s.val, nodup := ⋯ } }
A version of LawfulSingleton.insert_emptyc_eq
that works with dsimp
.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
To prove a proposition about an arbitrary Finset α
,
it suffices to prove it for the empty Finset
,
and to show that if it holds for some Finset α
,
then it holds for the Finset
obtained by inserting a new element.
To prove a proposition about S : Finset α
,
it suffices to prove it for the empty Finset
,
and to show that if it holds for some Finset α ⊆ S
,
then it holds for the Finset
obtained by inserting a new element of S
.
To prove a proposition about a nonempty s : Finset α
, it suffices to show it holds for all
singletons and that if it holds for nonempty t : Finset α
, then it also holds for the Finset
obtained by inserting an element in t
.
Inserting an element to a finite set is equivalent to the option type.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Split the added element of insert off a Pi type.
Equations
- Finset.insertPiProd f x = (x a ⋯, fun (i : α) (hi : i ∈ s) => x i ⋯)
Instances For
Combine a product with a pi type to pi of insert.
Equations
- Finset.prodPiInsert f x i hi = if h : i = a then cast ⋯ x.1 else x.2 i ⋯
Instances For
The equivalence between pi types on insert and the product.
Equations
- Finset.insertPiProdEquiv f has = { toFun := Finset.insertPiProd f, invFun := Finset.prodPiInsert f, left_inv := ⋯, right_inv := ⋯ }
Instances For
Useful in proofs by induction.