mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 15:07:37 +03:00
[ fix ] add missing recursive calls
This commit is contained in:
parent
3f7cd84ccb
commit
86bd19b64d
@ -1793,7 +1793,7 @@ addMetas res ns (TDelay fc x t y)
|
||||
= addMetas res (addMetas res ns t) y
|
||||
addMetas res ns (TForce fc r x) = addMetas res ns x
|
||||
addMetas res ns (PrimVal fc c) = ns
|
||||
addMetas res ns (Erased fc i) = ns
|
||||
addMetas res ns (Erased fc i) = foldr (flip $ addMetas res) ns i
|
||||
addMetas res ns (TType fc u) = ns
|
||||
|
||||
-- Get the metavariable names in a term
|
||||
@ -1828,7 +1828,7 @@ addRefs ua at ns (TDelay fc x t y)
|
||||
= addRefs ua at (addRefs ua at ns t) y
|
||||
addRefs ua at ns (TForce fc r x) = addRefs ua at ns x
|
||||
addRefs ua at ns (PrimVal fc c) = ns
|
||||
addRefs ua at ns (Erased fc i) = ns
|
||||
addRefs ua at ns (Erased fc i) = foldr (flip $ addRefs ua at) ns i
|
||||
addRefs ua at ns (TType fc u) = ns
|
||||
|
||||
-- As above, but for references. Also flag whether a name is under an
|
||||
|
@ -282,6 +282,7 @@ unifyArgs mode loc env _ _ = ufail loc ""
|
||||
getVars : {vars : _} ->
|
||||
List Nat -> List (NF vars) -> Maybe (List (Var vars))
|
||||
getVars got [] = Just []
|
||||
getVars got (NErased fc (Dotted t) :: xs) = getVars got (t :: xs)
|
||||
getVars got (NApp fc (NLocal r idx v) [] :: xs)
|
||||
= if inArgs idx got then Nothing
|
||||
else do xs' <- getVars (idx :: got) xs
|
||||
@ -639,6 +640,8 @@ solveIfUndefined env metavar@(Meta fc mname idx args) soln
|
||||
Just (Hole _ _) <- lookupDefExact (Resolved idx) (gamma defs)
|
||||
| _ => pure False
|
||||
updateSolution env metavar soln
|
||||
solveIfUndefined env (Erased _ (Dotted metavar)) soln
|
||||
= solveIfUndefined env metavar soln
|
||||
solveIfUndefined env metavar soln
|
||||
= pure False
|
||||
|
||||
|
Loading…
Reference in New Issue
Block a user