Namespaces were shadowed in all standard codegens.

This allows, for exmaple, to have apostrophes in module names.
Test was added only for chez, however this should be viable for all
targets with `:exec` implemented.
This commit is contained in:
Denis Buzdalov 2020-11-25 21:10:26 +03:00 committed by G. Allais
parent a7cb2745bc
commit 9990b5ad29
8 changed files with 21 additions and 4 deletions

View File

@ -109,7 +109,7 @@ keywordSafe "var" = "var_"
keywordSafe s = s
jsName : Name -> String
jsName (NS ns n) = showNSWithSep "_" ns ++ "_" ++ jsName n
jsName (NS ns n) = jsIdent (showNSWithSep "_" ns) ++ "_" ++ jsName n
jsName (UN n) = keywordSafe $ jsIdent n
jsName (MN n i) = jsIdent n ++ "_" ++ show i
jsName (PV n d) = "pat__" ++ jsName n

View File

@ -87,7 +87,7 @@ cCleanString : String -> String
cCleanString cs = showcCleanString (unpack cs) ""
cName : Name -> String
cName (NS ns n) = showNSWithSep "_" ns ++ "_" ++ cName n
cName (NS ns n) = cCleanString (showNSWithSep "_" ns) ++ "_" ++ cName n
cName (UN n) = cCleanString n
cName (MN n i) = cCleanString n ++ "_" ++ cCleanString (show i)
cName (PV n d) = "pat__" ++ cName n

View File

@ -31,7 +31,7 @@ schString s = concatMap okchar (unpack s)
export
schName : Name -> String
schName (NS ns n) = showNSWithSep "-" ns ++ "-" ++ schName n
schName (NS ns n) = schString (showNSWithSep "-" ns) ++ "-" ++ schName n
schName (UN n) = schString n
schName (MN n i) = schString n ++ "-" ++ show i
schName (PV n d) = "pat--" ++ schName n

View File

@ -44,7 +44,7 @@ idrisTests = MkTestPool []
"basic031", "basic032", "basic033", "basic034", "basic035",
"basic036", "basic037", "basic038", "basic039", "basic040",
"basic041", "basic042", "basic043", "basic044", "basic045",
"basic046", "basic047",
"basic046", "basic047", "basic048",
-- Coverage checking
"coverage001", "coverage002", "coverage003", "coverage004",
"coverage005", "coverage006", "coverage007", "coverage008",

View File

@ -0,0 +1,7 @@
module Module'
function' : Int -> Int
function' x = x + 1
main : IO ()
main = printLn . show $ function' 4

View File

@ -0,0 +1,4 @@
1/1: Building Module' (Module'.idr)
Module'> 2
Module'> Bye for now!
"5"

View File

@ -0,0 +1,2 @@
function' 1
:q

View File

@ -0,0 +1,4 @@
$1 --no-banner --no-color --console-width 0 "Module'.idr" < input
$1 --exec main --cg ${IDRIS2_TESTS_CG} "Module'.idr"
rm -rf build