Skip to content

Commit b5ff35f

Browse files
mo271jcommelin
andcommitted
feat: port Mathlib.Data.List.Duplicate (#1465)
[![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/) Co-authored-by: Moritz Firsching <firsching@google.com> Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent 1d47ea1 commit b5ff35f

File tree

2 files changed

+166
-0
lines changed

2 files changed

+166
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -251,6 +251,7 @@ import Mathlib.Data.List.Card
251251
import Mathlib.Data.List.Chain
252252
import Mathlib.Data.List.Count
253253
import Mathlib.Data.List.Defs
254+
import Mathlib.Data.List.Duplicate
254255
import Mathlib.Data.List.Forall2
255256
import Mathlib.Data.List.Func
256257
import Mathlib.Data.List.Infix

Mathlib/Data/List/Duplicate.lean

Lines changed: 165 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,165 @@
1+
/-
2+
Copyright (c) 2021 Yakov Pechersky. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yakov Pechersky, Chris Hughes
5+
6+
! This file was ported from Lean 3 source module data.list.duplicate
7+
! leanprover-community/mathlib commit 7b78d1776212a91ecc94cf601f83bdcc46b04213
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Data.List.Nodup
12+
13+
/-!
14+
# List duplicates
15+
16+
## Main definitions
17+
18+
* `List.Duplicate x l : Prop` is an inductive property that holds when `x` is a duplicate in `l`
19+
20+
## Implementation details
21+
22+
In this file, `x ∈+ l` notation is shorthand for `List.Duplicate x l`.
23+
24+
-/
25+
26+
27+
variable {α : Type _}
28+
29+
namespace List
30+
31+
/-- Property that an element `x : α` of `l : list α` can be found in the list more than once. -/
32+
inductive Duplicate (x : α) : List α → Prop
33+
| cons_mem {l : List α} : x ∈ l → Duplicate x (x :: l)
34+
| cons_duplicate {y : α} {l : List α} : Duplicate x l → Duplicate x (y :: l)
35+
#align list.duplicate List.Duplicate
36+
37+
-- mathport name: «expr ∈+ »
38+
local infixl:50 " ∈+ " => List.Duplicate
39+
40+
variable {l : List α} {x : α}
41+
42+
theorem Mem.duplicate_cons_self (h : x ∈ l) : x ∈+ x :: l :=
43+
Duplicate.cons_mem h
44+
#align list.mem.duplicate_cons_self List.Mem.duplicate_cons_self
45+
46+
theorem Duplicate.duplicate_cons (h : x ∈+ l) (y : α) : x ∈+ y :: l :=
47+
Duplicate.cons_duplicate h
48+
#align list.duplicate.duplicate_cons List.Duplicate.duplicate_cons
49+
50+
theorem Duplicate.mem (h : x ∈+ l) : x ∈ l :=
51+
by
52+
induction' h with l' _ y l' _ hm
53+
· exact mem_cons_self _ _
54+
· exact mem_cons_of_mem _ hm
55+
#align list.duplicate.mem List.Duplicate.mem
56+
57+
theorem Duplicate.mem_cons_self (h : x ∈+ x :: l) : x ∈ l :=
58+
by
59+
cases' h with _ h _ _ h
60+
· exact h
61+
· exact h.mem
62+
#align list.duplicate.mem_cons_self List.Duplicate.mem_cons_self
63+
64+
@[simp]
65+
theorem duplicate_cons_self_iff : x ∈+ x :: l ↔ x ∈ l :=
66+
⟨Duplicate.mem_cons_self, Mem.duplicate_cons_self⟩
67+
#align list.duplicate_cons_self_iff List.duplicate_cons_self_iff
68+
69+
theorem Duplicate.ne_nil (h : x ∈+ l) : l ≠ [] := fun H => (mem_nil_iff x).mp (H ▸ h.mem)
70+
#align list.duplicate.ne_nil List.Duplicate.ne_nil
71+
72+
@[simp]
73+
theorem not_duplicate_nil (x : α) : ¬x ∈+ [] := fun H => H.ne_nil rfl
74+
#align list.not_duplicate_nil List.not_duplicate_nil
75+
76+
theorem Duplicate.ne_singleton (h : x ∈+ l) (y : α) : l ≠ [y] :=
77+
by
78+
induction' h with l' h z l' h _
79+
· simp [ne_nil_of_mem h]
80+
· simp [ne_nil_of_mem h.mem]
81+
#align list.duplicate.ne_singleton List.Duplicate.ne_singleton
82+
83+
@[simp]
84+
theorem not_duplicate_singleton (x y : α) : ¬x ∈+ [y] := fun H => H.ne_singleton _ rfl
85+
#align list.not_duplicate_singleton List.not_duplicate_singleton
86+
87+
theorem Duplicate.elim_nil (h : x ∈+ []) : False :=
88+
not_duplicate_nil x h
89+
#align list.duplicate.elim_nil List.Duplicate.elim_nil
90+
91+
theorem Duplicate.elim_singleton {y : α} (h : x ∈+ [y]) : False :=
92+
not_duplicate_singleton x y h
93+
#align list.duplicate.elim_singleton List.Duplicate.elim_singleton
94+
95+
theorem duplicate_cons_iff {y : α} : x ∈+ y :: l ↔ y = x ∧ x ∈ l ∨ x ∈+ l :=
96+
by
97+
refine' ⟨fun h => _, fun h => _⟩
98+
· cases' h with _ hm _ _ hm
99+
· exact Or.inl ⟨rfl, hm⟩
100+
· exact Or.inr hm
101+
· rcases h with (⟨rfl | h⟩ | h)
102+
· simpa
103+
· exact h.cons_duplicate
104+
#align list.duplicate_cons_iff List.duplicate_cons_iff
105+
106+
theorem Duplicate.of_duplicate_cons {y : α} (h : x ∈+ y :: l) (hx : x ≠ y) : x ∈+ l := by
107+
simpa [duplicate_cons_iff, hx.symm] using h
108+
#align list.duplicate.of_duplicate_cons List.Duplicate.of_duplicate_cons
109+
110+
theorem duplicate_cons_iff_of_ne {y : α} (hne : x ≠ y) : x ∈+ y :: l ↔ x ∈+ l := by
111+
simp [duplicate_cons_iff, hne.symm]
112+
#align list.duplicate_cons_iff_of_ne List.duplicate_cons_iff_of_ne
113+
114+
theorem Duplicate.mono_sublist {l' : List α} (hx : x ∈+ l) (h : l <+ l') : x ∈+ l' :=
115+
by
116+
induction' h with l₁ l₂ y _ IH l₁ l₂ y h IH
117+
· exact hx
118+
· exact (IH hx).duplicate_cons _
119+
· rw [duplicate_cons_iff] at hx⊢
120+
rcases hx with (⟨rfl, hx⟩ | hx)
121+
· simp [h.subset hx]
122+
· simp [IH hx]
123+
#align list.duplicate.mono_sublist List.Duplicate.mono_sublist
124+
125+
/-- The contrapositive of `List.nodup_iff_sublist`. -/
126+
theorem duplicate_iff_sublist : x ∈+ l ↔ [x, x] <+ l :=
127+
by
128+
induction' l with y l IH
129+
· simp
130+
· by_cases hx : x = y
131+
· simp [hx, cons_sublist_cons_iff, singleton_sublist]
132+
· rw [duplicate_cons_iff_of_ne hx, IH]
133+
refine' ⟨sublist_cons_of_sublist y, fun h => _⟩
134+
cases h
135+
· assumption
136+
· contradiction
137+
#align list.duplicate_iff_sublist List.duplicate_iff_sublist
138+
139+
theorem nodup_iff_forall_not_duplicate : Nodup l ↔ ∀ x : α, ¬x ∈+ l := by
140+
simp_rw [nodup_iff_sublist, duplicate_iff_sublist]
141+
#align list.nodup_iff_forall_not_duplicate List.nodup_iff_forall_not_duplicate
142+
143+
theorem exists_duplicate_iff_not_nodup : (∃ x : α, x ∈+ l) ↔ ¬Nodup l := by
144+
simp [nodup_iff_forall_not_duplicate]
145+
#align list.exists_duplicate_iff_not_nodup List.exists_duplicate_iff_not_nodup
146+
147+
theorem Duplicate.not_nodup (h : x ∈+ l) : ¬Nodup l := fun H =>
148+
nodup_iff_forall_not_duplicate.mp H _ h
149+
#align list.duplicate.not_nodup List.Duplicate.not_nodup
150+
151+
theorem duplicate_iff_two_le_count [DecidableEq α] : x ∈+ l ↔ 2 ≤ count x l := by
152+
simp [duplicate_iff_sublist, le_count_iff_replicate_sublist]
153+
#align list.duplicate_iff_two_le_count List.duplicate_iff_two_le_count
154+
155+
instance decidableDuplicate [DecidableEq α] (x : α) : ∀ l : List α, Decidable (x ∈+ l)
156+
| [] => isFalse (not_duplicate_nil x)
157+
| y :: l =>
158+
match decidableDuplicate x l with
159+
| isTrue h => isTrue (h.duplicate_cons y)
160+
| isFalse h =>
161+
if hx : y = x ∧ x ∈ l then isTrue (hx.left.symm ▸ List.Mem.duplicate_cons_self hx.right)
162+
else isFalse (by simpa [duplicate_cons_iff, h] using hx)
163+
#align list.decidable_duplicate List.decidableDuplicate
164+
165+
end List

0 commit comments

Comments
 (0)