[ golden ] Truncate long test names when printing the results (#2553)

This commit is contained in:
Denis Buzdalov 2022-10-06 21:18:34 +03:00 committed by GitHub
parent c69f439c2d
commit 2744a3a5a2
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
2 changed files with 36 additions and 2 deletions

View File

@ -102,3 +102,33 @@ justifyLeft n c s = s ++ replicate (n `minus` length s) c
export export
justifyRight : Nat -> Char -> String -> String justifyRight : Nat -> Char -> String -> String
justifyRight n c s = replicate (n `minus` length s) c ++ s justifyRight n c s = replicate (n `minus` length s) c ++ s
||| Truncates a string to the given length.
||| If the given string is longer, replace first characters with given prefix.
|||
||| Say, `leftEllipsis 5 ".." "abcdefgh"` should give `"..fgh"` and
||| `leftEllipsis 5 "" "abcdefgh"` should give `"defgh"`.
|||
||| Notice, that the resulting string can be longer than max length if the prefix is longer.
export
leftEllipsis : (maxLen : Nat) -> (pref : String) -> String -> String
leftEllipsis maxLen pref str = do
let len = length str
case len `isLTE` maxLen of
Yes _ => str
No _ => pref ++ substr ((len + length pref) `minus` maxLen) len str
||| Truncates a string to the given length.
||| If the given string is longer, replace last characters with given suffix.
|||
||| Say, `rightEllipsis 5 ".." "abcdefgh"` should give `"abc.."` and
||| `rightEllipsis 5 "" "abcdefgh"` should give `"abcde"`.
|||
||| Notice, that the resulting string can be longer than max length if the suffix is longer.
export
rightEllipsis : (maxLen : Nat) -> (suff : String) -> String -> String
rightEllipsis maxLen suff str = do
let len = length str
case len `isLTE` maxLen of
Yes _ => str
No _ => take (maxLen `minus` length suff) str ++ suff

View File

@ -78,6 +78,7 @@ import Data.Maybe
import Data.List import Data.List
import Data.List1 import Data.List1
import Data.String import Data.String
import Data.String.Extra
import System import System
import System.Clock import System.Clock
@ -281,11 +282,14 @@ runTest opts testPath = do
printTiming False _ path msg = putStrLn $ concat [path, ": ", msg] printTiming False _ path msg = putStrLn $ concat [path, ": ", msg]
printTiming True clock path msg = printTiming True clock path msg =
let time = showTime 2 3 clock let time = showTime 2 3 clock
width = 72
-- We use 9 instead of `String.length msg` because: -- We use 9 instead of `String.length msg` because:
-- 1. ": success" and ": FAILURE" have the same length -- 1. ": success" and ": FAILURE" have the same length
-- 2. ANSI escape codes make the msg look longer than it is -- 2. ANSI escape codes make the msg look longer than it is
spent = String.length time + String.length path + 9 msgl = 9
pad = pack $ replicate (minus 72 spent) ' ' path = leftEllipsis (width `minus` (1 + msgl + length time)) "(...)" path
spent = length time + length path + msgl
pad = pack $ replicate (width `minus` spent) ' '
in putStrLn $ concat [path, ": ", msg, pad, time] in putStrLn $ concat [path, ": ", msg, pad, time]
||| Find the first occurrence of an executable on `PATH`. ||| Find the first occurrence of an executable on `PATH`.