Documentation

Mathlib.Data.Finsupp.Multiset

Equivalence between Multiset and -valued finitely supported functions #

This defines Finsupp.toMultiset the equivalence between α →₀ ℕ and Multiset α, along with Multiset.toFinsupp the reverse equivalence and Finsupp.orderIsoMultiset (the equivalence promoted to an order isomorphism).

def Finsupp.toMultiset {α : Type u_1} :

Given f : α →₀ ℕ, f.toMultiset is the multiset with multiplicities given by the values of f on the elements of α. We define this function as an AddMonoidHom.

Under the additional assumption of [DecidableEq α], this is available as Multiset.toFinsupp : Multiset α ≃+ (α →₀ ℕ); the two declarations are separate as this assumption is only needed for one direction.

Equations
Instances For
    theorem Finsupp.toMultiset_add {α : Type u_1} (m n : α →₀ ) :
    theorem Finsupp.toMultiset_apply {α : Type u_1} (f : α →₀ ) :
    toMultiset f = f.sum fun (a : α) (n : ) => n {a}
    @[simp]
    theorem Finsupp.toMultiset_single {α : Type u_1} (a : α) (n : ) :
    theorem Finsupp.toMultiset_sum {α : Type u_1} {ι : Type u_3} {f : ια →₀ } (s : Finset ι) :
    toMultiset (∑ is, f i) = is, toMultiset (f i)
    theorem Finsupp.toMultiset_sum_single {ι : Type u_3} (s : Finset ι) (n : ) :
    toMultiset (∑ is, single i n) = n s.val
    @[simp]
    theorem Finsupp.card_toMultiset {α : Type u_1} (f : α →₀ ) :
    (toMultiset f).card = f.sum fun (x : α) => id
    theorem Finsupp.toMultiset_map {α : Type u_1} {β : Type u_2} (f : α →₀ ) (g : αβ) :
    @[simp]
    theorem Finsupp.prod_toMultiset {α : Type u_1} [CommMonoid α] (f : α →₀ ) :
    (toMultiset f).prod = f.prod fun (a : α) (n : ) => a ^ n
    @[simp]
    theorem Finsupp.sum_toMultiset {α : Type u_1} [AddCommMonoid α] (f : α →₀ ) :
    (toMultiset f).sum = f.sum fun (a : α) (n : ) => n a
    @[simp]
    @[simp]
    theorem Finsupp.count_toMultiset {α : Type u_1} [DecidableEq α] (f : α →₀ ) (a : α) :
    @[simp]
    theorem Finsupp.mem_toMultiset {α : Type u_1} (f : α →₀ ) (i : α) :

    Given a multiset s, s.toFinsupp returns the finitely supported function on given by the multiplicities of the elements of s.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      @[simp]
      theorem Multiset.toFinsupp_apply {α : Type u_1} [DecidableEq α] (s : Multiset α) (a : α) :
      (toFinsupp s) a = count a s
      theorem Multiset.toFinsupp_add {α : Type u_1} [DecidableEq α] (s t : Multiset α) :
      @[simp]
      @[simp]
      theorem Multiset.toFinsupp_sum_eq {α : Type u_1} [DecidableEq α] (s : Multiset α) :
      ((toFinsupp s).sum fun (x : α) => id) = s.card

      As an order isomorphism #

      Finsupp.toMultiset as an order isomorphism.

      Equations
      Instances For
        theorem Finsupp.sum_id_lt_of_lt {ι : Type u_3} (m n : ι →₀ ) (h : m < n) :
        (m.sum fun (x : ι) => id) < n.sum fun (x : ι) => id
        theorem Finsupp.lt_wf (ι : Type u_3) :

        The order on ι →₀ ℕ is well-founded.

        Equations
        def Sym.equivNatSum (α : Type u_1) [DecidableEq α] (n : ) :
        Sym α n { P : α →₀ // (P.sum fun (x : α) => id) = n }

        The nth symmetric power of a type α is naturally equivalent to the subtype of finitely-supported maps α →₀ ℕ with total mass n.

        See also Sym.equivNatSumOfFintype when α is finite.

        Equations
        Instances For
          @[simp]
          theorem Sym.coe_equivNatSum_apply_apply (α : Type u_1) [DecidableEq α] (n : ) (s : Sym α n) (a : α) :
          ((equivNatSum α n) s) a = Multiset.count a s
          @[simp]
          theorem Sym.coe_equivNatSum_symm_apply (α : Type u_1) [DecidableEq α] (n : ) (P : { P : α →₀ // (P.sum fun (x : α) => id) = n }) :
          noncomputable def Sym.equivNatSumOfFintype (α : Type u_1) [DecidableEq α] (n : ) [Fintype α] :
          Sym α n { P : α // i : α, P i = n }

          The nth symmetric power of a finite type α is naturally equivalent to the subtype of maps α → ℕ with total mass n.

          See also Sym.equivNatSum when α is not necessarily finite.

          Equations
          Instances For
            @[simp]
            theorem Sym.coe_equivNatSumOfFintype_apply_apply (α : Type u_1) [DecidableEq α] (n : ) [Fintype α] (s : Sym α n) (a : α) :
            ((equivNatSumOfFintype α n) s) a = Multiset.count a s
            @[simp]
            theorem Sym.coe_equivNatSumOfFintype_symm_apply (α : Type u_1) [DecidableEq α] (n : ) [Fintype α] (P : { P : α // i : α, P i = n }) :
            ((equivNatSumOfFintype α n).symm P) = a : α, P a {a}