Idris2Doc : Data.List.AtIndex

Data.List.AtIndex(source)

Indexing into Lists.

`Elem` proofs give existential quantification that a given element
*is* in the list, but not where in the list it is!  Here we give a
predicate that presents proof that a given index in a list, given
by a natural number, will return a certain element.

Definitions

dataAtIndex : a->Lista->Nat->Type
  @AtIndex witnesses the fact that a natural number encodes a membership proof.
It is meant to be used as a runtime-irrelevant gadget to guarantee that the
natural number is indeed a valid index.

Totality: total
Visibility: public export
Constructors:
Z : AtIndexa (a::as) 0
S : AtIndexaasn->AtIndexa (b::as) (Sn)

Hint: 
Uninhabited (AtIndexa [] n)
inverseZ : AtIndexx (y::xs) 0->x=y
  Inversion principle for Z constructor

Totality: total
Visibility: export
inverseS : AtIndexx (y::xs) (Sn) ->AtIndexxxsn
  inversion principle for S constructor

Totality: total
Visibility: export
atIndexUnique : AtIndexaasn->AtIndexbasn->a=b
  For a given list and a given index, there is only one possible value
stored at that index in that list

Totality: total
Visibility: export
find : DecEqa=> (x : a) -> (xs : Lista) ->Dec (SubsetNat (AtIndexxxs))
  Provided that equality is decidable, we can look for the first occurrence
of a value inside of a list

Totality: total
Visibility: public export
lookup : (n : Nat) -> (xs : Lista) ->Dec (Subseta (\x=>AtIndexxxsn))
  Given an index, we can decide whether there is a value corresponding to it

Totality: total
Visibility: public export
inRange : (n : Nat) -> (xs : Lista) -> (0_ : AtIndexxxsn) ->LTEn (lengthxs)
  An AtIndex proof implies that n is less than the length of the list indexed into

Totality: total
Visibility: public export
weakenR : AtIndexxxsn->AtIndexx (xs++ys) n
  An index remains unchanged if we extend the list to the right

Totality: total
Visibility: export
weakenL : (p : SubsetNat (flipHasLengthws)) ->AtIndexxxsn->AtIndexx (ws++xs) (fstp+n)
  An index into `xs` is shifted by `m` if we prepend a list of size `m`
in front of it

Totality: total
Visibility: export
strengthenL : (p : SubsetNat (flipHasLengthxs)) ->ltn (fstp) =True->AtIndexx (xs++ys) n->AtIndexxxsn
  Conversely to `weakenR`, if an index is smaller than the length of
a prefix then it points into that prefix.

Totality: total
Visibility: export
strengthenR : (p : SubsetNat (flipHasLengthws)) ->lte (fstp) n=True->AtIndexx (ws++xs) n->AtIndexxxs (minusn (fstp))
  Conversely to `weakenL`, if an index is larger than the length of
a prefix then it points into the suffix.

Totality: total
Visibility: export