Skip to content

Commit

Permalink
[ fix ] Add escapeArg in executeExpr
Browse files Browse the repository at this point in the history
  • Loading branch information
spcfox committed Dec 19, 2024
1 parent 6ac9301 commit ccc7adf
Show file tree
Hide file tree
Showing 6 changed files with 18 additions and 6 deletions.
4 changes: 3 additions & 1 deletion src/Compiler/ES/Node.idr
Original file line number Diff line number Diff line change
Expand Up @@ -66,7 +66,9 @@ executeExpr :
executeExpr c s tmpDir tm =
do out <- compileExpr c s tmpDir tmpDir tm "_tmp_node.js"
node <- coreLift findNode
system $ "\"" ++ node ++ "\" " ++ out -- Windows often have a space in the path.
system $ "\"" ++ node ++ "\" " ++ escapeArg out -- Windows often have a space in the path.
-- TODO: Replace with the following call after fix idris-lang/Idris2#3436
-- system $ [node, out]

||| Codegen wrapper for Node implementation.
export
Expand Down
4 changes: 3 additions & 1 deletion src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -1006,7 +1006,9 @@ compileExpr c s _ outputDir tm outfile =
export
executeExpr : Ref Ctxt Defs -> Ref Syn SyntaxInfo ->
(execDir : String) -> ClosedTerm -> Core ExitCode
executeExpr c s tmpDir tm = system !(compileExpr c s tmpDir tmpDir tm "_tmp_refc")
executeExpr c s tmpDir tm = do
out <- compileExpr c s tmpDir tmpDir tm "_tmp_refc"
system $ escapeArg out

export
codegenRefC : Codegen
Expand Down
4 changes: 3 additions & 1 deletion src/Compiler/Scheme/Chez.idr
Original file line number Diff line number Diff line change
Expand Up @@ -648,7 +648,9 @@ executeExpr :
Ref Ctxt Defs ->
Ref Syn SyntaxInfo ->
(tmpDir : String) -> ClosedTerm -> Core ExitCode
executeExpr c s tmpDir tm = system !(compileExpr False c s tmpDir tmpDir tm "_tmpchez")
executeExpr c s tmpDir tm = do
out <- compileExpr False c s tmpDir tmpDir tm "_tmpchez"
system $ escapeArg out

incCompile :
Ref Ctxt Defs ->
Expand Down
4 changes: 3 additions & 1 deletion src/Compiler/Scheme/ChezSep.idr
Original file line number Diff line number Diff line change
Expand Up @@ -310,7 +310,9 @@ executeExpr :
Ref Ctxt Defs ->
Ref Syn SyntaxInfo ->
(tmpDir : String) -> ClosedTerm -> Core ExitCode
executeExpr c s tmpDir tm = system !(compileExpr False c s tmpDir tmpDir tm "_tmpchez")
executeExpr c s tmpDir tm = do
out <- compileExpr False c s tmpDir tmpDir tm "_tmpchez"
system $ escapeArg out

||| Codegen wrapper for Chez scheme implementation.
export
Expand Down
4 changes: 3 additions & 1 deletion src/Compiler/Scheme/Gambit.idr
Original file line number Diff line number Diff line change
Expand Up @@ -407,7 +407,9 @@ executeExpr :
Ref Ctxt Defs ->
Ref Syn SyntaxInfo ->
(tmpDir : String) -> ClosedTerm -> Core ExitCode
executeExpr c s tmpDir tm = system !(compileExpr c s tmpDir tmpDir tm "_tmpgambit") -- TODO: on windows, should add exe extension
executeExpr c s tmpDir tm = do
out <- compileExpr c s tmpDir tmpDir tm "_tmpgambit" -- TODO: on windows, should add exe extension
system $ escapeArg out

export
codegenGambit : Codegen
Expand Down
4 changes: 3 additions & 1 deletion src/Compiler/Scheme/Racket.idr
Original file line number Diff line number Diff line change
Expand Up @@ -465,7 +465,9 @@ executeExpr :
Ref Ctxt Defs ->
Ref Syn SyntaxInfo ->
(tmpDir : String) -> ClosedTerm -> Core ExitCode
executeExpr c s tmpDir tm = system !(compileExpr False c s tmpDir tmpDir tm "_tmpracket")
executeExpr c s tmpDir tm = do
out <- compileExpr False c s tmpDir tmpDir tm "_tmpracket"
system $ escapeArg out
export
codegenRacket : Codegen
Expand Down

0 comments on commit ccc7adf

Please sign in to comment.