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]
:
@[simp]
theorem
Multiset.sort_sorted
{α : Type u_1}
(s : Multiset α)
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
List.Sorted r (s.sort r)
@[simp]
theorem
Multiset.sort_eq
{α : Type u_1}
(s : Multiset α)
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
@[simp]
theorem
Multiset.sort_zero
{α : Type u_1}
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
@[simp]
theorem
Multiset.sort_singleton
{α : Type u_1}
(a : α)
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
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 : ∀ a ∈ s, ∀ b ∈ s, r a b ↔ r' (f a) (f b))
:
theorem
Multiset.sort_cons
{α : Type u_1}
(a : α)
(s : Multiset α)
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
@[simp]
@[simp]
theorem
Multiset.mem_sort
{α : Type u_1}
{a : α}
{s : Multiset α}
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
@[simp]
theorem
Multiset.length_sort
{α : Type u_1}
{s : Multiset α}
(r : α → α → Prop)
[DecidableRel r]
[IsTrans α r]
[IsAntisymm α r]
[IsTotal α r]
:
unsafe instance
Multiset.instToExprOfToLevel
{α : Type u}
[Lean.ToLevel]
[Lean.ToExpr α]
:
Lean.ToExpr (Multiset α)