Idris2/tests/idris2/reg033/DerivingEq.idr
Edwin Brady 04a0f5001f Correct multiplicities when checking Pi binders
We've always just used 0, which isn't correct if the function is going
to be used in a runtime pattern match. Now calculate correctly so that
we're explicit about which type level variables are used at runtime.

This might cause some programs to fail to compile, if they use functions
that calculate Pi types. The solution is to make those functions
explicitly 0 multiplicity. If that doesn't work, you may have been
accidentally trying to use compile-time only data at run time!

Fixes #1163
2021-03-09 17:23:05 +00:00

52 lines
1.7 KiB
Idris

module DerivingEq
import Language.Reflection
%language ElabReflection
public export
countArgs : (ty : TTImp) -> Nat
countArgs (IPi _ _ ExplicitArg _ _ retTy) = 1 + countArgs retTy
countArgs (IPi _ _ _ _ _ retTy) = countArgs retTy
countArgs _ = 0
-- %logging 5
public export
genEq : {t : _} -> Name -> Elab (t -> t -> Bool)
genEq typeName = do
let pos : FC = MkFC "generated code" (0,0) (0,0)
[(n, _)] <- getType typeName
| _ => fail "Ambiguous name"
constrs <- getCons n
let and : TTImp -> TTImp -> TTImp
and x y = `(~(x) && ~(y))
compareEq : String -> String -> TTImp
compareEq x y = `(~(IVar pos $ UN x) == ~(IVar pos $ UN y))
makeClause : Name -> Elab Clause
makeClause constr = do
[(_, ty)] <- getType constr
| _ => fail "ambiguous name for constr"
let nArgs = countArgs ty
let xs = map (\i => "x_" ++ show i) $ take nArgs [1..]
let ys = map (\i => "y_" ++ show i) $ take nArgs [1..]
let px = foldl (IApp pos) (IVar pos constr) $ map (IBindVar pos) xs
let py = foldl (IApp pos) (IVar pos constr) $ map (IBindVar pos) ys
pure $ PatClause pos `(MkPair ~(px) ~(py))
$ foldl and `(True) $ zipWith compareEq xs ys
finalClause : Clause
finalClause = PatClause pos `(_) `(False)
clauses <- traverse makeClause constrs
let allClauses = clauses ++ [finalClause]
caseExpr = ICase pos `(MkPair x y) (Implicit pos True) allClauses
result = `(\x, y => ~(caseExpr))
check result
%logging 0
-- This tree works
data TreeOne a = BranchOne (TreeOne a) a (TreeOne a)
| Leaf
Eq a => Eq (TreeOne a) where
(==) = %runElab genEq `{{ TreeOne }}