doc-next-gen

Init.Data.List.Pairwise

Module docstring

{"# Lemmas about List.Pairwise and List.Nodup. ","## Pairwise and Nodup ","### Pairwise ","### Nodup "}

List.Pairwise.sublist theorem
: l₁ <+ l₂ → l₂.Pairwise R → l₁.Pairwise R
Full source
theorem Pairwise.sublist : l₁ <+ l₂ → l₂.Pairwise R → l₁.Pairwise R
  | .slnil, h => h
  | .cons _ s, .cons _ h₂ => h₂.sublist s
  | .cons₂ _ s, .cons h₁ h₂ => (h₂.sublist s).cons fun _ h => h₁ _ (s.subset h)
Preservation of Pairwise Relation under Sublist Operation
Informal description
For any two lists `l₁` and `l₂` of elements of type `α`, if `l₁` is a sublist of `l₂` (denoted `l₁ <+ l₂`) and the relation `R` holds pairwise for all elements in `l₂`, then the relation `R` also holds pairwise for all elements in `l₁`. In other words, if every pair of elements in `l₂` satisfies `R` (with the first element appearing before the second), and `l₁` is obtained by deleting some elements from `l₂`, then every pair of elements in `l₁` also satisfies `R`.
List.Pairwise.imp theorem
{α R S} (H : ∀ {a b}, R a b → S a b) : ∀ {l : List α}, l.Pairwise R → l.Pairwise S
Full source
theorem Pairwise.imp {α R S} (H : ∀ {a b}, R a b → S a b) :
    ∀ {l : List α}, l.Pairwise R → l.Pairwise S
  | _, .nil => .nil
  | _, .cons h₁ h₂ => .cons (H ∘ h₁ ·) (h₂.imp H)
Pairwise Relation Implication for Lists
Informal description
For any relation $R$ and $S$ on a type $\alpha$, if $R(a, b)$ implies $S(a, b)$ for all $a, b \in \alpha$, then for any list $l$ of elements of type $\alpha$, if $l$ satisfies the pairwise relation $R$, then $l$ also satisfies the pairwise relation $S$.
List.rel_of_pairwise_cons theorem
(p : (a :: l).Pairwise R) : ∀ {a'}, a' ∈ l → R a a'
Full source
theorem rel_of_pairwise_cons (p : (a :: l).Pairwise R) : ∀ {a'}, a' ∈ l → R a a' :=
  (pairwise_cons.1 p).1 _
Pairwise Relation on Cons List Implies Relation Between Head and All Tail Elements
Informal description
For any relation $R$, element $a$, and list $l$, if the list $a :: l$ satisfies the pairwise relation $R$, then for every element $a'$ in $l$, the relation $R(a, a')$ holds.
List.Pairwise.of_cons theorem
(p : (a :: l).Pairwise R) : Pairwise R l
Full source
theorem Pairwise.of_cons (p : (a :: l).Pairwise R) : Pairwise R l :=
  (pairwise_cons.1 p).2
Tail of Pairwise List is Pairwise
Informal description
If a list with head $a$ and tail $l$ satisfies the pairwise relation $R$, then the tail $l$ itself satisfies the pairwise relation $R$.
List.Pairwise.tail theorem
: ∀ {l : List α} (h : Pairwise R l), Pairwise R l.tail
Full source
theorem Pairwise.tail : ∀ {l : List α} (h : Pairwise R l), Pairwise R l.tail
  | [], h => h
  | _ :: _, h => h.of_cons
Tail of a Pairwise List is Pairwise
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $R$ on $\alpha$, if $l$ satisfies the pairwise relation $R$ (i.e., for every pair of elements in $l$ where one appears before the other, $R$ holds between them), then the tail of $l$ (the list obtained by removing the first element) also satisfies the pairwise relation $R$.
List.Pairwise.imp_of_mem theorem
{S : α → α → Prop} (H : ∀ {a b}, a ∈ l → b ∈ l → R a b → S a b) (p : Pairwise R l) : Pairwise S l
Full source
theorem Pairwise.imp_of_mem {S : α → α → Prop}
    (H : ∀ {a b}, a ∈ lb ∈ l → R a b → S a b) (p : Pairwise R l) : Pairwise S l := by
  induction p with
  | nil => constructor
  | @cons a l r _ ih =>
    constructor
    · exact fun x h => H (mem_cons_self ..) (mem_cons_of_mem _ h) <| r x h
    · exact ih fun m m' => H (mem_cons_of_mem _ m) (mem_cons_of_mem _ m')
Implication of Pairwise Relations under Membership Condition
Informal description
Let $R$ and $S$ be binary relations on a type $\alpha$, and let $l$ be a list of elements of type $\alpha$. If for any two elements $a$ and $b$ in $l$, the relation $R(a, b)$ implies $S(a, b)$, and the list $l$ satisfies `Pairwise R l`, then $l$ also satisfies `Pairwise S l`.
List.Pairwise.and theorem
(hR : Pairwise R l) (hS : Pairwise S l) : l.Pairwise fun a b => R a b ∧ S a b
Full source
theorem Pairwise.and (hR : Pairwise R l) (hS : Pairwise S l) :
    l.Pairwise fun a b => R a b ∧ S a b := by
  induction hR with
  | nil => simp only [Pairwise.nil]
  | cons R1 _ IH =>
    simp only [Pairwise.nil, pairwise_cons] at hS ⊢
    exact ⟨fun b bl => ⟨R1 b bl, hS.1 b bl⟩, IH hS.2⟩
Conjunction of Pairwise Relations on a List
Informal description
For any list $l$ of elements of type $\alpha$, if $l$ satisfies the pairwise relation $R$ (i.e., `Pairwise R l`) and also satisfies the pairwise relation $S$ (i.e., `Pairwise S l`), then $l$ satisfies the pairwise relation $\lambda a b, R a b \land S a b$ (i.e., for every pair of elements in $l$ where one appears before the other, both $R$ and $S$ hold between them).
List.pairwise_and_iff theorem
: l.Pairwise (fun a b => R a b ∧ S a b) ↔ Pairwise R l ∧ Pairwise S l
Full source
theorem pairwise_and_iff : l.Pairwise (fun a b => R a b ∧ S a b) ↔ Pairwise R l ∧ Pairwise S l :=
  ⟨fun h => ⟨h.imp fun h => h.1, h.imp fun h => h.2⟩, fun ⟨hR, hS⟩ => hR.and hS⟩
Equivalence of Conjunctive Pairwise Relation to Separate Pairwise Relations
Informal description
For any list $l$ of elements of type $\alpha$ and any two binary relations $R$ and $S$ on $\alpha$, the list $l$ satisfies the pairwise relation $\lambda a b, R a b \land S a b$ if and only if $l$ satisfies both `Pairwise R l` and `Pairwise S l`.
List.Pairwise.imp₂ theorem
(H : ∀ a b, R a b → S a b → T a b) (hR : Pairwise R l) (hS : l.Pairwise S) : l.Pairwise T
Full source
theorem Pairwise.imp₂ (H : ∀ a b, R a b → S a b → T a b)
    (hR : Pairwise R l) (hS : l.Pairwise S) : l.Pairwise T :=
  (hR.and hS).imp fun ⟨h₁, h₂⟩ => H _ _ h₁ h₂
Pairwise Relation Implication for Two Relations on a List
Informal description
For any list $l$ of elements of type $\alpha$ and any relations $R$, $S$, and $T$ on $\alpha$, if for all $a, b \in \alpha$, $R(a, b)$ and $S(a, b)$ together imply $T(a, b)$, and if $l$ satisfies both `Pairwise R l` and `Pairwise S l`, then $l$ also satisfies `Pairwise T l`.
List.Pairwise.iff_of_mem theorem
{S : α → α → Prop} {l : List α} (H : ∀ {a b}, a ∈ l → b ∈ l → (R a b ↔ S a b)) : Pairwise R l ↔ Pairwise S l
Full source
theorem Pairwise.iff_of_mem {S : α → α → Prop} {l : List α}
    (H : ∀ {a b}, a ∈ lb ∈ l → (R a b ↔ S a b)) : PairwisePairwise R l ↔ Pairwise S l :=
  ⟨Pairwise.imp_of_mem fun m m' => (H m m').1, Pairwise.imp_of_mem fun m m' => (H m m').2⟩
Equivalence of Pairwise Relations under Membership Condition
Informal description
For any list $l$ of elements of type $\alpha$ and any two binary relations $R$ and $S$ on $\alpha$, if for all elements $a$ and $b$ in $l$ the relations $R(a, b)$ and $S(a, b)$ are equivalent, then $l$ satisfies `Pairwise R l` if and only if it satisfies `Pairwise S l`.
List.Pairwise.iff theorem
{S : α → α → Prop} (H : ∀ a b, R a b ↔ S a b) {l : List α} : Pairwise R l ↔ Pairwise S l
Full source
theorem Pairwise.iff {S : α → α → Prop} (H : ∀ a b, R a b ↔ S a b) {l : List α} :
    PairwisePairwise R l ↔ Pairwise S l :=
  Pairwise.iff_of_mem fun _ _ => H ..
Equivalence of Pairwise Relations Under Pointwise Equivalence
Informal description
For any binary relations $R$ and $S$ on a type $\alpha$ and any list $l$ of elements of type $\alpha$, if for all elements $a$ and $b$ in $\alpha$ the relations $R(a, b)$ and $S(a, b)$ are equivalent, then $l$ satisfies the pairwise relation `Pairwise R l` if and only if it satisfies `Pairwise S l`.
List.pairwise_of_forall theorem
{l : List α} (H : ∀ x y, R x y) : Pairwise R l
Full source
theorem pairwise_of_forall {l : List α} (H : ∀ x y, R x y) : Pairwise R l := by
  induction l <;> simp [*]
Universal Relation Implies Pairwise Relation on Lists
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $R$ on $\alpha$, if $R(x, y)$ holds for all elements $x$ and $y$ in $\alpha$, then the list $l$ satisfies the pairwise relation `Pairwise R l`. This means that for every pair of elements in $l$ where one appears before the other, the relation $R$ holds between them.
List.Pairwise.and_mem theorem
{l : List α} : Pairwise R l ↔ Pairwise (fun x y => x ∈ l ∧ y ∈ l ∧ R x y) l
Full source
theorem Pairwise.and_mem {l : List α} :
    PairwisePairwise R l ↔ Pairwise (fun x y => x ∈ l ∧ y ∈ l ∧ R x y) l :=
  Pairwise.iff_of_mem <| by simp +contextual
Equivalence of Pairwise Relation with Membership-Constrained Version
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $R$ on $\alpha$, the pairwise relation $\text{Pairwise } R\ l$ holds if and only if the pairwise relation holds for the modified relation $\lambda x y, x \in l \land y \in l \land R x y$ on $l$. In other words, a list $l$ is pairwise $R$ if and only if it is pairwise with respect to the relation that additionally requires both elements to be members of $l$ before applying $R$.
List.Pairwise.imp_mem theorem
{l : List α} : Pairwise R l ↔ Pairwise (fun x y => x ∈ l → y ∈ l → R x y) l
Full source
theorem Pairwise.imp_mem {l : List α} :
    PairwisePairwise R l ↔ Pairwise (fun x y => x ∈ l → y ∈ l → R x y) l :=
  Pairwise.iff_of_mem <| by simp +contextual
Equivalence of Pairwise Relation with Membership-Implication Form
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $R$ on $\alpha$, the list $l$ satisfies the pairwise relation $\text{Pairwise } R\ l$ if and only if it satisfies the pairwise relation where for any two elements $x$ and $y$ in $l$, the implication $x \in l \to y \in l \to R(x, y)$ holds.
List.Pairwise.forall_of_forall_of_flip theorem
(h₁ : ∀ x ∈ l, R x x) (h₂ : Pairwise R l) (h₃ : l.Pairwise (flip R)) : ∀ ⦃x⦄, x ∈ l → ∀ ⦃y⦄, y ∈ l → R x y
Full source
theorem Pairwise.forall_of_forall_of_flip (h₁ : ∀ x ∈ l, R x x) (h₂ : Pairwise R l)
    (h₃ : l.Pairwise (flip R)) : ∀ ⦃x⦄, x ∈ l → ∀ ⦃y⦄, y ∈ l → R x y := by
  induction l with
  | nil => exact forall_mem_nil _
  | cons a l ih =>
    rw [pairwise_cons] at h₂ h₃
    simp only [mem_cons]
    rintro x (rfl | hx) y (rfl | hy)
    · exact h₁ _ l.mem_cons_self
    · exact h₂.1 _ hy
    · exact h₃.1 _ hx
    · exact ih (fun x hx => h₁ _ <| mem_cons_of_mem _ hx) h₂.2 h₃.2 hx hy
Pairwise Relation Implies Total Relation on List Elements
Informal description
Let $R$ be a binary relation on a type $\alpha$ and $l$ be a list of elements of type $\alpha$. If: 1. $R$ is reflexive on $l$ (i.e., $\forall x \in l, R(x, x)$), 2. $l$ is pairwise $R$ (i.e., $\text{Pairwise } R\ l$), and 3. $l$ is pairwise under the flipped relation $\text{flip } R$ (i.e., $\text{Pairwise } (\text{flip } R)\ l$), then for any two elements $x, y \in l$, the relation $R(x, y)$ holds.
List.pairwise_singleton theorem
(R) (a : α) : Pairwise R [a]
Full source
theorem pairwise_singleton (R) (a : α) : Pairwise R [a] := by simp
Singleton List Satisfies Any Pairwise Relation
Informal description
For any binary relation $R$ on a type $\alpha$ and any element $a \in \alpha$, the singleton list $[a]$ satisfies the pairwise relation $\text{Pairwise } R$.
List.pairwise_pair theorem
{a b : α} : Pairwise R [a, b] ↔ R a b
Full source
theorem pairwise_pair {a b : α} : PairwisePairwise R [a, b] ↔ R a b := by simp
Pairwise Relation on Two-Element List: $R(a,b) \leftrightarrow \text{Pairwise } R [a, b]$
Informal description
For any two elements $a$ and $b$ of type $\alpha$ and any binary relation $R$ on $\alpha$, the list $[a, b]$ satisfies the pairwise relation `Pairwise R` if and only if $R(a, b)$ holds.
List.pairwise_map theorem
{l : List α} : (l.map f).Pairwise R ↔ l.Pairwise fun a b => R (f a) (f b)
Full source
theorem pairwise_map {l : List α} :
    (l.map f).Pairwise R ↔ l.Pairwise fun a b => R (f a) (f b) := by
  induction l
  · simp
  · simp only [map, pairwise_cons, forall_mem_map, *]
Pairwise Relation Preservation Under List Mapping: $R(f(a), f(b)) \leftrightarrow \text{Pairwise } R (l.map f)$
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f : \alpha \to \beta$, the mapped list $l.map(f)$ satisfies the pairwise relation $R$ if and only if the original list $l$ satisfies the pairwise relation $\lambda a b, R (f a) (f b)$.
List.Pairwise.of_map theorem
{S : β → β → Prop} (f : α → β) (H : ∀ a b : α, S (f a) (f b) → R a b) (p : Pairwise S (map f l)) : Pairwise R l
Full source
theorem Pairwise.of_map {S : β → β → Prop} (f : α → β) (H : ∀ a b : α, S (f a) (f b) → R a b)
    (p : Pairwise S (map f l)) : Pairwise R l :=
  (pairwise_map.1 p).imp (H _ _)
Lifting Pairwise Relations Through List Mapping
Informal description
Let $S$ be a binary relation on $\beta$ and $f : \alpha \to \beta$ a function. If for all $a, b \in \alpha$, $S(f(a), f(b))$ implies $R(a, b)$, and if the mapped list $\text{map } f l$ is pairwise related by $S$, then the original list $l$ is pairwise related by $R$.
List.Pairwise.map theorem
{S : β → β → Prop} (f : α → β) (H : ∀ a b : α, R a b → S (f a) (f b)) (p : Pairwise R l) : Pairwise S (map f l)
Full source
theorem Pairwise.map {S : β → β → Prop} (f : α → β) (H : ∀ a b : α, R a b → S (f a) (f b))
    (p : Pairwise R l) : Pairwise S (map f l) :=
  pairwise_map.2 <| p.imp (H _ _)
Pairwise Relation Preservation Under List Mapping: $R(a, b) \to S(f(a), f(b))$ Implies $\text{Pairwise } R l \to \text{Pairwise } S (\text{map } f l)$
Informal description
Let $R$ and $S$ be relations on types $\alpha$ and $\beta$ respectively, and let $f : \alpha \to \beta$ be a function such that for all $a, b \in \alpha$, $R(a, b)$ implies $S(f(a), f(b))$. If a list $l$ of elements of type $\alpha$ satisfies the pairwise relation $R$, then the mapped list $\text{map } f l$ satisfies the pairwise relation $S$.
List.pairwise_filterMap theorem
{f : β → Option α} {l : List β} : Pairwise R (filterMap f l) ↔ Pairwise (fun a a' : β => ∀ b ∈ f a, ∀ b' ∈ f a', R b b') l
Full source
theorem pairwise_filterMap {f : β → Option α} {l : List β} :
    PairwisePairwise R (filterMap f l) ↔ Pairwise (fun a a' : β => ∀ b ∈ f a, ∀ b' ∈ f a', R b b') l := by
  let _S (a a' : β) := ∀ b ∈ f a, ∀ b' ∈ f a', R b b'
  simp only [Option.mem_def]
  induction l with
  | nil => simp only [filterMap, Pairwise.nil]
  | cons a l IH => ?_
  match e : f a with
  | none =>
    rw [filterMap_cons_none e, pairwise_cons]
    simp only [e, false_implies, implies_true, true_and, IH, reduceCtorEq]
  | some b =>
    rw [filterMap_cons_some e]
    simpa [IH, e] using fun _ =>
      ⟨fun h a ha b hab => h _ _ ha hab, fun h a b ha hab => h _ ha _ hab⟩
Pairwise Relation Preservation under Filter-Map: $\text{Pairwise } R (\text{filterMap } f l) \leftrightarrow \text{Pairwise } (\lambda a a', \forall b \in f(a), \forall b' \in f(a'), R(b, b')) l$
Informal description
For a function $f : \beta \to \text{Option } \alpha$ and a list $l$ of elements in $\beta$, the filtered and mapped list $\text{filterMap } f l$ is pairwise related by $R$ if and only if the original list $l$ is pairwise related by the relation that for any two elements $a, a' \in \beta$, and any $b \in f(a)$, $b' \in f(a')$, the relation $R(b, b')$ holds.
List.Pairwise.filterMap theorem
{S : β → β → Prop} (f : α → Option β) (H : ∀ a a' : α, R a a' → ∀ b ∈ f a, ∀ b' ∈ f a', S b b') {l : List α} (p : Pairwise R l) : Pairwise S (filterMap f l)
Full source
theorem Pairwise.filterMap {S : β → β → Prop} (f : α → Option β)
    (H : ∀ a a' : α, R a a' → ∀ b ∈ f a, ∀ b' ∈ f a', S b b') {l : List α} (p : Pairwise R l) :
    Pairwise S (filterMap f l) :=
  pairwise_filterMap.2 <| p.imp (H _ _)
Preservation of Pairwise Relation under Filter-Map Operation
Informal description
Let $R$ be a relation on $\alpha$, $S$ a relation on $\beta$, and $f : \alpha \to \text{Option } \beta$ a function. Suppose that for any $a, a' \in \alpha$ with $R(a, a')$, and for any $b \in f(a)$ and $b' \in f(a')$, we have $S(b, b')$. Then for any list $l$ of elements of $\alpha$ that is pairwise related by $R$, the filtered and mapped list $\text{filterMap } f l$ is pairwise related by $S$.
List.pairwise_filter theorem
{p : α → Prop} [DecidablePred p] {l : List α} : Pairwise R (filter p l) ↔ Pairwise (fun x y => p x → p y → R x y) l
Full source
theorem pairwise_filter {p : α → Prop} [DecidablePred p] {l : List α} :
    PairwisePairwise R (filter p l) ↔ Pairwise (fun x y => p x → p y → R x y) l := by
  rw [← filterMap_eq_filter, pairwise_filterMap]
  simp
Pairwise Relation Preservation under List Filtering: $\text{Pairwise } R (\text{filter } p l) \leftrightarrow \text{Pairwise } (\lambda x y, p(x) \to p(y) \to R(x, y)) l$
Informal description
For any binary relation $R$ on a type $\alpha$, any decidable predicate $p : \alpha \to \text{Prop}$, and any list $l$ of elements in $\alpha$, the filtered list $\text{filter } p l$ is pairwise related by $R$ if and only if the original list $l$ is pairwise related by the relation that for any two elements $x$ and $y$ in $l$, if both $p(x)$ and $p(y)$ hold, then $R(x, y)$ holds.
List.Pairwise.filter theorem
(p : α → Bool) : Pairwise R l → Pairwise R (filter p l)
Full source
theorem Pairwise.filter (p : α → Bool) : Pairwise R l → Pairwise R (filter p l) :=
  Pairwise.sublist filter_sublist
Pairwise Relation is Preserved Under List Filtering
Informal description
For any binary relation $R$ on a type $\alpha$, any predicate $p : \alpha \to \text{Bool}$, and any list $l$ of elements in $\alpha$, if $R$ holds pairwise for all elements in $l$, then $R$ also holds pairwise for all elements in the filtered list $\text{filter } p l$.
List.pairwise_append theorem
{l₁ l₂ : List α} : (l₁ ++ l₂).Pairwise R ↔ l₁.Pairwise R ∧ l₂.Pairwise R ∧ ∀ a ∈ l₁, ∀ b ∈ l₂, R a b
Full source
theorem pairwise_append {l₁ l₂ : List α} :
    (l₁ ++ l₂).Pairwise R ↔ l₁.Pairwise R ∧ l₂.Pairwise R ∧ ∀ a ∈ l₁, ∀ b ∈ l₂, R a b := by
  induction l₁ <;> simp [*, or_imp, forall_and, and_assoc, and_left_comm]
Pairwise Relation on Concatenated Lists: $R$-pairwise$(l_1 ++ l_2) \leftrightarrow R$-pairwise$(l_1) \land R$-pairwise$(l_2) \land \forall a \in l_1, \forall b \in l_2, R(a, b)$
Informal description
For any binary relation $R$ on a type $\alpha$ and any two lists $l_1$ and $l_2$ of elements in $\alpha$, the concatenated list $l_1 ++ l_2$ satisfies the pairwise relation $R$ if and only if: 1. $l_1$ satisfies $R$ pairwise, 2. $l_2$ satisfies $R$ pairwise, and 3. For every element $a$ in $l_1$ and every element $b$ in $l_2$, the relation $R(a, b)$ holds.
List.pairwise_append_comm theorem
{R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {l₁ l₂ : List α} : Pairwise R (l₁ ++ l₂) ↔ Pairwise R (l₂ ++ l₁)
Full source
theorem pairwise_append_comm {R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {l₁ l₂ : List α} :
    PairwisePairwise R (l₁ ++ l₂) ↔ Pairwise R (l₂ ++ l₁) := by
  have (l₁ l₂ : List α) (H : ∀ x : α, x ∈ l₁ → ∀ y : α, y ∈ l₂ → R x y)
    (x : α) (xm : x ∈ l₂) (y : α) (ym : y ∈ l₁) : R x y := s (H y ym x xm)
  simp only [pairwise_append, and_left_comm]; rw [Iff.intro (this l₁ l₂) (this l₂ l₁)]
Pairwise Relation is Preserved Under List Concatenation Commutation for Symmetric Relations
Informal description
For any symmetric binary relation $R$ on a type $\alpha$ (i.e., $R x y$ implies $R y x$ for all $x, y \in \alpha$) and any two lists $l_1, l_2$ of elements in $\alpha$, the relation $R$ holds pairwise for the concatenated list $l_1 ++ l_2$ if and only if it holds pairwise for the concatenated list $l_2 ++ l_1$.
List.pairwise_middle theorem
{R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {a : α} {l₁ l₂ : List α} : Pairwise R (l₁ ++ a :: l₂) ↔ Pairwise R (a :: (l₁ ++ l₂))
Full source
theorem pairwise_middle {R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {a : α} {l₁ l₂ : List α} :
    PairwisePairwise R (l₁ ++ a :: l₂) ↔ Pairwise R (a :: (l₁ ++ l₂)) := by
  show PairwisePairwise R (l₁ ++ ([a] ++ l₂)) ↔ Pairwise R ([a] ++ l₁ ++ l₂)
  rw [← append_assoc, pairwise_append, @pairwise_append _ _ ([a] ++ l₁), pairwise_append_comm s]
  simp only [mem_append, or_comm]
Pairwise Relation for Middle-Inserted Element in Concatenated Lists
Informal description
For any symmetric binary relation $R$ on a type $\alpha$ (i.e., $R(x, y)$ implies $R(y, x)$ for all $x, y \in \alpha$), any element $a \in \alpha$, and any two lists $l_1, l_2$ of elements in $\alpha$, the following are equivalent: 1. The list $l_1 \mathbin{+\kern-1.5ex+} [a] \mathbin{+\kern-1.5ex+} l_2$ satisfies $R$ pairwise. 2. The list $[a] \mathbin{+\kern-1.5ex+} (l_1 \mathbin{+\kern-1.5ex+} l_2)$ satisfies $R$ pairwise.
List.pairwise_flatten theorem
{L : List (List α)} : Pairwise R (flatten L) ↔ (∀ l ∈ L, Pairwise R l) ∧ Pairwise (fun l₁ l₂ => ∀ x ∈ l₁, ∀ y ∈ l₂, R x y) L
Full source
theorem pairwise_flatten {L : List (List α)} :
    PairwisePairwise R (flatten L) ↔
      (∀ l ∈ L, Pairwise R l) ∧ Pairwise (fun l₁ l₂ => ∀ x ∈ l₁, ∀ y ∈ l₂, R x y) L := by
  induction L with
  | nil => simp
  | cons l L IH =>
    simp only [flatten, pairwise_append, IH, mem_flatten, exists_imp, and_imp, forall_mem_cons,
      pairwise_cons, and_assoc, and_congr_right_iff]
    rw [and_comm, and_congr_left_iff]
    intros; exact ⟨fun h l' b c d e => h c d e l' b, fun h c d e l' b => h l' b c d e⟩
Pairwise Relation on Flattened List of Lists: $\text{Pairwise}(R, \text{flatten}(L)) \leftrightarrow (\forall l \in L, \text{Pairwise}(R, l)) \land \text{Pairwise}(\forall x \in l_1, \forall y \in l_2, R(x, y), L)$
Informal description
For a list of lists $L$ of elements of type $\alpha$ and a binary relation $R$ on $\alpha$, the concatenated list $\text{flatten}(L)$ satisfies the pairwise relation $R$ if and only if: 1. Every sublist $l \in L$ satisfies the pairwise relation $R$, and 2. The list $L$ itself satisfies the pairwise relation where for any two sublists $l_1, l_2 \in L$ with $l_1$ appearing before $l_2$, the relation $R(x, y)$ holds for all $x \in l_1$ and $y \in l_2$.
List.pairwise_join abbrev
Full source
@[deprecated pairwise_flatten (since := "2024-10-14")] abbrev pairwise_join := @pairwise_flatten
Pairwise Relation on Joined List of Lists: $\text{Pairwise}(R, \text{join}(L)) \leftrightarrow (\forall l \in L, \text{Pairwise}(R, l)) \land \text{Pairwise}(\forall x \in l_1, \forall y \in l_2, R(x, y), L)$
Informal description
For a list of lists $L$ of elements of type $\alpha$ and a binary relation $R$ on $\alpha$, the concatenated list obtained by joining all sublists in $L$ (denoted as $\text{join}(L)$) satisfies the pairwise relation $R$ if and only if: 1. Every sublist $l \in L$ satisfies the pairwise relation $R$, and 2. The list $L$ itself satisfies the pairwise relation where for any two sublists $l_1, l_2 \in L$ with $l_1$ appearing before $l_2$, the relation $R(x, y)$ holds for all $x \in l_1$ and $y \in l_2$.
List.pairwise_flatMap theorem
{R : β → β → Prop} {l : List α} {f : α → List β} : List.Pairwise R (l.flatMap f) ↔ (∀ a ∈ l, Pairwise R (f a)) ∧ Pairwise (fun a₁ a₂ => ∀ x ∈ f a₁, ∀ y ∈ f a₂, R x y) l
Full source
theorem pairwise_flatMap {R : β → β → Prop} {l : List α} {f : α → List β} :
    List.PairwiseList.Pairwise R (l.flatMap f) ↔
      (∀ a ∈ l, Pairwise R (f a)) ∧ Pairwise (fun a₁ a₂ => ∀ x ∈ f a₁, ∀ y ∈ f a₂, R x y) l := by
  simp [List.flatMap, pairwise_flatten, pairwise_map]
Pairwise Relation on Flat-Mapped List: $\text{Pairwise}(R, \text{flatMap}\ f\ l) \leftrightarrow (\forall a \in l, \text{Pairwise}(R, f\ a)) \land \text{Pairwise}(\forall x \in f\ a_1, \forall y \in f\ a_2, R(x, y), l)$
Informal description
For a binary relation $R$ on elements of type $\beta$, a list $l$ of elements of type $\alpha$, and a function $f : \alpha \to \text{List}\ \beta$, the following equivalence holds: The flattened list $\text{flatMap}\ f\ l$ satisfies $\text{Pairwise}\ R$ if and only if: 1. For every element $a \in l$, the list $f(a)$ satisfies $\text{Pairwise}\ R$, and 2. The list $l$ satisfies $\text{Pairwise}$ with respect to the relation: for any $a_1, a_2 \in l$ where $a_1$ precedes $a_2$, the relation $R(x, y)$ holds for all $x \in f(a_1)$ and $y \in f(a_2)$.
List.pairwise_bind abbrev
Full source
@[deprecated pairwise_flatMap (since := "2024-10-14")] abbrev pairwise_bind := @pairwise_flatMap
Pairwise Relation on Bound List: $\text{Pairwise}(R, \text{bind}\ f\ l) \leftrightarrow (\forall a \in l, \text{Pairwise}(R, f\ a)) \land \text{Pairwise}(\forall x \in f\ a_1, \forall y \in f\ a_2, R(x, y), l)$
Informal description
For a binary relation $R$ on elements of type $\beta$, a list $l$ of elements of type $\alpha$, and a function $f : \alpha \to \text{List}\ \beta$, the following equivalence holds: The flattened list obtained by mapping $f$ over $l$ and then concatenating (denoted as $\text{bind}\ f\ l$) satisfies $\text{Pairwise}\ R$ if and only if: 1. For every element $a \in l$, the list $f(a)$ satisfies $\text{Pairwise}\ R$, and 2. The list $l$ satisfies $\text{Pairwise}$ with respect to the relation: for any $a_1, a_2 \in l$ where $a_1$ precedes $a_2$, the relation $R(x, y)$ holds for all $x \in f(a_1)$ and $y \in f(a_2)$.
List.pairwise_reverse theorem
{l : List α} : l.reverse.Pairwise R ↔ l.Pairwise (fun a b => R b a)
Full source
theorem pairwise_reverse {l : List α} :
    l.reverse.Pairwise R ↔ l.Pairwise (fun a b => R b a) := by
  induction l <;> simp [*, pairwise_append, and_comm]
Reversed List Pairwise Relation: $\text{Pairwise}(R, \text{reverse}(l)) \leftrightarrow \text{Pairwise}(\lambda a b, R(b, a), l)$
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $R$ on $\alpha$, the reversed list $\text{reverse}(l)$ satisfies the pairwise relation $R$ if and only if the original list $l$ satisfies the pairwise relation $\lambda a b, R(b, a)$. In other words, $\text{Pairwise}(R, \text{reverse}(l)) \leftrightarrow \text{Pairwise}(\lambda a b, R(b, a), l)$.
List.pairwise_replicate theorem
{n : Nat} {a : α} : (replicate n a).Pairwise R ↔ n ≤ 1 ∨ R a a
Full source
@[simp] theorem pairwise_replicate {n : Nat} {a : α} :
    (replicate n a).Pairwise R ↔ n ≤ 1 ∨ R a a := by
  induction n with
  | zero => simp
  | succ n ih =>
    simp only [replicate_succ, pairwise_cons, mem_replicate, ne_eq, and_imp,
      forall_eq_apply_imp_iff, ih]
    constructor
    · rintro ⟨h, h' | h'⟩
      · by_cases w : n = 0
        · left
          subst w
          simp
        · right
          exact h w
      · right
        exact h'
    · rintro (h | h)
      · obtain rfl := eq_zero_of_le_zero (le_of_lt_succ h)
        simp
      · exact ⟨fun _ => h, Or.inr h⟩
Pairwise Relation on Replicated List: $\text{replicate}\ n\ a\ \text{Pairwise}\ R \leftrightarrow n \leq 1 \lor R\ a\ a$
Informal description
For any natural number $n$ and any element $a$ of type $\alpha$, the list consisting of $n$ copies of $a$ satisfies the pairwise relation $R$ if and only if either $n \leq 1$ or $R(a, a)$ holds.
List.Pairwise.drop theorem
{l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.drop i)
Full source
theorem Pairwise.drop {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.drop i) :=
  h.sublist (drop_sublist _ _)
Pairwise Relation Preserved Under Dropping Elements
Informal description
For any list $l$ of elements of type $\alpha$ and any natural number $i$, if the relation $R$ holds pairwise for all elements in $l$, then $R$ also holds pairwise for all elements in the list obtained by dropping the first $i$ elements of $l$.
List.Pairwise.take theorem
{l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.take i)
Full source
theorem Pairwise.take {l : List α} {i : Nat} (h : List.Pairwise R l) : List.Pairwise R (l.take i) :=
  h.sublist (take_sublist _ _)
Pairwise Relation Preserved Under Taking Initial Segment
Informal description
For any list $l$ of elements of type $\alpha$ and any natural number $i$, if the relation $R$ holds pairwise for all elements in $l$, then $R$ also holds pairwise for the first $i$ elements of $l$ (i.e., for the sublist obtained by `take i l`).
List.pairwise_iff_forall_sublist theorem
: l.Pairwise R ↔ (∀ {a b}, [a, b] <+ l → R a b)
Full source
theorem pairwise_iff_forall_sublist : l.Pairwise R ↔ (∀ {a b}, [a,b] <+ l → R a b) := by
  induction l with
  | nil => simp
  | cons hd tl IH =>
    rw [List.pairwise_cons]
    constructor <;> intro h
    · intro
      | a, b, .cons _ hab => exact IH.mp h.2 hab
      | _, b, .cons₂ _ hab => refine h.1 _ (hab.subset ?_); simp
    · constructor
      · intro x hx
        apply h
        rw [List.cons_sublist_cons, List.singleton_sublist]
        exact hx
      · apply IH.mpr
        intro a b hab
        apply h; exact hab.cons _
Pairwise Relation Characterization via Sublists: $R$-Pairwise iff $R$ Holds on All Consecutive Pairs
Informal description
For any list $l$ and binary relation $R$, the list $l$ satisfies the pairwise relation $R$ if and only if for every pair of elements $a$ and $b$ in $l$, if the sublist $[a, b]$ appears in $l$, then $R(a, b)$ holds.
List.Pairwise.rel_of_mem_take_of_mem_drop theorem
{l : List α} (h : l.Pairwise R) (hx : x ∈ l.take i) (hy : y ∈ l.drop i) : R x y
Full source
theorem Pairwise.rel_of_mem_take_of_mem_drop
    {l : List α} (h : l.Pairwise R) (hx : x ∈ l.take i) (hy : y ∈ l.drop i) : R x y := by
  apply pairwise_iff_forall_sublist.mp h
  rw [← take_append_drop i l, sublist_append_iff]
  refine ⟨[x], [y], rfl, by simpa, by simpa⟩
Pairwise Relation Between Elements in Take and Drop Segments: $R(x, y)$ for $x \in \text{take}\ i\ l$ and $y \in \text{drop}\ i\ l$
Informal description
For any list $l$ of elements of type $\alpha$, if $l$ satisfies the pairwise relation $R$, then for any element $x$ in the first $i$ elements of $l$ (i.e., $x \in \text{take}\ i\ l$) and any element $y$ in the remaining elements after dropping the first $i$ elements (i.e., $y \in \text{drop}\ i\ l$), the relation $R(x, y)$ holds.
List.Pairwise.rel_of_mem_append theorem
{l₁ l₂ : List α} (h : (l₁ ++ l₂).Pairwise R) (hx : x ∈ l₁) (hy : y ∈ l₂) : R x y
Full source
theorem Pairwise.rel_of_mem_append
    {l₁ l₂ : List α} (h : (l₁ ++ l₂).Pairwise R) (hx : x ∈ l₁) (hy : y ∈ l₂) : R x y := by
  apply pairwise_iff_forall_sublist.mp h
  rw [sublist_append_iff]
  exact ⟨[x], [y], rfl, by simpa, by simpa⟩
Pairwise Relation Preserved Under List Concatenation
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if the concatenated list $l_1 \mathbin{+\!\!+} l_2$ satisfies the pairwise relation $R$, then for any element $x \in l_1$ and $y \in l_2$, the relation $R(x, y)$ holds.
List.pairwise_of_forall_mem_list theorem
{l : List α} {r : α → α → Prop} (h : ∀ a ∈ l, ∀ b ∈ l, r a b) : l.Pairwise r
Full source
theorem pairwise_of_forall_mem_list {l : List α} {r : α → α → Prop} (h : ∀ a ∈ l, ∀ b ∈ l, r a b) :
    l.Pairwise r := by
  rw [pairwise_iff_forall_sublist]
  intro a b hab
  apply h <;> (apply hab.subset; simp)
Universal Pairwise Relation Implies Pairwise on List
Informal description
For any list $l$ of elements of type $\alpha$ and any binary relation $r$ on $\alpha$, if for all elements $a$ and $b$ in $l$ the relation $r(a, b)$ holds, then the list $l$ is pairwise related by $r$.
List.pairwise_pmap theorem
{p : β → Prop} {f : ∀ b, p b → α} {l : List β} (h : ∀ x ∈ l, p x) : Pairwise R (l.pmap f h) ↔ Pairwise (fun b₁ b₂ => ∀ (h₁ : p b₁) (h₂ : p b₂), R (f b₁ h₁) (f b₂ h₂)) l
Full source
theorem pairwise_pmap {p : β → Prop} {f : ∀ b, p b → α} {l : List β} (h : ∀ x ∈ l, p x) :
    PairwisePairwise R (l.pmap f h) ↔
      Pairwise (fun b₁ b₂ => ∀ (h₁ : p b₁) (h₂ : p b₂), R (f b₁ h₁) (f b₂ h₂)) l := by
  induction l with
  | nil => simp
  | cons a l ihl =>
    obtain ⟨_, hl⟩ : p a ∧ ∀ b, b ∈ l → p b := by simpa using h
    simp only [ihl hl, pairwise_cons, exists₂_imp, pmap, and_congr_left_iff, mem_pmap]
    refine fun _ => ⟨fun H b hb _ hpb => H _ _ hb rfl, ?_⟩
    rintro H _ b hb rfl
    exact H b hb _ _
Pairwise Relation Preservation under Partial Mapping: $R(f(b_1), f(b_2)) \leftrightarrow \forall h_1\ h_2, R(f(b_1, h_1), f(b_2, h_2))$
Informal description
Let $p$ be a predicate on elements of type $\beta$, $f$ a function that maps each $b \in \beta$ with $p(b)$ to an element of type $\alpha$, and $l$ a list of elements of type $\beta$ where every element satisfies $p$. Then, the list obtained by mapping $f$ over $l$ (denoted as $l.\text{pmap}\ f\ h$) is pairwise related by $R$ if and only if the original list $l$ is pairwise related by the relation $\lambda b_1\ b_2, \forall (h_1 : p(b_1)) (h_2 : p(b_2)), R(f(b_1, h_1), f(b_2, h_2))$.
List.Pairwise.pmap theorem
{l : List α} (hl : Pairwise R l) {p : α → Prop} {f : ∀ a, p a → β} (h : ∀ x ∈ l, p x) {S : β → β → Prop} (hS : ∀ ⦃x⦄ (hx : p x) ⦃y⦄ (hy : p y), R x y → S (f x hx) (f y hy)) : Pairwise S (l.pmap f h)
Full source
theorem Pairwise.pmap {l : List α} (hl : Pairwise R l) {p : α → Prop} {f : ∀ a, p a → β}
    (h : ∀ x ∈ l, p x) {S : β → β → Prop}
    (hS : ∀ ⦃x⦄ (hx : p x) ⦃y⦄ (hy : p y), R x y → S (f x hx) (f y hy)) :
    Pairwise S (l.pmap f h) := by
  refine (pairwise_pmap h).2 (Pairwise.imp_of_mem ?_ hl)
  intros; apply hS; assumption
Preservation of Pairwise Relation under Partial Mapping
Informal description
Let $R$ be a binary relation on a type $\alpha$, and let $l$ be a list of elements of type $\alpha$ such that `Pairwise R l` holds. Let $p$ be a predicate on $\alpha$, $f$ a function mapping each $a \in \alpha$ with $p(a)$ to an element of type $\beta$, and assume every element in $l$ satisfies $p$. If for any $x, y \in l$ with $R(x, y)$, and any proofs $h_x : p(x)$, $h_y : p(y)$, the relation $S(f(x, h_x), f(y, h_y))$ holds, then the list obtained by mapping $f$ over $l$ (denoted as $l.\text{pmap}\ f\ h$) satisfies `Pairwise S`.
List.nodup_nil theorem
: @Nodup α []
Full source
@[simp]
theorem nodup_nil : @Nodup α [] :=
  Pairwise.nil
Empty List Has No Duplicates
Informal description
The empty list `[]` has no duplicate elements, i.e., `Nodup []` holds.
List.nodup_cons theorem
{a : α} {l : List α} : Nodup (a :: l) ↔ a ∉ l ∧ Nodup l
Full source
@[simp]
theorem nodup_cons {a : α} {l : List α} : NodupNodup (a :: l) ↔ a ∉ l ∧ Nodup l := by
  simp only [Nodup, pairwise_cons, forall_mem_ne]
Cons-Preservation of No-Duplicates Property in Lists
Informal description
For any element $a$ of type $\alpha$ and any list $l$ of elements of type $\alpha$, the list $a :: l$ has no duplicates if and only if $a$ does not appear in $l$ and $l$ itself has no duplicates.
List.Nodup.sublist theorem
: l₁ <+ l₂ → Nodup l₂ → Nodup l₁
Full source
theorem Nodup.sublist : l₁ <+ l₂Nodup l₂ → Nodup l₁ :=
  Pairwise.sublist
Sublist Preserves No-Duplicates Property
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 \subseteq l_2$) and $l_2$ has no duplicate elements, then $l_1$ also has no duplicate elements.
List.Sublist.nodup theorem
: l₁ <+ l₂ → Nodup l₂ → Nodup l₁
Full source
theorem Sublist.nodup : l₁ <+ l₂Nodup l₂ → Nodup l₁ :=
  Nodup.sublist
Sublist Preserves No-Duplicates Property
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$, if $l_1$ is a sublist of $l_2$ (denoted $l_1 \subseteq l_2$) and $l_2$ has no duplicate elements, then $l_1$ also has no duplicate elements.
List.getElem?_inj theorem
{xs : List α} (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs[i]? = xs[j]?) : i = j
Full source
theorem getElem?_inj {xs : List α}
    (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs[i]? = xs[j]?) : i = j := by
  induction xs generalizing i j with
  | nil => cases h₀
  | cons x xs ih =>
    match i, j with
    | 0, 0 => rfl
    | i+1, j+1 =>
      cases h₁ with
      | cons ha h₁ =>
        simp only [getElem?_cons_succ] at h₂
        exact congrArg (· + 1) (ih (Nat.lt_of_succ_lt_succ h₀) h₁ h₂)
    | i+1, 0 => ?_
    | 0, j+1 => ?_
    all_goals
      simp only [getElem?_cons_zero, getElem?_cons_succ] at h₂
      cases h₁; rename_i h' h
      have := h x ?_ rfl; cases this
      rw [mem_iff_getElem?]
    exact ⟨_, h₂⟩; exact ⟨_ , h₂.symm⟩
Injective Indexing in No-Duplicate Lists: `xs[i]? = xs[j]? → i = j` when `xs` has no duplicates and `i` is valid.
Informal description
For any list `xs` of elements of type `α`, if the index `i` is within the bounds of `xs` (i.e., `i < xs.length`), `xs` has no duplicate elements, and the optional lookup operations `xs[i]?` and `xs[j]?` return the same value, then the indices `i` and `j` must be equal.
List.nodup_replicate theorem
{n : Nat} {a : α} : (replicate n a).Nodup ↔ n ≤ 1
Full source
@[simp] theorem nodup_replicate {n : Nat} {a : α} :
    (replicate n a).Nodup ↔ n ≤ 1 := by simp [Nodup]
Replicated List Has No Duplicates if and only if Replication Count is at Most One
Informal description
For any natural number $n$ and any element $a$ of type $\alpha$, the list obtained by replicating $a$ exactly $n$ times has no duplicate elements if and only if $n \leq 1$.