List Permutations #
This file develops theory about the List.Perm
relation.
Notation #
The notation ~
is used for permutation equivalence.
theorem
List.forall₂_comp_perm_eq_perm_comp_forall₂
{α : Type u_1}
{β : Type u_2}
{r : α → β → Prop}
:
theorem
List.Perm.foldl_eq
{α : Type u_1}
{β : Type u_2}
{f : β → α → β}
{l₁ l₂ : List α}
[rcomm : RightCommutative f]
(p : l₁.Perm l₂)
(b : β)
:
theorem
List.Perm.foldr_eq
{α : Type u_1}
{β : Type u_2}
{f : α → β → β}
{l₁ l₂ : List α}
[lcomm : LeftCommutative f]
(p : l₁.Perm l₂)
(b : β)
:
theorem
List.Perm.foldl_op_eq
{α : Type u_1}
{op : α → α → α}
[IA : Std.Associative op]
[IC : Std.Commutative op]
{l₁ l₂ : List α}
{a : α}
(h : l₁.Perm l₂)
:
theorem
List.Perm.foldr_op_eq
{α : Type u_1}
{op : α → α → α}
[IA : Std.Associative op]
[IC : Std.Commutative op]
{l₁ l₂ : List α}
{a : α}
(h : l₁.Perm l₂)
:
@[deprecated List.Perm.foldl_op_eq (since := "2024-09-28")]
theorem
List.Perm.fold_op_eq
{α : Type u_1}
{op : α → α → α}
[IA : Std.Associative op]
[IC : Std.Commutative op]
{l₁ l₂ : List α}
{a : α}
(h : l₁.Perm l₂)
:
Alias of List.Perm.foldl_op_eq
.
@[deprecated List.perm_option_toList (since := "2024-10-16")]
Alias of List.perm_option_toList
.