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 vector
s.
We define the following operations:
fin.tail
: the tail of ann+1
tuple, i.e., its lastn
entries;fin.cons
: adding an element at the beginning of ann
-tuple, to get ann+1
-tuple;fin.init
: the beginning of ann+1
tuple, i.e., its firstn
entries;fin.snoc
: adding an element at the end of ann
-tuple, to get ann+1
-tuple. The namesnoc
comes fromcons
(i.e., adding an element to the left of a tuple) read in reverse order.fin.insert_nth
: insert an element to a tuple at a given position.fin.find p
: returns the first indexn
wherep n
is satisfied, andnone
if it is never satisfied.fin.append a b
: append two tuples.fin.repeat n a
: repeat a tuplen
times.
As a binary function, fin.cons
is injective.
Recurse on an n+1
-tuple by splitting it into a single element and an n
-tuple.
Equations
- fin.cons_cases h x = cast _ (h (x 0) (fin.tail x))
Recurse on an tuple by splitting into fin.elim0
and fin.cons
.
Equations
- fin.cons_induction h0 h x = fin.cons_cases (λ (x₀ : α) (x : fin n → α), h x₀ x (fin.cons_induction h0 h x)) x
- fin.cons_induction h0 h x = _.mpr h0
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
- fin.append a b = fin.add_cases a b
Repeat a
m
times. For example fin.repeat 2 ![0, 3, 7] = ![0, 3, 7, 0, 3, 7]
.
Equations
- fin.repeat m a i = a i.mod_nat
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.
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.
Updating a tuple and adding an element at the end commute.
Adding an element at the beginning of a tuple and then updating it amounts to adding it directly.
Updating an element and taking the beginning commute.
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.
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.
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
- i.insert_nth x p j = i.succ_above_cases x p j
find p
returns the first index n
where p n
is satisfied, and none
if it is never
satisfied.
Equations
- fin.find p = (fin.find (λ (i : fin n), p (i.cast_lt _))).cases_on (ite (p (fin.last n)) (option.some (fin.last n)) option.none) (λ (i : fin n), option.some (i.cast_lt _))
- fin.find p = option.none
If find p = some i
, then p i
holds
find p
returns none
if and only if p i
never holds.