mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 16:51:51 +03:00
[ fix ] implicits are not in scope under an implicit parameter
This commit is contained in:
parent
9e84b153bd
commit
c9f7f308cf
@ -728,7 +728,7 @@ implicitsAs n defs ns tm
|
||||
-- corresponding to the variables bound in the parent function.
|
||||
-- So we first peel off all of the explicit quantifiers corresponding
|
||||
-- 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))
|
||||
findImps ns es locals body
|
||||
-- ^ TODO? check that name of the pi matches name of local?
|
||||
|
7
tests/idris2/reg/reg053/ImplicitParam.idr
Normal file
7
tests/idris2/reg/reg053/ImplicitParam.idr
Normal 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
|
||||
|
4
tests/idris2/reg/reg053/Issue2444.idr
Normal file
4
tests/idris2/reg/reg053/Issue2444.idr
Normal file
@ -0,0 +1,4 @@
|
||||
parameters {0 res : Type}
|
||||
fun : {x : Nat} -> Nat
|
||||
fun = x
|
||||
|
8
tests/idris2/reg/reg053/expected
Normal file
8
tests/idris2/reg/reg053/expected
Normal 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)
|
2
tests/idris2/reg/reg053/input
Normal file
2
tests/idris2/reg/reg053/input
Normal file
@ -0,0 +1,2 @@
|
||||
:ti hole
|
||||
:q
|
3
tests/idris2/reg/reg053/run
Normal file
3
tests/idris2/reg/reg053/run
Normal file
@ -0,0 +1,3 @@
|
||||
. ../../../testutils.sh
|
||||
idris2 ImplicitParam.idr < input
|
||||
check Issue2444.idr
|
Loading…
Reference in New Issue
Block a user