Skip to content

Commit

Permalink
feat(data/list/nodup): nodup.pairwise_of_forall_ne (#7587)
Browse files Browse the repository at this point in the history
  • Loading branch information
pechersky committed Jun 6, 2021
1 parent 7021ff9 commit ba2c056
Show file tree
Hide file tree
Showing 2 changed files with 62 additions and 0 deletions.
14 changes: 14 additions & 0 deletions src/data/list/nodup.lean
Expand Up @@ -322,6 +322,20 @@ begin
{ simp [ne.symm H, H, update_nth, ← apply_ite (cons (f hd))] }
end

lemma nodup.pairwise_of_forall_ne {l : list α} {r : α → α → Prop}
(hl : l.nodup) (h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
begin
classical,
refine pairwise_of_reflexive_on_dupl_of_forall_ne _ h,
intros x hx,
rw nodup_iff_count_le_one at hl,
exact absurd (hl x) hx.not_le
end

lemma nodup.pairwise_of_set_pairwise_on {l : list α} {r : α → α → Prop}
(hl : l.nodup) (h : {x | x ∈ l}.pairwise_on r) : l.pairwise r :=
hl.pairwise_of_forall_ne h

end list

theorem option.to_list_nodup {α} : ∀ o : option α, o.to_list.nodup
Expand Down
48 changes: 48 additions & 0 deletions src/data/list/pairwise.lean
Expand Up @@ -227,6 +227,54 @@ from λ R l, ⟨λ p, reverse_reverse l ▸ this p, this⟩,
pairwise_cons, forall_prop_of_false (not_mem_nil _), forall_true_iff,
pairwise.nil, mem_reverse, mem_singleton, forall_eq, true_and] using h]

lemma pairwise.set_pairwise_on {l : list α} (h : pairwise R l) (hr : symmetric R) :
set.pairwise_on {x | x ∈ l} R :=
begin
induction h with hd tl imp h IH,
{ simp },
{ intros x hx y hy hxy,
simp only [mem_cons_iff, set.mem_set_of_eq] at hx hy,
rcases hx with rfl|hx;
rcases hy with rfl|hy,
{ contradiction },
{ exact imp y hy },
{ exact hr (imp x hx) },
{ exact IH x hx y hy hxy } }
end

lemma pairwise_of_reflexive_on_dupl_of_forall_ne [decidable_eq α] {l : list α} {r : α → α → Prop}
(hr : ∀ a, 1 < count a l → r a a)
(h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
begin
induction l with hd tl IH,
{ simp },
{ rw list.pairwise_cons,
split,
{ intros x hx,
by_cases H : hd = x,
{ rw H,
refine hr _ _,
simpa [count_cons, H, nat.succ_lt_succ_iff, count_pos] using hx },
{ exact h hd (mem_cons_self _ _) x (mem_cons_of_mem _ hx) H } },
{ refine IH _ _,
{ intros x hx,
refine hr _ _,
rw count_cons,
split_ifs,
{ exact hx.trans (nat.lt_succ_self _) },
{ exact hx } },
{ intros x hx y hy,
exact h x (mem_cons_of_mem _ hx) y (mem_cons_of_mem _ hy) } } }
end

lemma pairwise_of_reflexive_of_forall_ne {l : list α} {r : α → α → Prop}
(hr : reflexive r) (h : ∀ (a ∈ l) (b ∈ l), a ≠ b → r a b) : l.pairwise r :=
begin
classical,
refine pairwise_of_reflexive_on_dupl_of_forall_ne _ h,
exact λ _ _, hr _
end

theorem pairwise_iff_nth_le {R} : ∀ {l : list α},
pairwise R l ↔ ∀ i j (h₁ : j < length l) (h₂ : i < j),
R (nth_le l i (lt_trans h₂ h₁)) (nth_le l j h₁)
Expand Down

0 comments on commit ba2c056

Please sign in to comment.