mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 11:05:17 +03:00
parent
b3c80e0765
commit
ece1198074
@ -635,7 +635,10 @@ getArgName defs x bound allvars ty
|
||||
findNames nf = pure $ filter notBound $ fromMaybe defaultNames !(findNamesM nf)
|
||||
|
||||
getName : Name -> List String -> List Name -> String
|
||||
getName (UN (Basic n)) defs used = unique (n :: defs) (n :: defs) 0 used
|
||||
getName (UN (Basic n)) defs used =
|
||||
-- # 1742 Uppercase names are not valid for pattern variables
|
||||
let candidate = ifThenElse (lowerFirst n) n (toLower n) in
|
||||
unique (candidate :: defs) (candidate :: defs) 0 used
|
||||
getName _ defs used = unique defs defs 0 used
|
||||
|
||||
export
|
||||
|
@ -107,7 +107,8 @@ idrisTestsInteractive = MkTestPool "Interactive editing" [] Nothing
|
||||
"interactive029", "interactive030", "interactive031", "interactive032",
|
||||
"interactive033", "interactive034", "interactive035", "interactive036",
|
||||
"interactive037", "interactive038", "interactive039", "interactive040",
|
||||
"interactive041", "interactive042", "interactive043", "interactive044"]
|
||||
"interactive041", "interactive042", "interactive043", "interactive044",
|
||||
"interactive045"]
|
||||
|
||||
idrisTestsInterface : TestPool
|
||||
idrisTestsInterface = MkTestPool "Interface" [] Nothing
|
||||
|
5
tests/idris2/interactive045/Issue1742.idr
Normal file
5
tests/idris2/interactive045/Issue1742.idr
Normal file
@ -0,0 +1,5 @@
|
||||
data Foo : Type where
|
||||
MkFoo : (Bar : Bool) -> Foo
|
||||
|
||||
Gnu : Foo -> Nat
|
||||
Gnu x = ?Gnu_rhs_0
|
4
tests/idris2/interactive045/expected
Normal file
4
tests/idris2/interactive045/expected
Normal file
@ -0,0 +1,4 @@
|
||||
1/1: Building Issue1742Gen (Issue1742Gen.idr)
|
||||
Main> Gnu (MkFoo bar) = ?Gnu_rhs_1
|
||||
Main> Loaded file Issue1742Gen.idr
|
||||
Main> Bye for now!
|
3
tests/idris2/interactive045/input
Normal file
3
tests/idris2/interactive045/input
Normal file
@ -0,0 +1,3 @@
|
||||
:cs 5 4 x
|
||||
:r
|
||||
:q
|
8
tests/idris2/interactive045/run
Executable file
8
tests/idris2/interactive045/run
Executable file
@ -0,0 +1,8 @@
|
||||
rm -rf build
|
||||
rm -f Issue1742Gen.idr
|
||||
|
||||
cp Issue1742.idr Issue1742Gen.idr
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner Issue1742Gen.idr < input
|
||||
|
||||
rm Issue1742Gen.idr
|
Loading…
Reference in New Issue
Block a user