List duplicates #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Main definitions #
list.duplicate x l : Prop
is an inductive property that holds whenx
is a duplicate inl
Implementation details #
In this file, x ∈+ l
notation is shorthand for list.duplicate x l
.
- cons_mem : ∀ {α : Type u_1} {x : α} {l : list α}, x ∈ l → list.duplicate x (x :: l)
- cons_duplicate : ∀ {α : Type u_1} {x y : α} {l : list α}, list.duplicate x l → list.duplicate x (y :: l)
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) :
list.duplicate x (x :: l)
theorem
list.duplicate.duplicate_cons
{α : Type u_1}
{l : list α}
{x : α}
(h : list.duplicate x l)
(y : α) :
list.duplicate x (y :: 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_singleton
{α : Type u_1}
{l : list α}
{x : α}
(h : list.duplicate x l)
(y : α) :
@[simp]
theorem
list.duplicate_cons_iff
{α : Type u_1}
{l : list α}
{x y : α} :
list.duplicate x (y :: l) ↔ y = x ∧ x ∈ l ∨ list.duplicate x l
theorem
list.duplicate.of_duplicate_cons
{α : Type u_1}
{l : list α}
{x y : α}
(h : list.duplicate x (y :: l))
(hx : x ≠ y) :
list.duplicate x l
theorem
list.duplicate_cons_iff_of_ne
{α : Type u_1}
{l : list α}
{x y : α}
(hne : x ≠ y) :
list.duplicate x (y :: l) ↔ list.duplicate x l
theorem
list.duplicate.mono_sublist
{α : Type u_1}
{l : list α}
{x : α}
{l' : list α}
(hx : list.duplicate x l)
(h : l <+ l') :
list.duplicate x 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.duplicate_iff_two_le_count
{α : Type u_1}
{l : list α}
{x : α}
[decidable_eq α] :
list.duplicate x l ↔ 2 ≤ list.count x l
@[protected, instance]
def
list.decidable_duplicate
{α : Type u_1}
[decidable_eq α]
(x : α)
(l : list α) :
decidable (list.duplicate x l)
Equations
- list.decidable_duplicate x (y :: l) = list.decidable_duplicate._match_1 x y l (list.decidable_duplicate x l)
- list.decidable_duplicate x list.nil = decidable.is_false _
- list.decidable_duplicate._match_1 x y l (decidable.is_true h) = decidable.is_true _
- list.decidable_duplicate._match_1 x y l (decidable.is_false h) = dite (y = x ∧ x ∈ l) (λ (hx : y = x ∧ x ∈ l), decidable.is_true _) (λ (hx : ¬(y = x ∧ x ∈ l)), decidable.is_false _)