Symmetric powers of a finset #
This file defines the symmetric powers of a finset as Finset (Sym α n)
and Finset (Sym2 α)
.
Main declarations #
Finset.sym
: The symmetric power of a finset.s.sym n
is all the multisets of cardinalityn
whose elements are ins
.Finset.sym2
: The symmetric square of a finset.s.sym2
is all the pairs whose elements are ins
.- A
Fintype (Sym2 α)
instance that does not requireDecidableEq α
.
TODO #
Finset.sym
forms a Galois connection between Finset α
and Finset (Sym α n)
. Similar for
Finset.sym2
.
Equations
- Sym2.instFintype = { elems := Finset.univ.sym2, complete := ⋯ }
Alias of the reverse direction of Finset.sym2_nonempty
.
theorem
Finset.isDiag_mk_of_mem_diag
{α : Type u_1}
{s : Finset α}
[DecidableEq α]
{a : α × α}
(h : a ∈ s.diag)
:
Equations
- Finset.instDecidableEqSym = inferInstanceAs (DecidableEq { s : Multiset α // s.card = n })
theorem
Finset.replicate_mem_sym
{α : Type u_1}
{s : Finset α}
{a : α}
[DecidableEq α]
(ha : a ∈ s)
(n : ℕ)
:
theorem
Finset.Nonempty.sym
{α : Type u_1}
{s : Finset α}
[DecidableEq α]
(h : s.Nonempty)
(n : ℕ)
:
@[simp]
@[simp]
@[simp]
theorem
Finset.sym_filterNe_mem
{α : Type u_1}
{s : Finset α}
[DecidableEq α]
{n : ℕ}
{m : Sym α n}
(a : α)
(h : m ∈ s.sym n)
:
def
Finset.symInsertEquiv
{α : Type u_1}
{s : Finset α}
{a : α}
[DecidableEq α]
{n : ℕ}
(h : a ∉ s)
:
If a
does not belong to the finset s
, then the n
th symmetric power of {a} ∪ s
is
in 1-1 correspondence with the disjoint union of the n - i
th symmetric powers of s
,
for 0 ≤ i ≤ n
.
Equations
- One or more equations did not get rendered due to their size.