Documentation

Mathlib.Data.Multiset.Sort

Construct a sorted list from a multiset. #

def Multiset.sort {α : Type u_1} (s : Multiset α) (r : ααProp := by exact fun a b => a ≤ b) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
List α

sort s constructs a sorted list from the multiset s. (Uses merge sort algorithm.)

Equations
Instances For
    @[simp]
    theorem Multiset.coe_sort {α : Type u_1} (l : List α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    (↑l).sort r = l.mergeSort fun (x1 x2 : α) => decide (r x1 x2)
    @[simp]
    theorem Multiset.sort_sorted {α : Type u_1} (s : Multiset α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    @[simp]
    theorem Multiset.sort_eq {α : Type u_1} (s : Multiset α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    (s.sort r) = s
    @[simp]
    theorem Multiset.sort_zero {α : Type u_1} (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    sort 0 r = []
    @[simp]
    theorem Multiset.sort_singleton {α : Type u_1} (a : α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    {a}.sort r = [a]
    theorem Multiset.map_sort {α : Type u_1} {β : Type u_2} (f : αβ) (s : Multiset α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] (r' : ββProp) [DecidableRel r'] [IsTrans β r'] [IsAntisymm β r'] [IsTotal β r'] (hs : as, bs, r a b r' (f a) (f b)) :
    List.map f (s.sort r) = (map f s).sort r'
    theorem Multiset.sort_cons {α : Type u_1} (a : α) (s : Multiset α) (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    (∀ bs, r a b)(a ::ₘ s).sort r = a :: s.sort r
    @[simp]
    theorem Multiset.sort_range (n : ) :
    ((range n).sort fun (a b : ) => a b) = List.range n
    @[simp]
    theorem Multiset.mem_sort {α : Type u_1} {a : α} {s : Multiset α} (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    a s.sort r a s
    @[simp]
    theorem Multiset.length_sort {α : Type u_1} {s : Multiset α} (r : ααProp) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] :
    (s.sort r).length = s.card
    unsafe instance Multiset.instRepr {α : Type u_1} [Repr α] :
    Equations
    • One or more equations did not get rendered due to their size.