[ fix ] implicits are not in scope under an implicit parameter

This commit is contained in:
Steve Dunham 2024-06-12 12:03:25 -07:00
parent 9e84b153bd
commit c9f7f308cf
6 changed files with 25 additions and 1 deletions

View File

@ -728,7 +728,7 @@ implicitsAs n defs ns tm
-- corresponding to the variables bound in the parent function. -- corresponding to the variables bound in the parent function.
-- So we first peel off all of the explicit quantifiers corresponding -- So we first peel off all of the explicit quantifiers corresponding
-- to these variables. -- to these variables.
findImps ns es (_ :: locals) (NBind fc x (Pi _ _ Explicit _) sc) findImps ns es (_ :: locals) (NBind fc x (Pi _ _ _ _) sc)
= do body <- sc defs (toClosure defaultOpts [] (Erased fc Placeholder)) = do body <- sc defs (toClosure defaultOpts [] (Erased fc Placeholder))
findImps ns es locals body findImps ns es locals body
-- ^ TODO? check that name of the pi matches name of local? -- ^ TODO? check that name of the pi matches name of local?

View File

@ -0,0 +1,7 @@
import Data.Vect
-- m was not in scope when n is implicit
parameters {n : Nat}
foo : Vect m Nat -> Nat
foo xs = ?hole

View File

@ -0,0 +1,4 @@
parameters {0 res : Type}
fun : {x : Nat} -> Nat
fun = x

View File

@ -0,0 +1,8 @@
1/1: Building ImplicitParam (ImplicitParam.idr)
Main> {n : Nat}
0 m : Nat
xs : Vect m Nat
------------------------------
hole : Nat
Main> Bye for now!
1/1: Building Issue2444 (Issue2444.idr)

View File

@ -0,0 +1,2 @@
:ti hole
:q

View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
idris2 ImplicitParam.idr < input
check Issue2444.idr