An alternative (Fin-based) indexing function was added for lists.

This commit is contained in:
Denis Buzdalov 2020-12-04 04:49:26 +03:00 committed by G. Allais
parent 4364793484
commit 13cc27da1f
2 changed files with 7 additions and 1 deletions

View File

@ -2,7 +2,7 @@ module Data.Fin
import public Data.Maybe
import Data.Nat
import Decidable.Equality
import Decidable.Equality.Core
%default total

View File

@ -2,6 +2,7 @@ module Data.List
import Data.Nat
import Data.List1
import Data.Fin
public export
isNil : List a -> Bool
@ -62,6 +63,11 @@ index : (n : Nat) -> (xs : List a) -> {auto 0 ok : InBounds n xs} -> a
index Z (x :: xs) {ok = InFirst} = x
index (S k) (_ :: xs) {ok = InLater _} = index k xs
public export
index' : (xs : List a) -> Fin (length xs) -> a
index' (x::_) FZ = x
index' (_::xs) (FS i) = index' xs i
||| Generate a list by repeatedly applying a partial function until exhausted.
||| @ f the function to iterate
||| @ x the initial value that will be the head of the list