Skip to content

Commit

Permalink
[ test ] Add tests for GetOpt
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden authored and mattpolzin committed Dec 30, 2023
1 parent d5d4b2a commit 7815f20
Show file tree
Hide file tree
Showing 4 changed files with 355 additions and 0 deletions.
104 changes: 104 additions & 0 deletions tests/contrib/getOpt001/UseGetOpt.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,104 @@
module UseGetOpt

import Data.DPair
import Data.So

import Deriving.Show

import System
import System.Console.GetOpt

%default total

%language ElabReflection

data Command = Up | Down | Left | Right

%hint ShowCommand : Show Command; ShowCommand = %runElab derive

record Config where
constructor MkConfig
fieldA : Maybe Bits64
fieldB : Integer
fieldC : String
fieldD : SnocList String
fieldE : Bool
fieldF : Bool
commands : SnocList Command

%hint ShowConfig : Show Config; ShowConfig = %runElab derive

defaultConfig : Config
defaultConfig = MkConfig
{ fieldA = Nothing
, fieldB = -1
, fieldC = ""
, fieldD = [<]
, fieldE = False
, fieldF = False
, commands = [<]
}

parseFieldA : String -> Either String $ Config -> Config
parseFieldA str = case parsePositive str of
Just n => Right { fieldA := Just n }
Nothing => Left "Cannot parse field A with `\{str}`"

parseFieldB : String -> Either String $ Config -> Config
parseFieldB str = case parseInteger str of
Just n => Right { fieldB := n }
Nothing => Left "Cannot parse field B with `\{str}`"

parseFieldE : Maybe String -> Either String $ Config -> Config
parseFieldE Nothing = Right { fieldE := True }
parseFieldE (Just "true") = Right { fieldE := True }
parseFieldE (Just "false") = Right { fieldE := False }
parseFieldE (Just str) = Left "Unknown boolean value `\{str}`"

parseCommand : String -> Either String $ Config -> Config
parseCommand "up" = Right { commands $= (:< Up) }
parseCommand "down" = Right { commands $= (:< Down) }
parseCommand "left" = Right { commands $= (:< Left) }
parseCommand "right" = Right { commands $= (:< Right) }
parseCommand cmd = Left "Unknown command `\{cmd}`"

cliOpts : List $ OptDescr $ Config -> Config
cliOpts =
[ MkOpt [] ["field-a"]
(ReqArg' parseFieldA "<bits-64>")
"Sets the value of the field A"
, MkOpt ['b'] ["field-b"]
(ReqArg' parseFieldB "<integer>")
"Sets the value of the field B"
, MkOpt ['c'] ["field-c"]
(ReqArg (\s => the (Config -> Config) { fieldC := s }) "<string>")
"Sets the value of the field C"
, MkOpt ['d'] ["field-d"]
(ReqArg (\s => the (Config -> Config) { fieldD $= (:< s) }) "<string>")
"Adds a string to the value of the field D"
, MkOpt ['e'] ["field-e"]
(OptArg' parseFieldE "<bool>")
"Sets (by default), or resets the flag of field E"
, MkOpt ['f'] ["field-f"]
(NoArg $ the (Config -> Config) { fieldF := True })
"Sets the flag of field F"
]

printList : (name : String) -> List String -> IO ()
printList name [] = putStrLn "\{name}: none"
printList name lst = do
putStrLn "\{name}:"
for_ lst $ putStrLn . (" - " ++)

main : IO ()
main = do
putStrLn "----------"
-- let usage : Lazy String := usageInfo "\nUsage:" cliOpts
args <- fromMaybe [] . tail' <$> getArgs
putStrLn "raw args: \{show args}"
let result = getOpt (ReturnInOrder' parseCommand) cliOpts args
let conf = foldl (flip (.)) id result.options $ defaultConfig
printList "non-options" result.nonOptions
printList "unrecognised" result.unrecognized
printList "errors" result.errors
putStrLn "config: \{show conf}"
203 changes: 203 additions & 0 deletions tests/contrib/getOpt001/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,203 @@
1/1: Building UseGetOpt (UseGetOpt.idr)
Now compiling the executable: use-get-opt
======== Good ========
----------
raw args: []
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12"]
non-options: none
unrecognised: none
errors: none
config: MkConfig (Just 12) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12", "--field-a", "13"]
non-options: none
unrecognised: none
errors: none
config: MkConfig (Just 13) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a=12", "--field-a", "13"]
non-options: none
unrecognised: none
errors: none
config: MkConfig (Just 13) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12", "--field-a=13"]
non-options: none
unrecognised: none
errors: none
config: MkConfig (Just 13) (-1) "" [< ] False False [< ]
----------
raw args: ["up"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "" [< ] False False [< Up]
----------
raw args: ["up", "down", "right"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "" [< ] False False [< Up, Down, Right]
----------
raw args: ["-b", "12", "up"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing 12 "" [< ] False False [< Up]
----------
raw args: ["up", "-c", "12 "]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] False False [< Up]
----------
raw args: ["up", "-c", "12 "]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] False False [< Up]
----------
raw args: ["up", "-c", "12 ", "-e"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] True False [< Up]
----------
raw args: ["up", "-e", "-c", "12 "]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] True False [< Up]
----------
raw args: ["up", "-c12 ", "-e", "true"]
non-options: none
unrecognised: none
errors:
- Unknown command `true`
config: MkConfig Nothing (-1) "12 " [< ] True False [< Up]
----------
raw args: ["up", "-c12 ", "-e", "false"]
non-options: none
unrecognised: none
errors:
- Unknown command `false`
config: MkConfig Nothing (-1) "12 " [< ] True False [< Up]
----------
raw args: ["up", "-c12 ", "-etrue"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] True False [< Up]
----------
raw args: ["up", "-c12 ", "-efalse"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] False False [< Up]
----------
raw args: ["up", "-f", "-c12 ", "-etrue"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] True True [< Up]
----------
raw args: ["up", "-fc12 ", "-etrue"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-1) "12 " [< ] True True [< Up]
----------
raw args: ["-b-15"]
non-options: none
unrecognised: none
errors: none
config: MkConfig Nothing (-15) "" [< ] False False [< ]
----------
======== Bad ========
----------
raw args: ["--fled-a"]
non-options: none
unrecognised:
- --fled-a
errors: none
config: MkConfig Nothing (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12", "--fled-a", "--field-a", "13"]
non-options: none
unrecognised:
- --fled-a
errors: none
config: MkConfig (Just 13) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12", "--field-a", "--field-a", "13"]
non-options: none
unrecognised: none
errors:
- Cannot parse field A with `--field-a`
- Unknown command `13`
config: MkConfig (Just 12) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-a", "12", "--field-a", "13", "--field-a"]
non-options: none
unrecognised: none
errors:
- option `--field-a' requires an argument <bits-64>

config: MkConfig (Just 13) (-1) "" [< ] False False [< ]
----------
raw args: ["--field-c", "12", "--field-b", "13", "--field-a"]
non-options: none
unrecognised: none
errors:
- option `--field-a' requires an argument <bits-64>

config: MkConfig Nothing 13 "12" [< ] False False [< ]
----------
raw args: ["--field-c", "12", "--field-b", "13", "--field-a="]
non-options: none
unrecognised: none
errors:
- Cannot parse field A with ``
config: MkConfig Nothing 13 "12" [< ] False False [< ]
----------
raw args: ["--field-c", "12", "--field-a=", "--field-b", "13"]
non-options: none
unrecognised: none
errors:
- Cannot parse field A with ``
config: MkConfig Nothing 13 "12" [< ] False False [< ]
----------
raw args: ["--field-c", "12", "--field-b", "13", "--field-f=16"]
non-options: none
unrecognised: none
errors:
- option `--field-f' doesn't allow an argument

config: MkConfig Nothing 13 "12" [< ] False False [< ]
----------
raw args: ["--field-a", "12x"]
non-options: none
unrecognised: none
errors:
- Cannot parse field A with `12x`
config: MkConfig Nothing (-1) "" [< ] False False [< ]
----------
raw args: ["-b-15f"]
non-options: none
unrecognised: none
errors:
- Cannot parse field B with `-15f`
config: MkConfig Nothing (-1) "" [< ] False False [< ]
----------
raw args: ["lefc", "-fe", "up", "rigth"]
non-options: none
unrecognised: none
errors:
- Unknown command `lefc`
- Unknown command `rigth`
config: MkConfig Nothing (-1) "" [< ] True True [< Up]
42 changes: 42 additions & 0 deletions tests/contrib/getOpt001/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,42 @@
. ../../testutils.sh

idris2 --build test.ipkg

P="build/exec/use-get-opt"

echo "======== Good ========"

"$P"
"$P" --field-a 12
"$P" --field-a 12 --field-a 13
"$P" --field-a=12 --field-a 13
"$P" --field-a 12 --field-a=13
"$P" up
"$P" up down right
"$P" -b 12 up
"$P" up -c "12 "
"$P" up -c "12 "
"$P" up -c "12 " -e
"$P" up -e -c "12 "
"$P" up -c"12 " -e true
"$P" up -c"12 " -e false
"$P" up -c"12 " -etrue
"$P" up -c"12 " -efalse
"$P" up -f -c"12 " -etrue
"$P" up -fc"12 " -etrue
"$P" -b-15

echo '----------'
echo '======== Bad ========'

"$P" --fled-a
"$P" --field-a 12 --fled-a --field-a 13
"$P" --field-a 12 --field-a --field-a 13
"$P" --field-a 12 --field-a 13 --field-a
"$P" --field-c 12 --field-b 13 --field-a
"$P" --field-c 12 --field-b 13 --field-a=
"$P" --field-c 12 --field-a= --field-b 13
"$P" --field-c 12 --field-b 13 --field-f=16
"$P" --field-a 12x
"$P" -b-15f
"$P" lefc -fe up rigth
6 changes: 6 additions & 0 deletions tests/contrib/getOpt001/test.ipkg
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
package a-test

depends = contrib

main = UseGetOpt
executable = use-get-opt

0 comments on commit 7815f20

Please sign in to comment.