Idris2Doc : Data.Vect.Quantifiers

Data.Vect.Quantifiers(source)

Definitions

dataAny : (0_ : (a->Type)) ->Vectna->Type
  A proof that some element of a vector satisfies some property

@ p the property to be satsified

Totality: total
Visibility: public export
Constructors:
Here : px->Anyp (x::xs)
  A proof that the satisfying element is the first one in the `Vect`
There : Anypxs->Anyp (x::xs)
  A proof that the satsifying element is in the tail of the `Vect`

Hints:
All (Eq.p) xs=>Eq (Anypxs)
All (Show.p) xs=>Show (Anypxs)
All (Uninhabited.p) xs=>Uninhabited (Anypxs)
dataAll : (0_ : (a->Type)) ->Vectna->Type
  A proof that all elements of a vector satisfy a property. It is a list of
proofs, corresponding element-wise to the `Vect`.

Totality: total
Visibility: public export
Constructors:
Nil : Allp []
(::) : px->Allpxs->Allp (x::xs)

Hints:
All (Ord.p) xs=>All (Eq.p) xs
All (Monoid.p) xs=>All (Semigroup.p) xs
All (DecEq.p) xs=>DecEq (Allpxs)
All (Eq.p) xs=>Eq (Anypxs)
All (Eq.p) xs=>Eq (Allpxs)
All (Monoid.p) xs=>Monoid (Allpxs)
All (Ord.p) xs=>Ord (Allpxs)
All (Semigroup.p) xs=>Semigroup (Allpxs)
All (Show.p) xs=>Show (Anypxs)
All (Show.p) xs=>Show (Allpxs)
All (Uninhabited.p) xs=>Uninhabited (Anypxs)
Either (Uninhabited (px)) (Uninhabited (Allpxs)) =>Uninhabited (Allp (x::xs))
anyElim : (Anypxs->b) -> (px->b) ->Anyp (x::xs) ->b
  Eliminator for `Any`

Totality: total
Visibility: public export
any : ((x : a) ->Dec (px)) -> (xs : Vectna) ->Dec (Anypxs)
  Given a decision procedure for a property, determine if an element of a
vector satisfies it.

@ p the property to be satisfied
@ dec the decision procedure
@ xs the vector to examine

Totality: total
Visibility: public export
mapProperty : (px->qx) ->Anypl->Anyql
Totality: total
Visibility: export
toDPair : Anypxs->DPairap
  Forget the membership proof

Totality: total
Visibility: public export
toExists : Anypxs->Existsp
Totality: total
Visibility: export
anyToFin : Anypxs->Finn
  Get the bounded numeric position of the element satisfying the predicate

Totality: total
Visibility: public export
anyToFinCorrect : (witness : Anypxs) ->p (index (anyToFinwitness) xs)
  `anyToFin`'s return type satisfies the predicate

Totality: total
Visibility: export
anyToFinV : Anypxs-> (idx : Finn**p (indexidxxs))
Totality: total
Visibility: public export
pushOut : Functorp=>Any (p.q) xs->p (Anyqxs)
Totality: total
Visibility: public export
notAllHere : Not (px) ->Not (Allp (x::xs))
Totality: total
Visibility: export
notAllThere : Not (Allpxs) ->Not (Allp (x::xs))
Totality: total
Visibility: export
all : ((x : a) ->Dec (px)) -> (xs : Vectna) ->Dec (Allpxs)
  Given a decision procedure for a property, decide whether all elements of
a vector satisfy it.

@ p the property
@ dec the decision procedure
@ xs the vector to examine

Totality: total
Visibility: public export
mapProperty : (px->qx) ->Allpl->Allql
Totality: total
Visibility: export
mapPropertyRelevant : ((x : a) ->px->qx) ->Allpxs->Allqxs
  A variant of `mapProperty` that also allows accessing
the values of `xs` that the corresponding `ps` prove `p` about.

Totality: total
Visibility: export
imapProperty : (0i : (a->Type)) -> (ix=>px->qx) ->Allias=>Allpas->Allqas
Totality: total
Visibility: public export
forget : All (constp) xs->Vectnp
  If `All` witnesses a property that does not depend on the vector `xs`
it's indexed by, then it is really a `Vect`.

Totality: total
Visibility: public export
remember : (xs : Vectnty) ->All (constty) xs
  Any `Vect` can be lifted to become an `All`
witnessing the presence of elements of the `Vect`'s type.

Totality: total
Visibility: public export
forgetRememberId : (xs : Vectnty) ->forget (rememberxs) =xs
Totality: total
Visibility: export
castAllConst : All (constty) xs->All (constty) ys
Totality: total
Visibility: public export
rememberForgetId : (vs : All (constty) xs) ->castAllConst (remember (forgetvs)) =vs
Totality: total
Visibility: export
zipPropertyWith : (px->qx->rx) ->Allpxs->Allqxs->Allrxs
Totality: total
Visibility: export
traverseProperty : Applicativef=> (px->f (qx)) ->Allpxs->f (Allqxs)
  A `Traversable`'s `traverse` for `All`,
for traversals that don't care about the values of the associated `Vect`.

Totality: total
Visibility: export
traversePropertyRelevant : Applicativef=> ((x : a) ->px->f (qx)) ->Allpxs->f (Allqxs)
  A `Traversable`'s `traverse` for `All`,
in case the elements of the `Vect` that the `All` is proving `p` about are also needed.

Totality: total
Visibility: export
consInjective : x::xs=y::ys-> (x=y, xs=ys)
Totality: total
Visibility: export
tabulate : ((ix : Finn) ->p (indexixxs)) ->Allpxs
Totality: total
Visibility: public export
(++) : Allpxs->Allpys->Allp (xs++ys)
Totality: total
Visibility: public export
Fixity Declaration: infixr operator, level 7
HVect : VectnType->Type
  A heterogeneous vector of arbitrary types

Totality: total
Visibility: public export
head : Allp (x::xs) ->px
  Take the first element.

Totality: total
Visibility: public export
tail : Allp (x::xs) ->Allpxs
  Take all but the first element.

Totality: total
Visibility: public export
drop : (n : Nat) ->Allpxs->Allp (the (Vectma) (dropnxs))
  Drop the first n elements given knowledge that
there are at least n elements available.

Totality: total
Visibility: export
drop' : (l : Nat) ->Allpxs->Allp (drop'lxs)
  Drop up to the first l elements, stopping early
if all elements have been dropped.

Totality: total
Visibility: export
index : (i : Fink) ->Allpts->p (indexits)
  Extract an element from an All.

```idris example
> index 0 (the (HVect _) [1, "string"])
1
```

Totality: total
Visibility: export
deleteAt : (i : Fin (Sl)) ->Allpts->Allp (deleteAtits)
  Delete an element from an All at the given position.

```idris example
> deleteAt 0 (the (HVect _) [1, "string"])
["string"]
```

Totality: total
Visibility: export
replaceAt : (i : Fink) ->pt->Allpts->Allp (replaceAtitts)
  Replace an element in an All at the given position.

```idris example
> replaceAt 0 "firstString" (the (HVect _) [1, "string"])
["firstString", "string"]
```

Totality: total
Visibility: export
updateAt : (i : Fink) -> (p (indexits) ->pt) ->Allpts->Allp (replaceAtitts)
  Update an element in an All at the given position.

```idris example
> updateAt 0 (const True) (the (HVect _) [1, "string"])
[True, "string"]
```

Totality: total
Visibility: export
get : Allpts->Elemxts->px
  Extract an element of an All.

```idris example
> get [1, "string"] {p = Here}
1
```

Totality: total
Visibility: export
replaceElem : Allpts-> (e : Elemtts) ->pt'->Allp (replaceByElemtset')
  Replace an element of an All.

```idris example
> replaceElem 2 [1, "string"]
[2, "string"]
```

Totality: total
Visibility: export
updateElem : (pt->pt') ->Allpts-> (e : Elemtts) ->Allp (replaceByElemtset')
  Update an element of an All.

```idris example
> updateElem (const "hello world!") [1, "string"]
[1, "hello world!"]
```

Totality: total
Visibility: public export
pushIn : (xs : Vectna) -> (0_ : Allpxs) ->Vectn (Subsetap)
  Push in the property from the list level with element level

Totality: total
Visibility: public export
pullOut : Vectn (Subsetap) ->Subset (Vectna) (Allp)
  Pull the elementwise property out to the list level

Totality: total
Visibility: public export
pushInOutInverse : (xs : Vectna) -> (0prf : Allpxs) ->pullOut (pushInxsprf) =Elementxsprf
Totality: total
Visibility: export
pushOutInInverse : (xs : Vectn (Subsetap)) ->uncurrypushIn (pullOutxs) =xs
Totality: total
Visibility: export
pushOut : Applicativep=>All (p.q) xs->p (Allqxs)
Totality: total
Visibility: public export
negAnyAll : Not (Anypxs) ->All (Not.p) xs
  If there does not exist an element that satifies the property, then it is
the case that all elements do not satisfy.

Totality: total
Visibility: export
anyNegAll : Any (Not.p) xs->Not (Allpxs)
  If there exists an element that doesn't satify the property, then it is
not the case that all elements satisfy it.

Totality: total
Visibility: export
allNegAny : All (Not.p) xs->Not (Anypxs)
  If none of the elements satisfy the property, then not any single one can.

Totality: total
Visibility: export
allAnies : Allpxs->Vectn (Anypxs)
Totality: total
Visibility: public export
altAll : Alternativep=>All (p.q) xs->p (Anyqxs)
Totality: total
Visibility: public export