mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 11:05:17 +03:00
[ fix ] Use right reflection constructor name for unambiguing with
This commit is contained in:
parent
ccdd049335
commit
ab3862798b
@ -633,7 +633,7 @@ mutual
|
||||
= do fc' <- reflect fc defs lhs env tfc
|
||||
ns' <- reflect fc defs lhs env ns
|
||||
t' <- reflect fc defs lhs env t
|
||||
appCon fc defs (reflectionttimp "WithUnambigNames") [fc', ns', t']
|
||||
appCon fc defs (reflectionttimp "IWithUnambigNames") [fc', ns', t']
|
||||
|
||||
export
|
||||
Reflect IFieldUpdate where
|
||||
|
@ -215,6 +215,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
|
||||
"reflection001", "reflection002", "reflection003", "reflection004",
|
||||
"reflection005", "reflection006", "reflection007", "reflection008",
|
||||
"reflection009", "reflection010", "reflection011", "reflection012",
|
||||
"reflection013",
|
||||
-- The 'with' rule
|
||||
"with001", "with002", "with004", "with005", "with006",
|
||||
-- with-disambiguation
|
||||
|
12
tests/idris2/reflection013/WithUnambig.idr
Normal file
12
tests/idris2/reflection013/WithUnambig.idr
Normal file
@ -0,0 +1,12 @@
|
||||
import Language.Reflection
|
||||
|
||||
x : TTImp
|
||||
x = `(with Prelude.Nil [])
|
||||
|
||||
%language ElabReflection
|
||||
|
||||
x' : Elab (List Nat)
|
||||
x' = check x
|
||||
|
||||
l : List Nat
|
||||
l = 1 :: %runElab x'
|
1
tests/idris2/reflection013/expected
Normal file
1
tests/idris2/reflection013/expected
Normal file
@ -0,0 +1 @@
|
||||
1/1: Building WithUnambig (WithUnambig.idr)
|
4
tests/idris2/reflection013/run
Executable file
4
tests/idris2/reflection013/run
Executable file
@ -0,0 +1,4 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner --check WithUnambig.idr
|
||||
|
Loading…
Reference in New Issue
Block a user