Idris2/tests/base/deriving_foldable/DeriveFoldable.idr

297 lines
6.4 KiB
Idris
Raw Normal View History

2022-09-23 13:47:35 +03:00
module DeriveFoldable
2022-07-04 10:58:18 +03:00
2022-09-23 13:47:35 +03:00
import Deriving.Foldable
2022-07-04 10:58:18 +03:00
%language ElabReflection
%default covering
2022-09-23 13:47:35 +03:00
%logging "derive.foldable.clauses" 1
%logging "derive.foldable.assumption" 10
2022-07-04 10:58:18 +03:00
2022-09-23 13:47:35 +03:00
list : Foldable List
2022-07-04 10:58:18 +03:00
list = %runElab derive
2022-09-23 13:47:35 +03:00
maybe : Foldable Maybe
2022-07-04 10:58:18 +03:00
maybe = %runElab derive
2022-09-23 13:47:35 +03:00
either : Foldable (Either err)
2022-07-04 10:58:18 +03:00
either = %runElab derive
2022-09-23 13:47:35 +03:00
pair : Foldable (Pair a)
pair = %runElab derive
2022-07-04 10:58:18 +03:00
namespace Constant
record Constant (a, b : Type) where
constructor MkConstant
runConstant : a
2022-09-23 13:47:35 +03:00
constant : Foldable (Constant a)
2022-07-04 10:58:18 +03:00
constant = %runElab derive
namespace Vect
public export
data Vect : Nat -> Type -> Type where
Nil : Vect Z a
(::) : a -> Vect n a -> Vect (S n) a
export %hint
total
2022-09-23 13:47:35 +03:00
vect : Foldable (Vect n)
2022-07-04 10:58:18 +03:00
vect = %runElab derive
namespace BigTree
data BigTree a
= End a
| Branch String (List a) (Bool -> BigTree a)
2022-09-23 13:47:35 +03:00
-- this (Bool ->) is problematic, at least until we have an interface for finite types
2022-07-04 10:58:18 +03:00
| Rose (List (BigTree a))
2022-09-23 13:47:35 +03:00
failing "Unsupported type"
total
bigTree : Foldable BigTree
bigTree = %runElab derive
2022-07-04 10:58:18 +03:00
namespace Matrix
record Matrix m n a where
constructor MkMatrix
runMatrix : Vect m (Vect n a)
total
2022-09-23 13:47:35 +03:00
matrix : Foldable (Matrix m n)
2022-07-04 10:58:18 +03:00
matrix = %runElab derive
namespace Tm
data Op : Nat -> Type where
Neg : Op 1
Add : Op 2
data Tm : Type -> Type where
Var : a -> Tm a
Call : Op n -> Vect n (Tm a) -> Tm a
Lam : Tm (Maybe a) -> Tm a
total
2022-09-23 13:47:35 +03:00
tm : Foldable Tm
2022-07-04 10:58:18 +03:00
tm = %runElab derive
namespace Forest
data Tree : Type -> Type
data Forest : Type -> Type
data Tree : Type -> Type where
Leaf : a -> Tree a
Node : Forest a -> Tree a
data Forest : Type -> Type where
Empty : Forest a
Plant : Tree a -> Forest a -> Forest a
2022-09-23 13:47:35 +03:00
%hint total tree : Foldable Tree
%hint total forest : Foldable Forest
2022-07-04 10:58:18 +03:00
tree = %runElab derive {mutualWith = [`{Forest}]}
forest = %runElab derive {mutualWith = [`{Tree}]}
namespace List1
2022-09-23 13:47:35 +03:00
%hide List1
2022-07-04 10:58:18 +03:00
data List1 : Type -> Type where
MkList1 : (a, Maybe (List1 a)) -> List1 a
total
2022-09-23 13:47:35 +03:00
list1 : Foldable List1
2022-07-04 10:58:18 +03:00
list1 = %runElab derive
namespace Full
data Full a = Leaf a | Node (Full (a, a))
total
2022-09-23 13:47:35 +03:00
full : Foldable Full
2022-07-04 10:58:18 +03:00
full = %runElab derive
namespace Colist
data Colist a = Nil | (::) a (Inf (Colist a))
2022-09-23 13:47:35 +03:00
failing "Cannot fold over an infinite structure"
total
colist : Foldable Colist
colist = %runElab derive
2022-07-04 10:58:18 +03:00
namespace LAZY
record LAZY (a : Type) where
constructor MkLAZY
wrapped : Lazy a
total
2022-09-23 13:47:35 +03:00
lazy : Foldable LAZY
2022-07-04 10:58:18 +03:00
lazy = %runElab derive
namespace Rose
data Rose a = Node (List (Lazy (Rose a)))
total
2022-09-23 13:47:35 +03:00
rose : Foldable Rose
2022-07-04 10:58:18 +03:00
rose = %runElab derive
namespace Free
data Free : (Type -> Type) -> Type -> Type where
Pure : a -> Free f a
Bind : f a -> (a -> Free f b) -> Free f b
2022-09-23 13:47:35 +03:00
-- Once again the pi type prevents us from defining Foldable
failing "Unsupported type"
total
free : Foldable (Free f)
free = %runElab derive
2022-07-04 10:58:18 +03:00
namespace MaybeT
record MaybeT (m : Type -> Type) (a : Type) where
constructor MkMaybeT
runMaybeT : m (Maybe a)
total
2022-09-23 13:47:35 +03:00
maybeT : Foldable m => Foldable (MaybeT m)
2022-07-04 10:58:18 +03:00
maybeT = %runElab derive
namespace TreeT
2022-11-02 14:57:07 +03:00
data TreeT : (lbl : Type) -> (Type -> Type -> Type) -> Type -> Type where
MkTreeT : lbl -> layer a (TreeT lbl layer a) -> TreeT lbl layer a
2022-07-04 10:58:18 +03:00
%hint
2022-11-02 14:57:07 +03:00
treeT : Bifoldable layer => Foldable (TreeT lbl layer)
2022-07-04 10:58:18 +03:00
treeT = %runElab derive {treq = CoveringOnly}
record Tree (a : Type) where
constructor MkTree
2022-11-02 14:57:07 +03:00
runTree : TreeT () Either a
2022-07-04 10:58:18 +03:00
2022-09-23 13:47:35 +03:00
tree : Foldable Tree
2022-07-04 10:58:18 +03:00
tree = %runElab derive {treq = CoveringOnly}
namespace Implicit
data IVect : {n : Nat} -> (a : Type) -> Type where
MkIVect : (v : Vect n a) -> IVect {n} a
2022-09-23 13:47:35 +03:00
ivect : {m : Nat} -> Foldable (IVect {n = m})
ivect = %runElab derive
namespace EqMap
data EqMap : (key : Type) -> Eq key => (val : Type) -> Type where
MkEqMap : (eq : Eq key) => List (key, val) -> EqMap key @{eq} val
empty : Eq key => EqMap key val
empty = MkEqMap []
insert : (eq : Eq key) => key -> val -> EqMap key @{eq} val -> EqMap key @{eq} val
insert k v (MkEqMap kvs) = MkEqMap ((k, v) :: filter ((k /=) . fst) kvs)
fromList : Eq key => List (key, val) -> EqMap key val
fromList = foldr (uncurry insert) empty
toList : EqMap key @{eq} val -> List (key, val)
toList (MkEqMap kvs) = kvs
test : EqMap.toList (fromList [(1,2), (1,3), (2, 4), (5, 3), (1, 0)])
=== [(1,2), (2, 4), (5, 3)]
test = Refl
2022-09-23 13:47:35 +03:00
eqMap : (eq : Eq key) => Foldable (EqMap key @{eq})
eqMap = %runElab derive
namespace Implicit
data WithImplicits : Type -> Type where
MkImplicit : {x : a} -> (0 y : a) -> WithImplicits a
OtherImplicit : {0 x : a} -> a => WithImplicits a
LastOne : {auto 0 _ : a} -> a -> WithImplicits a
2022-09-23 13:47:35 +03:00
failing "Cannot fold over a runtime irrelevant value"
total
withImplicits : Foldable WithImplicits
withImplicits = %runElab derive
2022-09-23 13:47:35 +03:00
failing "Couldn't find a `Foldable' instance for the type constructor DeriveFoldable.Wrap"
2022-07-04 10:58:18 +03:00
record Wrap (a : Type) where
constructor MkWrap
unWrap : a
data Indirect : Type -> Type where
MkIndirect : Wrap a -> Indirect a
total
2022-09-23 13:47:35 +03:00
indirect : Foldable Indirect
2022-07-04 10:58:18 +03:00
indirect = %runElab derive
2022-09-23 13:47:35 +03:00
namespace BifoldableFail
2022-07-04 10:58:18 +03:00
data Tree : (l, n : Type) -> Type where
Leaf : l -> Tree l n
Node : Tree l n -> n -> Tree l n -> Tree l n
-- this one will succeed
total
2022-09-23 13:47:35 +03:00
tree : Foldable (Tree l)
2022-07-04 10:58:18 +03:00
tree = %runElab derive
2022-09-23 13:47:35 +03:00
failing "Couldn't find a `Bifoldable' instance for the type constructor DeriveFoldable.BifoldableFail.Tree"
2022-07-04 10:58:18 +03:00
record Tree' (a : Type) where
constructor MkTree'
getTree : Tree a a
-- and this one will fail
2022-09-23 13:47:35 +03:00
tree' : Foldable Tree'
2022-07-04 10:58:18 +03:00
tree' = %runElab derive
failing "Expected a type constructor, got: Prelude.Basics.id {a = Type}"
total
2022-09-23 13:47:35 +03:00
foldable : Foldable Prelude.id
foldable = %runElab derive
namespace Triple
data Triple a b c = MkTriple a b c
%hint
2022-09-23 13:47:35 +03:00
triple : Foldable (Triple a b)
triple = %runElab derive
data Tree3 a = Node (Triple a () (Tree3 a))
2022-09-23 13:47:35 +03:00
failing "The term DeriveFoldable.Triple.Triple a Builtin.Unit is not free of a"
2022-09-23 13:47:35 +03:00
tree : Foldable Tree3
tree = %runElab derive
namespace WriterList
data WList : (w, u, a : Type) -> Type where
(::) : (w, a) -> WList {- oops -} a u a -> WList w u a
Nil : WList w u a
2022-09-23 13:47:35 +03:00
failing "The term DeriveFoldable.WriterList.WList a u is not free of a"
2022-09-23 13:47:35 +03:00
wlist : Foldable (WList w ())
wlist = %runElab derive