Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ba2c056

Browse files
committed
feat(data/list/nodup): nodup.pairwise_of_forall_ne (#7587)
1 parent 7021ff9 commit ba2c056

File tree

2 files changed

+62
-0
lines changed

2 files changed

+62
-0
lines changed

src/data/list/nodup.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -322,6 +322,20 @@ begin
322322
{ simp [ne.symm H, H, update_nth, ← apply_ite (cons (f hd))] }
323323
end
324324

325+
lemma nodup.pairwise_of_forall_ne {l : list α} {r : α → α → Prop}
326+
(hl : l.nodup) (h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
327+
begin
328+
classical,
329+
refine pairwise_of_reflexive_on_dupl_of_forall_ne _ h,
330+
intros x hx,
331+
rw nodup_iff_count_le_one at hl,
332+
exact absurd (hl x) hx.not_le
333+
end
334+
335+
lemma nodup.pairwise_of_set_pairwise_on {l : list α} {r : α → α → Prop}
336+
(hl : l.nodup) (h : {x | x ∈ l}.pairwise_on r) : l.pairwise r :=
337+
hl.pairwise_of_forall_ne h
338+
325339
end list
326340

327341
theorem option.to_list_nodup {α} : ∀ o : option α, o.to_list.nodup

src/data/list/pairwise.lean

Lines changed: 48 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -227,6 +227,54 @@ from λ R l, ⟨λ p, reverse_reverse l ▸ this p, this⟩,
227227
pairwise_cons, forall_prop_of_false (not_mem_nil _), forall_true_iff,
228228
pairwise.nil, mem_reverse, mem_singleton, forall_eq, true_and] using h]
229229

230+
lemma pairwise.set_pairwise_on {l : list α} (h : pairwise R l) (hr : symmetric R) :
231+
set.pairwise_on {x | x ∈ l} R :=
232+
begin
233+
induction h with hd tl imp h IH,
234+
{ simp },
235+
{ intros x hx y hy hxy,
236+
simp only [mem_cons_iff, set.mem_set_of_eq] at hx hy,
237+
rcases hx with rfl|hx;
238+
rcases hy with rfl|hy,
239+
{ contradiction },
240+
{ exact imp y hy },
241+
{ exact hr (imp x hx) },
242+
{ exact IH x hx y hy hxy } }
243+
end
244+
245+
lemma pairwise_of_reflexive_on_dupl_of_forall_ne [decidable_eq α] {l : list α} {r : α → α → Prop}
246+
(hr : ∀ a, 1 < count a l → r a a)
247+
(h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
248+
begin
249+
induction l with hd tl IH,
250+
{ simp },
251+
{ rw list.pairwise_cons,
252+
split,
253+
{ intros x hx,
254+
by_cases H : hd = x,
255+
{ rw H,
256+
refine hr _ _,
257+
simpa [count_cons, H, nat.succ_lt_succ_iff, count_pos] using hx },
258+
{ exact h hd (mem_cons_self _ _) x (mem_cons_of_mem _ hx) H } },
259+
{ refine IH _ _,
260+
{ intros x hx,
261+
refine hr _ _,
262+
rw count_cons,
263+
split_ifs,
264+
{ exact hx.trans (nat.lt_succ_self _) },
265+
{ exact hx } },
266+
{ intros x hx y hy,
267+
exact h x (mem_cons_of_mem _ hx) y (mem_cons_of_mem _ hy) } } }
268+
end
269+
270+
lemma pairwise_of_reflexive_of_forall_ne {l : list α} {r : α → α → Prop}
271+
(hr : reflexive r) (h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
272+
begin
273+
classical,
274+
refine pairwise_of_reflexive_on_dupl_of_forall_ne _ h,
275+
exact λ _ _, hr _
276+
end
277+
230278
theorem pairwise_iff_nth_le {R} : ∀ {l : list α},
231279
pairwise R l ↔ ∀ i j (h₁ : j < length l) (h₂ : i < j),
232280
R (nth_le l i (lt_trans h₂ h₁)) (nth_le l j h₁)

0 commit comments

Comments
 (0)