mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 20:51:43 +03:00
Fix repeated argument check
As it was, it could break if the argument was repeated more than twice. When checking dot patterns, we need to check that no further holes are solved, and that the pattern variable doesn't unify with some other pattern variable, but if it had already made progress (either for a good or bad reason) we missed this. Fixes #536
This commit is contained in:
parent
7f960ed938
commit
04e05e3f86
@ -1586,6 +1586,13 @@ checkDots
|
|||||||
ust <- get UST
|
ust <- get UST
|
||||||
put UST (record { dotConstraints = [] } ust)
|
put UST (record { dotConstraints = [] } ust)
|
||||||
where
|
where
|
||||||
|
getHoleName : Term [] -> Core (Maybe Name)
|
||||||
|
getHoleName tm
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
NApp _ (NMeta n' i args) _ <- nf defs [] tm
|
||||||
|
| _ => pure Nothing
|
||||||
|
pure (Just n')
|
||||||
|
|
||||||
checkConstraint : (Name, DotReason, Constraint) -> Core ()
|
checkConstraint : (Name, DotReason, Constraint) -> Core ()
|
||||||
checkConstraint (n, reason, MkConstraint fc wl blocked env x y)
|
checkConstraint (n, reason, MkConstraint fc wl blocked env x y)
|
||||||
= do logTermNF 10 "Dot" env y
|
= do logTermNF 10 "Dot" env y
|
||||||
@ -1595,8 +1602,10 @@ checkDots
|
|||||||
ust <- get UST
|
ust <- get UST
|
||||||
handleUnify
|
handleUnify
|
||||||
(do defs <- get Ctxt
|
(do defs <- get Ctxt
|
||||||
Just olddef <- lookupDefExact n (gamma defs)
|
-- get the hole name that 'n' is currently resolved to,
|
||||||
| Nothing => throw (UndefinedName fc n)
|
-- if indeed it is still a hole
|
||||||
|
(i, _) <- getPosition n (gamma defs)
|
||||||
|
oldholen <- getHoleName (Meta fc n i [])
|
||||||
|
|
||||||
-- Check that what was given (x) matches what was
|
-- Check that what was given (x) matches what was
|
||||||
-- solved by unification (y).
|
-- solved by unification (y).
|
||||||
@ -1605,25 +1614,25 @@ checkDots
|
|||||||
-- must be complete.
|
-- must be complete.
|
||||||
cs <- unify inMatch fc env x y
|
cs <- unify inMatch fc env x y
|
||||||
defs <- get Ctxt
|
defs <- get Ctxt
|
||||||
Just ndef <- lookupDefExact n (gamma defs)
|
|
||||||
| Nothing => throw (UndefinedName fc n)
|
|
||||||
|
|
||||||
-- If the name standing for the dot wasn't solved
|
-- If the name standing for the dot wasn't solved
|
||||||
-- earlier, but is now (even with another metavariable)
|
-- earlier, but is now (even with another metavariable)
|
||||||
-- this is bad (it most likely means there's a non-linear
|
-- this is bad (it most likely means there's a non-linear
|
||||||
-- variable)
|
-- variable)
|
||||||
let hBefore = case olddef of
|
dotSolved <-
|
||||||
Hole _ _ => True -- dot not solved
|
maybe (pure False)
|
||||||
_ => False
|
(\n => do Just ndef <- lookupDefExact n (gamma defs)
|
||||||
let h = case ndef of
|
| Nothing => throw (UndefinedName fc n)
|
||||||
Hole _ _ => True -- dot not solved
|
case ndef of
|
||||||
_ => False
|
Hole _ _ => pure False
|
||||||
|
_ => pure True)
|
||||||
|
oldholen
|
||||||
|
|
||||||
-- If any of the things we solved have the same definition,
|
-- If any of the things we solved have the same definition,
|
||||||
-- we've sneaked a non-linear pattern variable in
|
-- we've sneaked a non-linear pattern variable in
|
||||||
argsSame <- checkArgsSame (namesSolved cs)
|
argsSame <- checkArgsSame (namesSolved cs)
|
||||||
when (not (isNil (constraints cs))
|
when (not (isNil (constraints cs))
|
||||||
|| (hBefore && not h) || argsSame) $
|
|| dotSolved || argsSame) $
|
||||||
throw (InternalError "Dot pattern match fail"))
|
throw (InternalError "Dot pattern match fail"))
|
||||||
(\err =>
|
(\err =>
|
||||||
case err of
|
case err of
|
||||||
|
@ -104,7 +104,7 @@ idrisTests
|
|||||||
"reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014",
|
"reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014",
|
||||||
"reg015", "reg016", "reg017", "reg018", "reg019", "reg020", "reg021",
|
"reg015", "reg016", "reg017", "reg018", "reg019", "reg020", "reg021",
|
||||||
"reg022", "reg023", "reg024", "reg025", "reg026", "reg027", "reg028",
|
"reg022", "reg023", "reg024", "reg025", "reg026", "reg027", "reg028",
|
||||||
"reg029", "reg030", "reg031", "reg032", "reg033",
|
"reg029", "reg030", "reg031", "reg032", "reg033", "reg034",
|
||||||
-- Totality checking
|
-- Totality checking
|
||||||
"total001", "total002", "total003", "total004", "total005",
|
"total001", "total002", "total003", "total004", "total005",
|
||||||
"total006", "total007", "total008", "total009",
|
"total006", "total007", "total008", "total009",
|
||||||
|
6
tests/idris2/reg034/expected
Normal file
6
tests/idris2/reg034/expected
Normal file
@ -0,0 +1,6 @@
|
|||||||
|
1/1: Building void (void.idr)
|
||||||
|
void.idr:18:19--18:20:While processing left hand side of Calc at void.idr:18:1--20:1:
|
||||||
|
Can't match on ?y [no locals in scope] (Non linear pattern variable) at
|
||||||
|
18 Calc {y} ((~~) {z=y} {y=y} der (MkDPair y Refl)) = Calc der
|
||||||
|
^
|
||||||
|
|
3
tests/idris2/reg034/run
Executable file
3
tests/idris2/reg034/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
$1 --check void.idr
|
||||||
|
|
||||||
|
rm -rf build
|
21
tests/idris2/reg034/void.idr
Normal file
21
tests/idris2/reg034/void.idr
Normal file
@ -0,0 +1,21 @@
|
|||||||
|
infixl 0 ~~
|
||||||
|
prefix 1 |~
|
||||||
|
infix 1 ...
|
||||||
|
|
||||||
|
public export
|
||||||
|
(...) : (x : a) -> (y ~=~ x) -> (z : a ** y ~=~ z)
|
||||||
|
(...) x pf = (x ** pf)
|
||||||
|
|
||||||
|
public export
|
||||||
|
data FastDerivation : (x : a) -> (y : b) -> Type where
|
||||||
|
(|~) : (x : a) -> FastDerivation x x
|
||||||
|
(~~) : FastDerivation x y ->
|
||||||
|
(step : (z : c ** y ~=~ z)) -> FastDerivation x z
|
||||||
|
|
||||||
|
public export
|
||||||
|
Calc : {x : a} -> {y : b} -> FastDerivation x y -> x = y
|
||||||
|
Calc (|~ x) = Refl
|
||||||
|
Calc {y} ((~~) {z=y} {y=y} der (MkDPair y Refl)) = Calc der
|
||||||
|
|
||||||
|
bad : Z = S Z
|
||||||
|
bad = Calc $ |~ Z ~~ Z ...(Refl)
|
Loading…
Reference in New Issue
Block a user