mathlib3 documentation

data.list.duplicate

List duplicates #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

Main definitions #

Implementation details #

In this file, x ∈+ l notation is shorthand for list.duplicate x l.

inductive list.duplicate {α : Type u_1} (x : α) :
list α Prop

Property that an element x : α of l : list α can be found in the list more than once.

Instances for list.duplicate
theorem list.mem.duplicate_cons_self {α : Type u_1} {l : list α} {x : α} (h : x l) :
theorem list.duplicate.duplicate_cons {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x l) (y : α) :
theorem list.duplicate.mem {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x l) :
x l
theorem list.duplicate.mem_cons_self {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x (x :: l)) :
x l
@[simp]
theorem list.duplicate_cons_self_iff {α : Type u_1} {l : list α} {x : α} :
list.duplicate x (x :: l) x l
theorem list.duplicate.ne_nil {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x l) :
@[simp]
theorem list.not_duplicate_nil {α : Type u_1} (x : α) :
theorem list.duplicate.ne_singleton {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x l) (y : α) :
l [y]
@[simp]
theorem list.not_duplicate_singleton {α : Type u_1} (x y : α) :
theorem list.duplicate.elim_nil {α : Type u_1} {x : α} (h : list.duplicate x list.nil) :
theorem list.duplicate.elim_singleton {α : Type u_1} {x y : α} (h : list.duplicate x [y]) :
theorem list.duplicate_cons_iff {α : Type u_1} {l : list α} {x y : α} :
theorem list.duplicate.of_duplicate_cons {α : Type u_1} {l : list α} {x y : α} (h : list.duplicate x (y :: l)) (hx : x y) :
theorem list.duplicate_cons_iff_of_ne {α : Type u_1} {l : list α} {x y : α} (hne : x y) :
theorem list.duplicate.mono_sublist {α : Type u_1} {l : list α} {x : α} {l' : list α} (hx : list.duplicate x l) (h : l <+ l') :
theorem list.duplicate_iff_sublist {α : Type u_1} {l : list α} {x : α} :
list.duplicate x l [x, x] <+ l

The contrapositive of list.nodup_iff_sublist.

theorem list.nodup_iff_forall_not_duplicate {α : Type u_1} {l : list α} :
l.nodup (x : α), ¬list.duplicate x l
theorem list.exists_duplicate_iff_not_nodup {α : Type u_1} {l : list α} :
( (x : α), list.duplicate x l) ¬l.nodup
theorem list.duplicate.not_nodup {α : Type u_1} {l : list α} {x : α} (h : list.duplicate x l) :
theorem list.duplicate_iff_two_le_count {α : Type u_1} {l : list α} {x : α} [decidable_eq α] :
@[protected, instance]
def list.decidable_duplicate {α : Type u_1} [decidable_eq α] (x : α) (l : list α) :
Equations