mathlib documentation

data.finset.sort

Construct a sorted list from a finset.

sort

def finset.sort {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] (s : finset α) :
list α

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

Equations
@[simp]
theorem finset.sort_sorted {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] (s : finset α) :

@[simp]
theorem finset.sort_eq {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] (s : finset α) :

@[simp]
theorem finset.sort_nodup {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] (s : finset α) :

@[simp]
theorem finset.sort_to_finset {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] [decidable_eq α] (s : finset α) :

@[simp]
theorem finset.mem_sort {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] {s : finset α} {a : α} :

@[simp]
theorem finset.length_sort {α : Type u_1} (r : α → α → Prop) [decidable_rel r] [is_trans α r] [is_antisymm α r] [is_total α r] {s : finset α} :

theorem finset.sorted_zero_eq_min'_aux {α : Type u_1} [decidable_linear_order α] (s : finset α) (h : 0 < (finset.sort has_le.le s).length) (H : s.nonempty) :

theorem finset.sorted_zero_eq_min' {α : Type u_1} [decidable_linear_order α] {s : finset α} {h : 0 < (finset.sort has_le.le s).length} :

theorem finset.min'_eq_sorted_zero {α : Type u_1} [decidable_linear_order α] {s : finset α} {h : s.nonempty} :

def finset.mono_of_fin {α : Type u_1} [decidable_linear_order α] (s : finset α) {k : } (h : s.card = k) (i : fin k) :
α

Given a finset s of cardinal k in a linear order α, the map mono_of_fin s h is the increasing bijection between fin k and s as an α-valued map. Here, h is a proof that the cardinality of s is k. We use this instead of a map fin s.card → α to avoid casting issues in further uses of this function.

Equations
theorem finset.mono_of_fin_strict_mono {α : Type u_1} [decidable_linear_order α] (s : finset α) {k : } (h : s.card = k) :

theorem finset.mono_of_fin_bij_on {α : Type u_1} [decidable_linear_order α] (s : finset α) {k : } (h : s.card = k) :

theorem finset.mono_of_fin_injective {α : Type u_1} [decidable_linear_order α] (s : finset α) {k : } (h : s.card = k) :

theorem finset.mono_of_fin_zero {α : Type u_1} [decidable_linear_order α] {s : finset α} {k : } (h : s.card = k) (hz : 0 < k) :
s.mono_of_fin h 0, hz⟩ = s.min' _

The bijection mono_of_fin s h sends 0 to the minimum of s.

theorem finset.mono_of_fin_last {α : Type u_1} [decidable_linear_order α] {s : finset α} {k : } (h : s.card = k) (hz : 0 < k) :
s.mono_of_fin h k - 1, _⟩ = s.max' _

The bijection mono_of_fin s h sends k-1 to the maximum of s.

@[simp]
theorem finset.mono_of_fin_singleton {α : Type u_1} [decidable_linear_order α] (a : α) (i : fin 1) {h : {a}.card = 1} :
{a}.mono_of_fin h i = a

mono_of_fin {a} h sends any argument to a.

@[simp]
theorem finset.range_mono_of_fin {α : Type u_1} [decidable_linear_order α] {s : finset α} {k : } (h : s.card = k) :

The range of mono_of_fin.

theorem finset.mono_of_fin_unique {α : Type u_1} [decidable_linear_order α] {s : finset α} {k : } (h : s.card = k) {f : fin k → α} (hbij : set.bij_on f set.univ s) (hmono : strict_mono f) :

Any increasing bijection between fin k and a finset of cardinality k has to coincide with the increasing bijection mono_of_fin s h. For a statement assuming only that f maps univ to s, see mono_of_fin_unique'.

theorem finset.mono_of_fin_unique' {α : Type u_1} [decidable_linear_order α] {s : finset α} {k : } (h : s.card = k) {f : fin k → α} (fmap : set.maps_to f set.univ s) (hmono : strict_mono f) :

Any increasing map between fin k and a finset of cardinality k has to coincide with the increasing bijection mono_of_fin s h.

@[simp]
theorem finset.mono_of_fin_eq_mono_of_fin_iff {α : Type u_1} [decidable_linear_order α] {k l : } {s : finset α} {i : fin k} {j : fin l} {h : s.card = k} {h' : s.card = l} :

Two parametrizations mono_of_fin of the same set take the same value on i and j if and only if i = j. Since they can be defined on a priori not defeq types fin k and fin l (although necessarily k = l), the conclusion is rather written (i : ℕ) = (j : ℕ).

def finset.mono_equiv_of_fin {α : Type u_1} [decidable_linear_order α] (s : finset α) {k : } (h : s.card = k) :
fin k {x // x s}

Given a finset s of cardinal k in a linear order α, the equiv mono_equiv_of_fin s h is the increasing bijection between fin k and s as an s-valued map. Here, h is a proof that the cardinality of s is k. We use this instead of a map fin s.card → α to avoid casting issues in further uses of this function.

Equations
@[instance]
def finset.has_repr {α : Type u_1} [has_repr α] :

Equations