[ cleanup ] Fewer assert_total in deriving Functor

This commit is contained in:
Guillaume Allais 2022-07-04 10:08:03 +01:00 committed by G. Allais
parent aeeb338e6c
commit 7db20d38a3
2 changed files with 5 additions and 2 deletions

View File

@ -446,7 +446,10 @@ namespace Functor
res <- withError (WhenCheckingArg (mapTTImp cleanup arg)) $
typeView f paras para arg
pure $ case res of
Left sp => functorFun fc mutualWith Nothing sp mapName funName (Just v)
Left sp => -- do not bother with assert_total if you're generating
-- a covering/partial definition
let useTot = False <$ guard (treq /= Total) in
functorFun fc mutualWith useTot sp mapName funName (Just v)
Right free => v
pure $ PatClause fc
(apply fc (IVar fc mapName) [ fun, apply fc (IVar fc cName) vars])

View File

@ -67,7 +67,7 @@ LOG derive.functor.clauses:1:
LOG derive.functor.assumption:10: I am assuming that the parameter layer is a Bifunctor
LOG derive.functor.clauses:1:
mapTreeT : {0 layer : _} -> Bifunctor layer => {0 a, b : Type} -> (a -> b) -> TreeT layer a -> TreeT layer b
mapTreeT f (MkTreeT x0) = MkTreeT (bimap f (assert_total (mapTreeT f)) x0)
mapTreeT f (MkTreeT x0) = MkTreeT (bimap f (mapTreeT f) x0)
LOG derive.functor.clauses:1:
mapTree : {0 a, b : Type} -> (a -> b) -> Tree a -> Tree b
mapTree f (MkTree x0) = MkTree (map f x0)