Lists from functions #
Theorems and lemmas for dealing with List.ofFn, which converts a function on Fin n to a list
of length n.
Main Statements #
The main statements pertain to lists generated using List.ofFn
- List.length_ofFn, which tells us the length of such a list
- List.get?_ofFn, which tells us the nth element of such a list
- List.equivSigmaTuple, which is an- Equivbetween lists and the functions that generate them via- List.ofFn.
theorem
List.get_ofFn_go
{α : Type u}
{n : ℕ}
(f : Fin n → α)
(i : ℕ)
(j : ℕ)
(h : i + j = n)
(k : ℕ)
(hk : k < (List.ofFn.go f i j h).length)
 :
(List.ofFn.go f i j h).get ⟨k, hk⟩ = f ⟨j + k, ⋯⟩
@[simp]
theorem
List.get?_ofFn
{α : Type u}
{n : ℕ}
(f : Fin n → α)
(i : ℕ)
 :
(List.ofFn f).get? i = List.ofFnNthVal f i
The nth element of a list
theorem
List.ofFn_add
{α : Type u}
{m : ℕ}
{n : ℕ}
(f : Fin (m + n) → α)
 :
List.ofFn f = (List.ofFn fun (i : Fin m) => f (Fin.castAdd n i)) ++ List.ofFn fun (j : Fin n) => f (Fin.natAdd m j)
Note this matches the convention of List.ofFn_succ', putting the Fin m elements first.
@[simp]
theorem
List.ofFn_const
{α : Type u}
(n : ℕ)
(c : α)
 :
(List.ofFn fun (x : Fin n) => c) = List.replicate n c
@[simp]
theorem
List.ofFn_fin_repeat
{α : Type u}
{m : ℕ}
(a : Fin m → α)
(n : ℕ)
 :
List.ofFn (Fin.repeat n a) = (List.replicate n (List.ofFn a)).join
@[simp]
theorem
List.equivSigmaTuple_apply_fst
{α : Type u}
(l : List α)
 :
(List.equivSigmaTuple l).fst = l.length
Lists are equivalent to the sigma type of tuples of a given length.
Equations
Instances For
def
List.ofFnRec
{α : Type u}
{C : List α → Sort u_1}
(h : (n : ℕ) → (f : Fin n → α) → C (List.ofFn f))
(l : List α)
 :
C l
A recursor for lists that expands a list into a function mapping to its elements.
This can be used with induction l using List.ofFnRec.
Equations
- List.ofFnRec h l = cast ⋯ (h l.length l.get)
Instances For
Note we can only state this when the two functions are indexed by defeq n.