Idris2/libs/base/System/Escape.idr
Mathew Polzin 266e06cab7
[cleanup] Small round of unused import culling. (#2231)
* small round of unused import culling.

* and also the libraries.
2021-12-29 20:42:29 -08:00

35 lines
1.2 KiB
Idris

module System.Escape
import Data.List
import System.Info
||| Escape special characters in an Idris string, for use as a string literal
||| in the shell
public export
escapeArg : String -> String
escapeArg cmd = let escapedCmdChars = pack $ unpack cmd >>= escapeArgChar in
if isWindows
then escapedCmdChars
else "\"" ++ escapedCmdChars ++ "\""
where
escapeArgChar : Char -> List Char
escapeArgChar c =
if isWindows
then if c == '%' || c == '^' || c == '&' || c == '<' || c == '>' || c == '|' ||
c == '\'' || c == '"' || c == '`' ||
c == ' ' || c == '\t' || c == '\n' || c == ';' || c == ',' || c == '=' || c == '\x0B' || c == '\x0C' || c == '\xFF' ||
c == '(' || c == ')' || c == '!'
then ['^', c]
else [c]
else if c == '$' || c == '`' || c == '\\' || c == '"'
then ['\\', c]
else [c]
||| Escape special characters in a list of shell arguments, as a single command
||| for the shell.
||| eg. the list ["a", "b", "c d"] is interpreted as the command `a b "c d"`
public export
escapeCmd : List String -> String
escapeCmd cmd = concat $ intersperse " " $ map escapeArg cmd