2021-01-16 17:18:34 +03:00
|
|
|
module Data.Fun.Graph
|
|
|
|
|
|
|
|
||| A relation corresponding to the graph of `f`.
|
|
|
|
public export
|
|
|
|
record Graph {0 a : Type} {0 b : a -> Type}
|
|
|
|
(f : (x : a) -> b x) (x : a) (y : b x) where
|
|
|
|
constructor MkGraph
|
2021-03-25 13:02:06 +03:00
|
|
|
equality : f x === y
|
2021-01-16 17:18:34 +03:00
|
|
|
|
|
|
|
||| An alternative for 'Syntax.WithProof' that allows to keep the
|
|
|
|
||| proof certificate in non-reduced form after nested matching.
|
|
|
|
||| Inspired by https://agda.github.io/agda-stdlib/README.Inspect.html
|
|
|
|
public export
|
|
|
|
remember : {0 a : Type} -> {0 b : a -> Type} ->
|
|
|
|
(f : (x : a) -> b x) -> (x : a) -> Graph f x (f x)
|
2021-06-16 17:02:26 +03:00
|
|
|
remember _ _ = MkGraph Refl
|