Idris2Doc : Search.HDecidable

Search.HDecidable(source)

The content of this module is based on the paper
Applications of Applicative Proof Search
by Liam O'Connor
https://doi.org/10.1145/2976022.2976030

Definitions

recordHDec : Type->Type
  Half a decider: when the search succeeds we bother building the proof

Totality: total
Visibility: public export
Constructor: 
MkHDec : (isTrue : Bool) -> (SoisTrue->a) ->HDeca

Projections:
.evidence : ({rec:0} : HDeca) ->So (isTrue{rec:0}) ->a
.isTrue : HDeca->Bool

Hints:
AlternativeHDec
AnHDecHDec
ApplicativeHDec
FunctorHDec
MonadHDec
Showf=>Show (HDecf)
.isTrue : HDeca->Bool
Totality: total
Visibility: public export
isTrue : HDeca->Bool
Totality: total
Visibility: public export
.evidence : ({rec:0} : HDeca) ->So (isTrue{rec:0}) ->a
Totality: total
Visibility: public export
evidence : ({rec:0} : HDeca) ->So (isTrue{rec:0}) ->a
Totality: total
Visibility: public export
yes : a->HDeca
  Happy path: we have found a proof!

Totality: total
Visibility: public export
no : HDeca
  Giving up

Totality: total
Visibility: public export
fromDec : Deca->HDeca
Totality: total
Visibility: public export
fromMaybe : Maybea->HDeca
Totality: total
Visibility: public export
toMaybe : HDeca->Maybea
Totality: total
Visibility: public export
interfaceAnHDec : (Type->Type) ->Type
  A type constructor satisfying AnHdec is morally an HDec i.e. we can
turn values of this type constructor into half deciders
It may be more powerful (like Dec) or more basic (like Maybe).

Parameters: t
Methods:
toHDec : ta->HDeca

Implementations:
AnHDecDec
AnHDecHDec
AnHDecMaybe
toHDec : AnHDect=>ta->HDeca
Totality: total
Visibility: public export
(&&) : (AnHDecl, AnHDecr) =>la->rb->HDec (a, b)
  Half deciders are closed under product

Totality: total
Visibility: public export
Fixity Declaration: infixr operator, level 5
(||) : (AnHDecl, AnHDecr) =>la->rb->HDec (Eitherab)
  Half deciders are closed under sum

Totality: total
Visibility: public export
Fixity Declaration: infixr operator, level 4
not : (AnHDecl, Negatesnaa) =>lna->HDec (Nota)
  Half deciders are closed negation. Here we use the `Negates` interface
so that we end up looking for *positive* evidence of something which is
much easier to find than negative one.

Totality: total
Visibility: public export
(==>) : (AnHDecl, (AnHDecr, Negatesnaa)) =>lna->rb->HDec (a->b)
  Half deciders are closed under implication

Totality: total
Visibility: public export
Fixity Declaration: infixr operator, level 3
any : AnHDecl=> (xs : Lista) -> ((x : a) ->l (px)) ->HDec (Anypxs)
  Half deciders are closed under the list quantifier any

Totality: total
Visibility: public export
all : AnHDecl=> (xs : Lista) -> ((x : a) ->l (px)) ->HDec (Allpxs)
  Half deciders are closed under the list quantifier all

Totality: total
Visibility: public export
any : AnHDecl=> (xs : LazyLista) -> ((x : a) ->l (px)) ->HDec (Anypxs)
  Half deciders are closed under the lazy list quantifier any

Totality: total
Visibility: public export