mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 09:12:34 +03:00
2ae330785b
* [contrib] Add misc libraries to contrib Expose some `private` function in libs/base that I needed, and seem like their visibility was forgotten I'd appreciate a code review, especially to tell me I'm re-implementing something that's already elsewhere in the library Mostly extending existing functionality: * `Data/Void.idr`: add some utility functions for manipulating absurdity. * `Decidable/Decidable/Extra.idr`: add support for double negation elimination in decidable relations * `Data/Fun/Extra.idr`: + add `application` (total and partil) for n-ary functions + add (slightly) dependent versions of these operations * `Decidable/Order/Strict.idr`: a strict preorder is what you get when you remove the diagonal from a pre-order. For example, `<` is the associated preorder for `<=` over `Nat`. Analogous to `Decidable.Order`. The proof search mechanism struggled a bit, so I had to hack it --- sorry. Eventually we should move `Data.Fun.Extra.Pointwise` to `Data.Vect.Quantifiers` in base but we don't have any interesting uses for it at the moment so it's not urgent. Co-authored by @gallais
22 lines
528 B
Idris
22 lines
528 B
Idris
module Decidable.Decidable
|
|
|
|
import Data.Rel
|
|
import Data.Fun
|
|
|
|
|
|
|
|
||| Interface for decidable n-ary Relations
|
|
public export
|
|
interface Decidable (ts : Vect k Type) (p : Rel ts) where
|
|
total decide : liftRel ts p Dec
|
|
|
|
||| Given a `Decidable` n-ary relation, provides a decision procedure for
|
|
||| this relation.
|
|
decision : (ts : Vect k Type) -> (p : Rel ts) -> (Decidable ts p) => liftRel ts p Dec
|
|
decision ts p = decide {ts} {p}
|
|
|
|
using (a : Type, x : a)
|
|
public export
|
|
data Given : Dec a -> Type where
|
|
Always : Given (Yes x)
|