mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
Merge pull request #465 from memoryruins/case-declarations
Wrap Javascript case clauses in brackets to prevent conflicting declarations
This commit is contained in:
commit
6233bbd583
@ -383,8 +383,8 @@ mutual
|
||||
pure $ nSpaces indent ++ "while(true){\n" ++ !(imperative2es (indent+1) x) ++ "\n" ++ nSpaces indent ++ "}"
|
||||
|
||||
alt2es : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Nat -> (ImperativeExp, ImperativeStatement) -> Core String
|
||||
alt2es indent (e, b) = pure $ nSpaces indent ++ "case " ++ !(impExp2es e) ++ ":\n" ++
|
||||
!(imperative2es (indent+1) b) ++ "\n" ++ nSpaces (indent+1) ++ "break;\n"
|
||||
alt2es indent (e, b) = pure $ nSpaces indent ++ "case " ++ !(impExp2es e) ++ ": {\n" ++
|
||||
!(imperative2es (indent+1) b) ++ "\n" ++ nSpaces (indent+1) ++ "break; }\n"
|
||||
|
||||
static_preamble : List String
|
||||
static_preamble =
|
||||
|
@ -131,6 +131,7 @@ nodeTests
|
||||
, "node011", "node012", "node015", "node017", "node018", "node019" -- node014
|
||||
, "node021" --, "node020"
|
||||
, "reg001"
|
||||
, "syntax001"
|
||||
, "tailrec001"
|
||||
]
|
||||
|
||||
|
27
tests/node/syntax001/caseBlock.idr
Normal file
27
tests/node/syntax001/caseBlock.idr
Normal file
@ -0,0 +1,27 @@
|
||||
module Main
|
||||
|
||||
-- When generating javasript from idris, it is possible that the same name may appear in
|
||||
-- different clauses when translating cases from Idris. To avoid a javascript syntax error
|
||||
-- from occuring due to duplicate declations, each case clause is wrapped in brackets
|
||||
-- to create unique block scopes.
|
||||
|
||||
-- https://developer.mozilla.org/en-US/docs/Web/JavaScript/Reference/Statements/switch
|
||||
|
||||
-- Without wrapping in a block, the test error may look like the following:
|
||||
-- syntax001/build/exec/_tmp_node.js:39
|
||||
-- const x = ({h:1});
|
||||
-- ^
|
||||
-- SyntaxError: Identifier 'x' has already been declared
|
||||
|
||||
test : Bool -> Int
|
||||
test x @ True =
|
||||
case x of
|
||||
_ => 0
|
||||
test x @ False =
|
||||
case x of
|
||||
_ => 1
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
printLn $ test True
|
||||
printLn $ test False
|
2
tests/node/syntax001/expected
Normal file
2
tests/node/syntax001/expected
Normal file
@ -0,0 +1,2 @@
|
||||
0
|
||||
1
|
3
tests/node/syntax001/run
Normal file
3
tests/node/syntax001/run
Normal file
@ -0,0 +1,3 @@
|
||||
$1 --cg node caseBlock.idr -x main
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user