Double universal quantification on a list #
This file provides an API for List.Forall₂
(definition in Data.List.Defs
).
Forall₂ R l₁ l₂
means that l₁
and l₂
have the same length, and whenever a
is the nth element
of l₁
, and b
is the nth element of l₂
, then R a b
is satisfied.
theorem
List.Forall₂.flip
{α : Type u_1}
{β : Type u_2}
{R : α → β → Prop}
{a : List α}
{b : List β}
:
Forall₂ (_root_.flip R) b a → Forall₂ R a b
theorem
Relator.LeftUnique.forall₂
{α : Type u_1}
{β : Type u_2}
{R : α → β → Prop}
(hr : LeftUnique R)
:
theorem
Relator.RightUnique.forall₂
{α : Type u_1}
{β : Type u_2}
{R : α → β → Prop}
(hr : RightUnique R)
:
theorem
Relator.BiUnique.forall₂
{α : Type u_1}
{β : Type u_2}
{R : α → β → Prop}
(hr : BiUnique R)
:
BiUnique (List.Forall₂ R)
Given a relation R
, sublist_forall₂ r l₁ l₂
indicates that there is a sublist of l₂
such
that forall₂ r l₁ l₂
.
- nil {α : Type u_1} {β : Type u_2} {R : α → β → Prop} {l : List β} : SublistForall₂ R [] l
- cons {α : Type u_1} {β : Type u_2} {R : α → β → Prop} {a₁ : α} {a₂ : β} {l₁ : List α} {l₂ : List β} : R a₁ a₂ → SublistForall₂ R l₁ l₂ → SublistForall₂ R (a₁ :: l₁) (a₂ :: l₂)
- cons_right {α : Type u_1} {β : Type u_2} {R : α → β → Prop} {a : β} {l₁ : List α} {l₂ : List β} : SublistForall₂ R l₁ l₂ → SublistForall₂ R l₁ (a :: l₂)
Instances For
instance
List.SublistForall₂.is_refl
{α : Type u_1}
{Rₐ : α → α → Prop}
[IsRefl α Rₐ]
:
IsRefl (List α) (SublistForall₂ Rₐ)
instance
List.SublistForall₂.is_trans
{α : Type u_1}
{Rₐ : α → α → Prop}
[IsTrans α Rₐ]
:
IsTrans (List α) (SublistForall₂ Rₐ)
theorem
List.Sublist.sublistForall₂
{α : Type u_1}
{Rₐ : α → α → Prop}
{l₁ l₂ : List α}
(h : l₁.Sublist l₂)
[IsRefl α Rₐ]
:
SublistForall₂ Rₐ l₁ l₂
theorem
List.tail_sublistForall₂_self
{α : Type u_1}
{Rₐ : α → α → Prop}
[IsRefl α Rₐ]
(l : List α)
:
SublistForall₂ Rₐ l.tail l
@[simp]
theorem
List.sublistForall₂_map_left_iff
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{R : α → β → Prop}
{f : γ → α}
{l₁ : List γ}
{l₂ : List β}
:
SublistForall₂ R (map f l₁) l₂ ↔ SublistForall₂ (fun (c : γ) (b : β) => R (f c) b) l₁ l₂
@[simp]
theorem
List.sublistForall₂_map_right_iff
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
{R : α → β → Prop}
{f : γ → β}
{l₁ : List α}
{l₂ : List γ}
:
SublistForall₂ R l₁ (map f l₂) ↔ SublistForall₂ (fun (a : α) (c : γ) => R a (f c)) l₁ l₂