Documentation

Mathlib.Data.Multiset.Fintype

Multiset coercion to type #

This module defines a CoeSort instance for multisets and gives it a Fintype instance. It also defines Multiset.toEnumFinset, which is another way to enumerate the elements of a multiset. These coercions and definitions make it easier to sum over multisets using existing Finset theory.

Main definitions #

Tags #

multiset enumeration

def Multiset.ToType {α : Type u_1} [DecidableEq α] (m : Multiset α) :
Type u_1

Auxiliary definition for the CoeSort instance. This prevents the CoeOut m α instance from inadvertently applying to other sigma types.

Equations
Instances For
    instance Multiset.instCoeSortType {α : Type u_1} [DecidableEq α] :

    Create a type that has the same number of elements as the multiset. Terms of this type are triples ⟨x, ⟨i, h⟩⟩ where x : α, i : ℕ, and h : i < m.count x. This way repeated elements of a multiset appear multiple times from different values of i.

    Equations
    @[reducible, match_pattern]
    def Multiset.mkToType {α : Type u_1} [DecidableEq α] (m : Multiset α) (x : α) (i : Fin (count x m)) :

    Constructor for terms of the coercion of m to a type. This helps Lean pick up the correct instances.

    Equations
    Instances For

      As a convenience, there is a coercion from m : Type* to α by projecting onto the first component.

      Equations
      theorem Multiset.coe_mk {α : Type u_1} [DecidableEq α] {m : Multiset α} {x : α} {i : Fin (count x m)} :
      (m.mkToType x i).fst = x
      @[simp]
      theorem Multiset.coe_mem {α : Type u_1} [DecidableEq α] {m : Multiset α} {x : m.ToType} :
      x.fst m
      @[simp]
      theorem Multiset.forall_coe {α : Type u_1} [DecidableEq α] {m : Multiset α} (p : m.ToTypeProp) :
      (∀ (x : m.ToType), p x) ∀ (x : α) (i : Fin (count x m)), p x, i
      @[simp]
      theorem Multiset.exists_coe {α : Type u_1} [DecidableEq α] {m : Multiset α} (p : m.ToTypeProp) :
      (∃ (x : m.ToType), p x) ∃ (x : α) (i : Fin (count x m)), p x, i
      def Multiset.toEnumFinset {α : Type u_1} [DecidableEq α] (m : Multiset α) :

      Construct a finset whose elements enumerate the elements of the multiset m. The component is used to differentiate between equal elements: if x appears n times then (x, 0), ..., and (x, n-1) appear in the Finset.

      Equations
      Instances For
        @[simp]
        theorem Multiset.mem_toEnumFinset {α : Type u_1} [DecidableEq α] (m : Multiset α) (p : α × ) :
        p m.toEnumFinset p.2 < count p.1 m
        theorem Multiset.mem_of_mem_toEnumFinset {α : Type u_1} [DecidableEq α] {m : Multiset α} {p : α × } (h : p m.toEnumFinset) :
        p.1 m
        @[simp]
        theorem Multiset.toEnumFinset_filter_eq {α : Type u_1} [DecidableEq α] (m : Multiset α) (a : α) :
        {xm.toEnumFinset | x.1 = a} = {a} ×ˢ Finset.range (count a m)
        @[simp]
        theorem Multiset.map_fst_le_of_subset_toEnumFinset {α : Type u_1} [DecidableEq α] {m : Multiset α} {s : Finset (α × )} (hsm : s m.toEnumFinset) :
        theorem Multiset.toEnumFinset_mono {α : Type u_1} [DecidableEq α] {m₁ m₂ : Multiset α} (h : m₁ m₂) :
        @[simp]
        theorem Multiset.toEnumFinset_subset_iff {α : Type u_1} [DecidableEq α] {m₁ m₂ : Multiset α} :
        m₁.toEnumFinset m₂.toEnumFinset m₁ m₂
        def Multiset.coeEmbedding {α : Type u_1} [DecidableEq α] (m : Multiset α) :

        The embedding from a multiset into α × ℕ where the second coordinate enumerates repeats. If you are looking for the function m → α, that would be plain (↑).

        Equations
        Instances For
          @[simp]
          theorem Multiset.coeEmbedding_apply {α : Type u_1} [DecidableEq α] (m : Multiset α) (x : m.ToType) :
          def Multiset.coeEquiv {α : Type u_1} [DecidableEq α] (m : Multiset α) :

          Another way to coerce a Multiset to a type is to go through m.toEnumFinset and coerce that Finset to a type.

          Equations
          Instances For
            @[simp]
            theorem Multiset.coeEquiv_apply_coe {α : Type u_1} [DecidableEq α] (m : Multiset α) (x : m.ToType) :
            (m.coeEquiv x) = m.coeEmbedding x
            @[simp]
            theorem Multiset.coeEquiv_symm_apply_fst {α : Type u_1} [DecidableEq α] (m : Multiset α) (x : m.toEnumFinset) :
            (m.coeEquiv.symm x).fst = (↑x).1
            @[simp]
            theorem Multiset.coeEquiv_symm_apply_snd_val {α : Type u_1} [DecidableEq α] (m : Multiset α) (x : m.toEnumFinset) :
            (m.coeEquiv.symm x).snd = (↑x).2
            @[irreducible]
            instance Multiset.fintypeCoe {α : Type u_1} [DecidableEq α] {m : Multiset α} :
            Equations
            @[simp]
            theorem Multiset.map_univ_coe {α : Type u_1} [DecidableEq α] (m : Multiset α) :
            map (fun (x : m.ToType) => x.fst) Finset.univ.val = m
            theorem Multiset.map_univ_comp_coe {α : Type u_1} [DecidableEq α] {β : Type u_3} (m : Multiset α) (f : αβ) :
            map (f fun (x : m.ToType) => x.fst) Finset.univ.val = map f m
            @[simp]
            theorem Multiset.map_univ {α : Type u_1} [DecidableEq α] {β : Type u_3} (m : Multiset α) (f : αβ) :
            map (fun (x : m.ToType) => f x.fst) Finset.univ.val = map f m
            @[simp]
            @[simp]
            theorem Multiset.card_coe {α : Type u_1} [DecidableEq α] (m : Multiset α) :
            theorem Multiset.prod_eq_prod_coe {α : Type u_1} [DecidableEq α] [CommMonoid α] (m : Multiset α) :
            m.prod = x : m.ToType, x.fst
            theorem Multiset.sum_eq_sum_coe {α : Type u_1} [DecidableEq α] [AddCommMonoid α] (m : Multiset α) :
            m.sum = x : m.ToType, x.fst
            theorem Multiset.prod_eq_prod_toEnumFinset {α : Type u_1} [DecidableEq α] [CommMonoid α] (m : Multiset α) :
            m.prod = xm.toEnumFinset, x.1
            theorem Multiset.sum_eq_sum_toEnumFinset {α : Type u_1} [DecidableEq α] [AddCommMonoid α] (m : Multiset α) :
            m.sum = xm.toEnumFinset, x.1
            theorem Multiset.prod_toEnumFinset {α : Type u_1} [DecidableEq α] {β : Type u_3} [CommMonoid β] (m : Multiset α) (f : αβ) :
            xm.toEnumFinset, f x.1 x.2 = x : m.ToType, f x.fst x.snd
            theorem Multiset.sum_toEnumFinset {α : Type u_1} [DecidableEq α] {β : Type u_3} [AddCommMonoid β] (m : Multiset α) (f : αβ) :
            xm.toEnumFinset, f x.1 x.2 = x : m.ToType, f x.fst x.snd
            def Multiset.cast {α : Type u_1} [DecidableEq α] {s t : Multiset α} (h : s = t) :

            If s = t then there's an equivalence between the appropriate types.

            Equations
            Instances For
              @[simp]
              theorem Multiset.cast_symm_apply_fst {α : Type u_1} [DecidableEq α] {s t : Multiset α} (h : s = t) (x : t.ToType) :
              ((cast h).symm x).fst = x.fst
              @[simp]
              theorem Multiset.cast_apply_snd {α : Type u_1} [DecidableEq α] {s t : Multiset α} (h : s = t) (x : s.ToType) :
              ((cast h) x).snd = Fin.cast x.snd
              @[simp]
              theorem Multiset.cast_symm_apply_snd {α : Type u_1} [DecidableEq α] {s t : Multiset α} (h : s = t) (x : t.ToType) :
              ((cast h).symm x).snd = Fin.cast x.snd
              @[simp]
              theorem Multiset.cast_apply_fst {α : Type u_1} [DecidableEq α] {s t : Multiset α} (h : s = t) (x : s.ToType) :
              ((cast h) x).fst = x.fst
              def Multiset.consEquiv {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} :

              v ::ₘ m is equivalent to Option m by mapping one v to none and everything else to m.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem Multiset.consEquiv_symm_none {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} :
                @[simp]
                theorem Multiset.consEquiv_symm_some {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} {x : m.ToType} :
                theorem Multiset.coe_consEquiv_of_ne {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} (x : (v ::ₘ m).ToType) (hx : x.fst v) :
                theorem Multiset.coe_consEquiv_of_eq_of_eq {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} (x : (v ::ₘ m).ToType) (hx : x.fst = v) (hx2 : x.snd = count v m) :
                theorem Multiset.coe_consEquiv_of_eq_of_lt {α : Type u_1} [DecidableEq α] {m : Multiset α} {v : α} (x : (v ::ₘ m).ToType) (hx : x.fst = v) (hx2 : x.snd < count v m) :
                def Multiset.mapEquiv_aux {α : Type u_1} {β : Type u_2} [DecidableEq α] [DecidableEq β] (m : Multiset α) (f : αβ) :
                Squash { v : m.ToType (map f m).ToType // ∀ (a : m.ToType), (v a).fst = f a.fst }

                There is some equivalence between m and m.map f which respects f.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  noncomputable def Multiset.mapEquiv {α : Type u_1} {β : Type u_2} [DecidableEq α] [DecidableEq β] (s : Multiset α) (f : αβ) :

                  One of the possible equivalences from Multiset.mapEquiv_aux, selected using choice.

                  Equations
                  Instances For
                    @[simp]
                    theorem Multiset.mapEquiv_apply {α : Type u_1} {β : Type u_2} [DecidableEq α] [DecidableEq β] (s : Multiset α) (f : αβ) (v : s.ToType) :
                    ((s.mapEquiv f) v).fst = f v.fst