mathlib3 documentation

data.fin.tuple.basic

Operation on tuples #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

We interpret maps Π i : fin n, α i as n-tuples of elements of possibly varying type α i, (α 0, …, α (n-1)). A particular case is fin n → α of elements with all the same type. In this case when α i is a constant map, then tuples are isomorphic (but not definitionally equal) to vectors.

We define the following operations:

@[simp]
theorem fin.tuple0_le {α : fin 0 Type u_1} [Π (i : fin 0), preorder (α i)] (f g : Π (i : fin 0), α i) :
f g
def fin.tail {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (i : fin n) :
α i.succ

The tail of an n+1 tuple, i.e., its last n entries.

Equations
theorem fin.tail_def {n : } {α : fin (n + 1) Type u_1} {q : Π (i : fin (n + 1)), α i} :
fin.tail (λ (k : fin (n + 1)), q k) = λ (k : fin n), q k.succ
def fin.cons {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) (i : fin (n + 1)) :
α i

Adding an element at the beginning of an n-tuple, to get an n+1-tuple.

Equations
@[simp]
theorem fin.tail_cons {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) :
@[simp]
theorem fin.cons_succ {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) (i : fin n) :
fin.cons x p i.succ = p i
@[simp]
theorem fin.cons_zero {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) :
fin.cons x p 0 = x
@[simp]
theorem fin.cons_update {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) (i : fin n) (y : α i.succ) :

Updating a tuple and adding an element at the beginning commute.

theorem fin.cons_injective2 {n : } {α : fin (n + 1) Type u} :

As a binary function, fin.cons is injective.

@[simp]
theorem fin.cons_eq_cons {n : } {α : fin (n + 1) Type u} {x₀ y₀ : α 0} {x y : Π (i : fin n), α i.succ} :
fin.cons x₀ x = fin.cons y₀ y x₀ = y₀ x = y
theorem fin.cons_left_injective {n : } {α : fin (n + 1) Type u} (x : Π (i : fin n), α i.succ) :
function.injective (λ (x₀ : α 0), fin.cons x₀ x)
theorem fin.cons_right_injective {n : } {α : fin (n + 1) Type u} (x₀ : α 0) :
theorem fin.update_cons_zero {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (i : fin n), α i.succ) (z : α 0) :

Adding an element at the beginning of a tuple and then updating it amounts to adding it directly.

@[simp]
theorem fin.cons_self_tail {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) :
fin.cons (q 0) (fin.tail q) = q

Concatenating the first element of a tuple with its tail gives back the original tuple

def fin.cons_cases {n : } {α : fin (n + 1) Type u} {P : (Π (i : fin n.succ), α i) Sort v} (h : Π (x₀ : α 0) (x : Π (i : fin n), α i.succ), P (fin.cons x₀ x)) (x : Π (i : fin n.succ), α i) :
P x

Recurse on an n+1-tuple by splitting it into a single element and an n-tuple.

Equations
@[simp]
theorem fin.cons_cases_cons {n : } {α : fin (n + 1) Type u} {P : (Π (i : fin n.succ), α i) Sort v} (h : Π (x₀ : α 0) (x : Π (i : fin n), α i.succ), P (fin.cons x₀ x)) (x₀ : α 0) (x : Π (i : fin n), α i.succ) :
fin.cons_cases h (fin.cons x₀ x) = h x₀ x
def fin.cons_induction {α : Type u_1} {P : Π {n : }, (fin n α) Sort v} (h0 : P fin.elim0) (h : Π {n : } (x₀ : α) (x : fin n α), P x P (fin.cons x₀ x)) {n : } (x : fin n α) :
P x

Recurse on an tuple by splitting into fin.elim0 and fin.cons.

Equations
theorem fin.cons_injective_of_injective {n : } {α : Type u_1} {x₀ : α} {x : fin n α} (hx₀ : x₀ set.range x) (hx : function.injective x) :
theorem fin.cons_injective_iff {n : } {α : Type u_1} {x₀ : α} {x : fin n α} :
@[simp]
theorem fin.forall_fin_zero_pi {α : fin 0 Sort u_1} {P : (Π (i : fin 0), α i) Prop} :
( (x : Π (i : fin 0), α i), P x) P fin_zero_elim
@[simp]
theorem fin.exists_fin_zero_pi {α : fin 0 Sort u_1} {P : (Π (i : fin 0), α i) Prop} :
( (x : Π (i : fin 0), α i), P x) P fin_zero_elim
theorem fin.forall_fin_succ_pi {n : } {α : fin (n + 1) Type u} {P : (Π (i : fin (n + 1)), α i) Prop} :
( (x : Π (i : fin (n + 1)), α i), P x) (a : α 0) (v : Π (i : fin n), α i.succ), P (fin.cons a v)
theorem fin.exists_fin_succ_pi {n : } {α : fin (n + 1) Type u} {P : (Π (i : fin (n + 1)), α i) Prop} :
( (x : Π (i : fin (n + 1)), α i), P x) (a : α 0) (v : Π (i : fin n), α i.succ), P (fin.cons a v)
@[simp]
theorem fin.tail_update_zero {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (z : α 0) :

Updating the first element of a tuple does not change the tail.

@[simp]
theorem fin.tail_update_succ {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (i : fin n) (y : α i.succ) :

Updating a nonzero element and taking the tail commute.

theorem fin.comp_cons {n : } {α : Type u_1} {β : Type u_2} (g : α β) (y : α) (q : fin n α) :
g fin.cons y q = fin.cons (g y) (g q)
theorem fin.comp_tail {n : } {α : Type u_1} {β : Type u_2} (g : α β) (q : fin n.succ α) :
theorem fin.le_cons {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {x : α 0} {q : Π (i : fin (n + 1)), α i} {p : Π (i : fin n), α i.succ} :
q fin.cons x p q 0 x fin.tail q p
theorem fin.cons_le {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {x : α 0} {q : Π (i : fin (n + 1)), α i} {p : Π (i : fin n), α i.succ} :
fin.cons x p q x q 0 p fin.tail q
theorem fin.cons_le_cons {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {x₀ y₀ : α 0} {x y : Π (i : fin n), α i.succ} :
fin.cons x₀ x fin.cons y₀ y x₀ y₀ x y
theorem fin.pi_lex_lt_cons_cons {n : } {α : fin (n + 1) Type u} {x₀ y₀ : α 0} {x y : Π (i : fin n), α i.succ} (s : Π {i : fin n.succ}, α i α i Prop) :
pi.lex has_lt.lt s (fin.cons x₀ x) (fin.cons y₀ y) s x₀ y₀ x₀ = y₀ pi.lex has_lt.lt (λ (i : fin n), s) x y
theorem fin.range_fin_succ {n : } {α : Type u_1} (f : fin (n + 1) α) :
@[simp]
theorem fin.range_cons {α : Type u_1} {n : } (x : α) (b : fin n α) :
def fin.append {m n : } {α : Type u_1} (a : fin m α) (b : fin n α) :
fin (m + n) α

Append a tuple of length m to a tuple of length n to get a tuple of length m + n. This is a non-dependent version of fin.add_cases.

Equations
@[simp]
theorem fin.append_left {m n : } {α : Type u_1} (u : fin m α) (v : fin n α) (i : fin m) :
fin.append u v ((fin.cast_add n) i) = u i
@[simp]
theorem fin.append_right {m n : } {α : Type u_1} (u : fin m α) (v : fin n α) (i : fin n) :
fin.append u v ((fin.nat_add m) i) = v i
theorem fin.append_right_nil {m n : } {α : Type u_1} (u : fin m α) (v : fin n α) (hv : n = 0) :
@[simp]
theorem fin.append_elim0' {m : } {α : Type u_1} (u : fin m α) :
theorem fin.append_left_nil {m n : } {α : Type u_1} (u : fin m α) (v : fin n α) (hu : m = 0) :
@[simp]
theorem fin.elim0'_append {n : } {α : Type u_1} (v : fin n α) :
theorem fin.append_assoc {m n p : } {α : Type u_1} (a : fin m α) (b : fin n α) (c : fin p α) :
theorem fin.append_left_eq_cons {α : Type u_1} {n : } (x₀ : fin 1 α) (x : fin n α) :
fin.append x₀ x = fin.cons (x₀ 0) x (fin.cast _)

Appending a one-tuple to the left is the same as fin.cons.

@[simp]
def fin.repeat {n : } {α : Type u_1} (m : ) (a : fin n α) :
fin (m * n) α

Repeat a m times. For example fin.repeat 2 ![0, 3, 7] = ![0, 3, 7, 0, 3, 7].

Equations
@[simp]
theorem fin.repeat_zero {n : } {α : Type u_1} (a : fin n α) :
@[simp]
theorem fin.repeat_one {n : } {α : Type u_1} (a : fin n α) :
theorem fin.repeat_succ {n : } {α : Type u_1} (a : fin n α) (m : ) :
@[simp]
theorem fin.repeat_add {n : } {α : Type u_1} (a : fin n α) (m₁ m₂ : ) :
fin.repeat (m₁ + m₂) a = fin.append (fin.repeat m₁ a) (fin.repeat m₂ a) (fin.cast _)

In the previous section, we have discussed inserting or removing elements on the left of a tuple. In this section, we do the same on the right. A difference is that fin (n+1) is constructed inductively from fin n starting from the left, not from the right. This implies that Lean needs more help to realize that elements belong to the right types, i.e., we need to insert casts at several places.

def fin.init {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (i : fin n) :

The beginning of an n+1 tuple, i.e., its first n entries

Equations
theorem fin.init_def {n : } {α : fin (n + 1) Type u_1} {q : Π (i : fin (n + 1)), α i} :
fin.init (λ (k : fin (n + 1)), q k) = λ (k : fin n), q (fin.cast_succ k)
def fin.snoc {n : } {α : fin (n + 1) Type u} (p : Π (i : fin n), α (fin.cast_succ i)) (x : α (fin.last n)) (i : fin (n + 1)) :
α i

Adding an element at the end of an n-tuple, to get an n+1-tuple. The name snoc comes from cons (i.e., adding an element to the left of a tuple) read in reverse order.

Equations
@[simp]
theorem fin.init_snoc {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (i : fin n), α (fin.cast_succ i)) :
@[simp]
theorem fin.snoc_cast_succ {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (i : fin n), α (fin.cast_succ i)) (i : fin n) :
@[simp]
theorem fin.snoc_comp_cast_succ {n : } {α : Type u_1} {a : α} {f : fin n α} :
@[simp]
theorem fin.snoc_last {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (i : fin n), α (fin.cast_succ i)) :
fin.snoc p x (fin.last n) = x
@[simp]
theorem fin.snoc_comp_nat_add {n m : } {α : Type u_1} (f : fin (m + n) α) (a : α) :
@[simp]
theorem fin.snoc_cast_add {m n : } {α : fin (n + m + 1) Type u_1} (f : Π (i : fin (n + m)), α (fin.cast_succ i)) (a : α (fin.last (n + m))) (i : fin n) :
fin.snoc f a ((fin.cast_add (m + 1)) i) = f ((fin.cast_add m) i)
@[simp]
theorem fin.snoc_comp_cast_add {n m : } {α : Type u_1} (f : fin (n + m) α) (a : α) :
@[simp]
theorem fin.snoc_update {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (i : fin n), α (fin.cast_succ i)) (i : fin n) (y : α (fin.cast_succ i)) :

Updating a tuple and adding an element at the end commute.

theorem fin.update_snoc_last {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (i : fin n), α (fin.cast_succ i)) (z : α (fin.last n)) :

Adding an element at the beginning of a tuple and then updating it amounts to adding it directly.

@[simp]
theorem fin.snoc_init_self {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) :
fin.snoc (fin.init q) (q (fin.last n)) = q

Concatenating the first element of a tuple with its tail gives back the original tuple

@[simp]
theorem fin.init_update_last {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (z : α (fin.last n)) :

Updating the last element of a tuple does not change the beginning.

@[simp]
theorem fin.init_update_cast_succ {n : } {α : fin (n + 1) Type u} (q : Π (i : fin (n + 1)), α i) (i : fin n) (y : α (fin.cast_succ i)) :

Updating an element and taking the beginning commute.

theorem fin.tail_init_eq_init_tail {n : } {β : Type u_1} (q : fin (n + 2) β) :

tail and init commute. We state this lemma in a non-dependent setting, as otherwise it would involve a cast to convince Lean that the two types are equal, making it harder to use.

theorem fin.cons_snoc_eq_snoc_cons {n : } {β : Type u_1} (a : β) (q : fin n β) (b : β) :

cons and snoc commute. We state this lemma in a non-dependent setting, as otherwise it would involve a cast to convince Lean that the two types are equal, making it harder to use.

theorem fin.comp_snoc {n : } {α : Type u_1} {β : Type u_2} (g : α β) (q : fin n α) (y : α) :
g fin.snoc q y = fin.snoc (g q) (g y)
theorem fin.append_right_eq_snoc {α : Type u_1} {n : } (x : fin n α) (x₀ : fin 1 α) :
fin.append x x₀ = fin.snoc x (x₀ 0)

Appending a one-tuple to the right is the same as fin.snoc.

theorem fin.comp_init {n : } {α : Type u_1} {β : Type u_2} (g : α β) (q : fin n.succ α) :
def fin.succ_above_cases {n : } {α : fin (n + 1) Sort u} (i : fin (n + 1)) (x : α i) (p : Π (j : fin n), α ((i.succ_above) j)) (j : fin (n + 1)) :
α j

Define a function on fin (n + 1) from a value on i : fin (n + 1) and values on each fin.succ_above i j, j : fin n. This version is elaborated as eliminator and works for propositions, see also fin.insert_nth for a version without an @[elab_as_eliminator] attribute.

Equations
theorem fin.forall_iff_succ_above {n : } {p : fin (n + 1) Prop} (i : fin (n + 1)) :
( (j : fin (n + 1)), p j) p i (j : fin n), p ((i.succ_above) j)
def fin.insert_nth {n : } {α : fin (n + 1) Type u} (i : fin (n + 1)) (x : α i) (p : Π (j : fin n), α ((i.succ_above) j)) (j : fin (n + 1)) :
α j

Insert an element into a tuple at a given position. For i = 0 see fin.cons, for i = fin.last n see fin.snoc. See also fin.succ_above_cases for a version elaborated as an eliminator.

Equations
@[simp]
theorem fin.insert_nth_apply_same {n : } {α : fin (n + 1) Type u} (i : fin (n + 1)) (x : α i) (p : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth x p i = x
@[simp]
theorem fin.insert_nth_apply_succ_above {n : } {α : fin (n + 1) Type u} (i : fin (n + 1)) (x : α i) (p : Π (j : fin n), α ((i.succ_above) j)) (j : fin n) :
i.insert_nth x p ((i.succ_above) j) = p j
@[simp]
theorem fin.insert_nth_comp_succ_above {n : } {β : Type v} (i : fin (n + 1)) (x : β) (p : fin n β) :
theorem fin.insert_nth_eq_iff {n : } {α : fin (n + 1) Type u} {i : fin (n + 1)} {x : α i} {p : Π (j : fin n), α ((i.succ_above) j)} {q : Π (j : fin (n + 1)), α j} :
i.insert_nth x p = q q i = x p = λ (j : fin n), q ((i.succ_above) j)
theorem fin.eq_insert_nth_iff {n : } {α : fin (n + 1) Type u} {i : fin (n + 1)} {x : α i} {p : Π (j : fin n), α ((i.succ_above) j)} {q : Π (j : fin (n + 1)), α j} :
q = i.insert_nth x p q i = x p = λ (j : fin n), q ((i.succ_above) j)
theorem fin.insert_nth_apply_below {n : } {α : fin (n + 1) Type u} {i j : fin (n + 1)} (h : j < i) (x : α i) (p : Π (k : fin n), α ((i.succ_above) k)) :
i.insert_nth x p j = _.rec_on (p (j.cast_lt _))
theorem fin.insert_nth_apply_above {n : } {α : fin (n + 1) Type u} {i j : fin (n + 1)} (h : i < j) (x : α i) (p : Π (k : fin n), α ((i.succ_above) k)) :
i.insert_nth x p j = _.rec_on (p (j.pred _))
theorem fin.insert_nth_zero {n : } {α : fin (n + 1) Type u} (x : α 0) (p : Π (j : fin n), α ((0.succ_above) j)) :
0.insert_nth x p = fin.cons x (λ (j : fin n), cast _ (p j))
@[simp]
theorem fin.insert_nth_zero' {n : } {β : Type v} (x : β) (p : fin n β) :
theorem fin.insert_nth_last {n : } {α : fin (n + 1) Type u} (x : α (fin.last n)) (p : Π (j : fin n), α (((fin.last n).succ_above) j)) :
(fin.last n).insert_nth x p = fin.snoc (λ (j : fin n), cast _ (p j)) x
@[simp]
theorem fin.insert_nth_last' {n : } {β : Type v} (x : β) (p : fin n β) :
@[simp]
theorem fin.insert_nth_zero_right {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), has_zero (α j)] (i : fin (n + 1)) (x : α i) :
theorem fin.insert_nth_binop {n : } {α : fin (n + 1) Type u} (op : Π (j : fin (n + 1)), α j α j α j) (i : fin (n + 1)) (x y : α i) (p q : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth (op i x y) (λ (j : fin n), op ((i.succ_above) j) (p j) (q j)) = λ (j : fin (n + 1)), op j (i.insert_nth x p j) (i.insert_nth y q j)
@[simp]
theorem fin.insert_nth_mul {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), has_mul (α j)] (i : fin (n + 1)) (x y : α i) (p q : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth (x * y) (p * q) = i.insert_nth x p * i.insert_nth y q
@[simp]
theorem fin.insert_nth_add {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), has_add (α j)] (i : fin (n + 1)) (x y : α i) (p q : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth (x + y) (p + q) = i.insert_nth x p + i.insert_nth y q
@[simp]
theorem fin.insert_nth_div {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), has_div (α j)] (i : fin (n + 1)) (x y : α i) (p q : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth (x / y) (p / q) = i.insert_nth x p / i.insert_nth y q
@[simp]
theorem fin.insert_nth_sub {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), has_sub (α j)] (i : fin (n + 1)) (x y : α i) (p q : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth (x - y) (p - q) = i.insert_nth x p - i.insert_nth y q
@[simp]
theorem fin.insert_nth_sub_same {n : } {α : fin (n + 1) Type u} [Π (j : fin (n + 1)), add_group (α j)] (i : fin (n + 1)) (x y : α i) (p : Π (j : fin n), α ((i.succ_above) j)) :
i.insert_nth x p - i.insert_nth y p = pi.single i (x - y)
theorem fin.insert_nth_le_iff {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {i : fin (n + 1)} {x : α i} {p : Π (j : fin n), α ((i.succ_above) j)} {q : Π (j : fin (n + 1)), α j} :
i.insert_nth x p q x q i p λ (j : fin n), q ((i.succ_above) j)
theorem fin.le_insert_nth_iff {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {i : fin (n + 1)} {x : α i} {p : Π (j : fin n), α ((i.succ_above) j)} {q : Π (j : fin (n + 1)), α j} :
q i.insert_nth x p q i x (λ (j : fin n), q ((i.succ_above) j)) p
theorem fin.insert_nth_mem_Icc {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {i : fin (n + 1)} {x : α i} {p : Π (j : fin n), α ((i.succ_above) j)} {q₁ q₂ : Π (j : fin (n + 1)), α j} :
i.insert_nth x p set.Icc q₁ q₂ x set.Icc (q₁ i) (q₂ i) p set.Icc (λ (j : fin n), q₁ ((i.succ_above) j)) (λ (j : fin n), q₂ ((i.succ_above) j))
theorem fin.preimage_insert_nth_Icc_of_mem {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {i : fin (n + 1)} {x : α i} {q₁ q₂ : Π (j : fin (n + 1)), α j} (hx : x set.Icc (q₁ i) (q₂ i)) :
i.insert_nth x ⁻¹' set.Icc q₁ q₂ = set.Icc (λ (j : fin n), q₁ ((i.succ_above) j)) (λ (j : fin n), q₂ ((i.succ_above) j))
theorem fin.preimage_insert_nth_Icc_of_not_mem {n : } {α : fin (n + 1) Type u} [Π (i : fin (n + 1)), preorder (α i)] {i : fin (n + 1)} {x : α i} {q₁ q₂ : Π (j : fin (n + 1)), α j} (hx : x set.Icc (q₁ i) (q₂ i)) :
i.insert_nth x ⁻¹' set.Icc q₁ q₂ =
def fin.find {n : } (p : fin n Prop) [decidable_pred p] :

find p returns the first index n where p n is satisfied, and none if it is never satisfied.

Equations
theorem fin.find_spec {n : } (p : fin n Prop) [decidable_pred p] {i : fin n} (hi : i fin.find p) :
p i

If find p = some i, then p i holds

theorem fin.is_some_find_iff {n : } {p : fin n Prop} [decidable_pred p] :
((fin.find p).is_some) (i : fin n), p i

find p does not return none if and only if p i holds at some index i.

theorem fin.find_eq_none_iff {n : } {p : fin n Prop} [decidable_pred p] :

find p returns none if and only if p i never holds.

theorem fin.find_min {n : } {p : fin n Prop} [decidable_pred p] {i : fin n} (hi : i fin.find p) {j : fin n} (hj : j < i) :
¬p j

If find p returns some i, then p j does not hold for j < i, i.e., i is minimal among the indices where p holds.

theorem fin.find_min' {n : } {p : fin n Prop} [decidable_pred p] {i : fin n} (h : i fin.find p) {j : fin n} (hj : p j) :
i j
theorem fin.nat_find_mem_find {n : } {p : fin n Prop} [decidable_pred p] (h : (i : ) (hin : i < n), p i, hin⟩) :
theorem fin.mem_find_iff {n : } {p : fin n Prop} [decidable_pred p] {i : fin n} :
i fin.find p p i (j : fin n), p j i j
theorem fin.find_eq_some_iff {n : } {p : fin n Prop} [decidable_pred p] {i : fin n} :
fin.find p = option.some i p i (j : fin n), p j i j
theorem fin.mem_find_of_unique {n : } {p : fin n Prop} [decidable_pred p] (h : (i j : fin n), p i p j i = j) {i : fin n} (hi : p i) :
def fin.contract_nth {n : } {α : Type u_1} (j : fin (n + 1)) (op : α α α) (g : fin (n + 1) α) (k : fin n) :
α

Sends (g₀, ..., gₙ) to (g₀, ..., op gⱼ gⱼ₊₁, ..., gₙ).

Equations
theorem fin.contract_nth_apply_of_lt {n : } {α : Type u_1} (j : fin (n + 1)) (op : α α α) (g : fin (n + 1) α) (k : fin n) (h : k < j) :
theorem fin.contract_nth_apply_of_eq {n : } {α : Type u_1} (j : fin (n + 1)) (op : α α α) (g : fin (n + 1) α) (k : fin n) (h : k = j) :
j.contract_nth op g k = op (g (fin.cast_succ k)) (g k.succ)
theorem fin.contract_nth_apply_of_gt {n : } {α : Type u_1} (j : fin (n + 1)) (op : α α α) (g : fin (n + 1) α) (k : fin n) (h : j < k) :
j.contract_nth op g k = g k.succ
theorem fin.contract_nth_apply_of_ne {n : } {α : Type u_1} (j : fin (n + 1)) (op : α α α) (g : fin (n + 1) α) (k : fin n) (hjk : j k) :
j.contract_nth op g k = g ((j.succ_above) k)
theorem fin.sigma_eq_of_eq_comp_cast {α : Type u_1} {a b : Σ (ii : ), fin ii α} (h : a.fst = b.fst) :
a.snd = b.snd (fin.cast h) a = b

To show two sigma pairs of tuples agree, it to show the second elements are related via fin.cast.

theorem fin.sigma_eq_iff_eq_comp_cast {α : Type u_1} {a b : Σ (ii : ), fin ii α} :
a = b (h : a.fst = b.fst), a.snd = b.snd (fin.cast h)

fin.sigma_eq_of_eq_comp_cast as an iff.