Automatic console width detection

This commit is contained in:
Giuseppe Lomurno 2020-08-16 13:20:12 +02:00 committed by G. Allais
parent f658ce357f
commit 42404c2d9d
305 changed files with 432 additions and 356 deletions

View File

@ -173,6 +173,7 @@ modules =
Utils.Path,
Utils.Shunting,
Utils.String,
Utils.Term,
Yaffle.Main,
Yaffle.REPL

View File

@ -74,7 +74,7 @@ data CLOpt
||| Run Idris 2 in verbose mode (cancels quiet if it's the default)
Verbose |
||| Set the console width for REPL output
ConsoleWidth Nat |
ConsoleWidth (Maybe Nat) |
||| Whether to use color in the console output
Color Bool |
||| Set the log level globally
@ -129,18 +129,20 @@ ideSocketModeAddress (_ :: rest) = ideSocketModeAddress rest
formatSocketAddress : (String, Int) -> String
formatSocketAddress (host, port) = host ++ ":" ++ show port
data OptType = Required String | Optional String | RequiredNat String
data OptType = Required String | Optional String | RequiredNat String | AutoNat String
Show OptType where
show (Required a) = "<" ++ a ++ ">"
show (RequiredNat a) = "<" ++ a ++ ">"
show (Optional a) = "[" ++ a ++ "]"
show (AutoNat a) = "<" ++ a ++ ">"
ActType : List OptType -> Type
ActType [] = List CLOpt
ActType (Required a :: as) = String -> ActType as
ActType (RequiredNat a :: as) = Nat -> ActType as
ActType (Optional a :: as) = Maybe String -> ActType as
ActType (AutoNat a :: as) = Maybe Nat -> ActType as
record OptDesc where
constructor MkOpt
@ -218,8 +220,8 @@ options = [MkOpt ["--check", "-c"] [] [CheckOnly]
(Just "Suppress the banner"),
MkOpt ["--quiet", "-q"] [] [Quiet]
(Just "Quiet mode; display fewer messages"),
MkOpt ["--consolewidth"] [RequiredNat "console width"] (\l => [ConsoleWidth l])
(Just "Width for console output (0 for unbounded) (120 by default)"),
MkOpt ["--consolewidth"] [AutoNat "console width"] (\l => [ConsoleWidth l])
(Just "Width for console output (0 for unbounded) (auto by default)"),
MkOpt ["--color", "--colour"] [] ([Color True])
(Just "Forces colored console output (enabled by default)"),
MkOpt ["--no-color", "--no-colour"] [] ([Color False])
@ -303,6 +305,14 @@ processArgs flag (RequiredNat a :: as) (x :: xs) f =
do arg <- maybeToEither ("Expected Nat argument " ++ show x ++ " for flag " ++ flag)
(parseInteger x >>= checkNat)
processArgs flag as xs (f arg)
processArgs flag (opt@(AutoNat _) :: as) [] f =
Left $ "Missing required argument " ++ show opt ++ " for flag " ++ flag
processArgs flag (AutoNat a :: as) ("auto" :: xs) f =
processArgs flag as xs (f Nothing)
processArgs flag (AutoNat a :: as) (x :: xs) f =
do arg <- maybeToEither ("Expected Nat or \"auto\" argument " ++ show x ++ " for flag " ++ flag)
(parseInteger x >>= checkNat)
processArgs flag as xs (f (Just arg))
processArgs flag (Required a :: as) (x :: xs) f =
processArgs flag as xs (f x)
processArgs flag (Optional a :: as) (x :: xs) f =

View File

@ -1576,6 +1576,9 @@ data CmdArg : Type where
||| The command takes a number.
NumberArg : CmdArg
||| The command takes a number or auto.
AutoNumberArg : CmdArg
||| The command takes an option.
OptionArg : CmdArg
@ -1601,6 +1604,7 @@ Show CmdArg where
show ExprArg = "<expr>"
show DeclsArg = "<decls>"
show NumberArg = "<number>"
show AutoNumberArg = "<number|auto>"
show OptionArg = "<option>"
show FileArg = "<file>"
show ModuleArg = "<module>"
@ -1732,6 +1736,18 @@ numberArgCmd parseCmd command doc = (names, NumberArg, doc, parse)
i <- intLit
pure (command (fromInteger i))
autoNumberArgCmd : ParseCmd -> (Maybe Nat -> REPLCmd) -> String -> CommandDefinition
autoNumberArgCmd parseCmd command doc = (names, AutoNumberArg, doc, parse)
where
names : List String
names = extractNames parseCmd
parse : Rule REPLCmd
parse = do
symbol ":"
runParseCmd parseCmd
(do keyword "auto"; pure (command Nothing)) <|> (do i <- intLit; pure (command (Just (fromInteger i))))
onOffArgCmd : ParseCmd -> (Bool -> REPLCmd) -> String -> CommandDefinition
onOffArgCmd parseCmd command doc = (names, OnOffArg, doc, parse)
where
@ -1781,7 +1797,7 @@ parserCommandsForHelp =
, nameArgCmd (ParseIdentCmd "doc") Doc "Show documentation for a name"
, moduleArgCmd (ParseIdentCmd "browse") Browse "Browse contents of a namespace"
, numberArgCmd (ParseREPLCmd ["log", "logging"]) SetLog "Set logging level"
, numberArgCmd (ParseREPLCmd ["consolewidth"]) SetConsoleWidth "Set the width of the console output (0 for unbounded)"
, autoNumberArgCmd (ParseREPLCmd ["consolewidth"]) SetConsoleWidth "Set the width of the console output (0 for unbounded) (auto by default)"
, onOffArgCmd (ParseREPLCmd ["color", "colour"]) SetColor "Whether to use color for the console output (enabled by default)"
, noArgCmd (ParseREPLCmd ["m", "metavars"]) Metavars "Show remaining proof obligations (metavariables or holes)"
, noArgCmd (ParseREPLCmd ["version"]) ShowVersion "Display the Idris version"

View File

@ -7,6 +7,7 @@ import public Text.PrettyPrint.Prettyprinter.Util
import Idris.REPLOpts
import Idris.Syntax
import Utils.Term
%default covering
@ -252,9 +253,11 @@ render : {auto o : Ref ROpts REPLOpts} -> Doc IdrisAnn -> Core String
render doc = do
consoleWidth <- getConsoleWidth
color <- getColor
let opts = MkLayoutOptions $ if consoleWidth == 0
then Unbounded
else AvailablePerLine (cast consoleWidth) 1
opts <- case consoleWidth of
Nothing => do cols <- coreLift getTermCols
pure $ MkLayoutOptions (AvailablePerLine cols 1)
Just 0 => pure $ MkLayoutOptions Unbounded
Just cw => pure $ MkLayoutOptions (AvailablePerLine (cast cw) 1)
let layout = layoutPretty opts doc
pure $ renderString $ if color then reAnnotateS colorAnn layout else unAnnotateS layout
@ -262,8 +265,10 @@ export
renderWithoutColor : {auto o : Ref ROpts REPLOpts} -> Doc IdrisAnn -> Core String
renderWithoutColor doc = do
consoleWidth <- getConsoleWidth
let opts = MkLayoutOptions $ if consoleWidth == 0
then Unbounded
else AvailablePerLine (cast consoleWidth) 1
opts <- case consoleWidth of
Nothing => do cols <- coreLift getTermCols
pure $ MkLayoutOptions (AvailablePerLine cols 1)
Just 0 => pure $ MkLayoutOptions Unbounded
Just cw => pure $ MkLayoutOptions (AvailablePerLine (cast cw) 1)
let layout = layoutPretty opts doc
pure $ renderString $ unAnnotateS layout

View File

@ -535,7 +535,7 @@ data REPLResult : Type where
FoundHoles : List HoleData -> REPLResult
OptionsSet : List REPLOpt -> REPLResult
LogLevelSet : Nat -> REPLResult
ConsoleWidthSet : Nat -> REPLResult
ConsoleWidthSet : Maybe Nat -> REPLResult
ColorSet : Bool -> REPLResult
VersionIs : Version -> REPLResult
DefDeclared : REPLResult
@ -969,7 +969,8 @@ mutual
let holes = concatWith (surround (pretty ", ")) (pretty . name <$> xs)
printResult (pretty (length xs) <++> pretty "holes" <+> colon <++> holes)
displayResult (LogLevelSet k) = printResult (reflow "Set loglevel to" <++> pretty k)
displayResult (ConsoleWidthSet k) = printResult (reflow "Set consolewidth to" <++> pretty k)
displayResult (ConsoleWidthSet (Just k)) = printResult (reflow "Set consolewidth to" <++> pretty k)
displayResult (ConsoleWidthSet Nothing) = printResult (reflow "Set consolewidth to auto")
displayResult (ColorSet b) = printResult (reflow (if b then "Set color on" else "Set color off"))
displayResult (VersionIs x) = printResult (pretty (showVersion True x))
displayResult (RequestedHelp) = printResult (pretty displayHelp) -- FIXME

View File

@ -34,14 +34,14 @@ record REPLOpts where
-- TODO: Move extraCodegens from here, it doesn't belong, but there's nowhere
-- better to stick it now.
extraCodegens : List (String, Codegen)
consoleWidth : Nat
consoleWidth : Maybe Nat -- Nothing is auto
color : Bool
export
defaultOpts : Maybe String -> OutputMode -> List (String, Codegen) -> REPLOpts
defaultOpts fname outmode cgs
= MkREPLOpts False NormaliseAll fname (litStyle fname) "" "vim"
Nothing outmode "" Nothing Nothing cgs 120 True
Nothing outmode "" Nothing Nothing cgs Nothing True
where
litStyle : Maybe String -> Maybe LiterateStyle
litStyle Nothing = Nothing
@ -145,12 +145,12 @@ getCodegen s = do opts <- get ROpts
pure $ lookup s (extraCodegens opts)
export
getConsoleWidth : {auto o : Ref ROpts REPLOpts} -> Core Nat
getConsoleWidth : {auto o : Ref ROpts REPLOpts} -> Core (Maybe Nat)
getConsoleWidth = do opts <- get ROpts
pure $ opts.consoleWidth
export
setConsoleWidth : {auto o : Ref ROpts REPLOpts} -> Nat -> Core ()
setConsoleWidth : {auto o : Ref ROpts REPLOpts} -> Maybe Nat -> Core ()
setConsoleWidth n = do opts <- get ROpts
put ROpts (record { consoleWidth = n } opts)

View File

@ -433,7 +433,7 @@ data REPLCmd : Type where
Doc : Name -> REPLCmd
Browse : List String -> REPLCmd
SetLog : Nat -> REPLCmd
SetConsoleWidth : Nat -> REPLCmd
SetConsoleWidth : Maybe Nat -> REPLCmd
SetColor : Bool -> REPLCmd
Metavars : REPLCmd
Editing : EditCmd -> REPLCmd

22
src/Utils/Term.idr Normal file
View File

@ -0,0 +1,22 @@
module Utils.Term
import System.FFI
%default total
libterm : String -> String
libterm s = "C:" ++ s ++ ", libidris2_support"
%foreign libterm "idris2_getTermCols"
prim__getTermCols : PrimIO Int
%foreign libterm "idris2_getTermLines"
prim__getTermLines : PrimIO Int
export
getTermCols : IO Int
getTermCols = primIO prim__getTermCols
export
getTermLines : IO Int
getTermLines = primIO prim__getTermLines

33
support/c/idris_term.c Normal file
View File

@ -0,0 +1,33 @@
#ifdef _WIN32
#include <windows.h>
int idris2_getTermCols() {
CONSOLE_SCREEN_BUFFER_INFO csbi;
GetConsoleScreenBufferInfo(GetStdHandle(STD_OUTPUT_HANDLE), &csbi);
return (int) csbi.srWindow.Right - csbi.srWindow.Left + 1;
}
int idris2_getTermLines() {
CONSOLE_SCREEN_BUFFER_INFO csbi;
GetConsoleScreenBufferInfo(GetStdHandle(STD_OUTPUT_HANDLE), &csbi);
return (int) csbi.srWindow.Bottom - csbi.srWindow.Top + 1;
}
#else
#include <sys/ioctl.h>
int idris2_getTermCols() {
struct winsize ts;
ioctl(0, TIOCGWINSZ, &ts);
return (int) ts.ws_col;
}
int idris2_getTermLines() {
struct winsize ts;
ioctl(0, TIOCGWINSZ, &ts);
return (int) ts.ws_row;
}
#endif

7
support/c/idris_term.h Normal file
View File

@ -0,0 +1,7 @@
#ifndef __IDRIS_TERM_H
#define __IDRIS_TERM_H
int idris2_getTermCols();
int idris2_getTermLines();
#endif

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Total.idr < input
$1 --no-color --consolewidth 0 --no-banner Total.idr < input
rm -rf build

View File

@ -1,2 +1,2 @@
$1 --no-color --no-banner Pythag.idr < input
$1 --no-color --consolewidth 0 --no-banner Pythag.idr < input
rm -rf build

View File

@ -1,2 +1,2 @@
$1 --no-color --no-banner IORef.idr < input
$1 --no-color --consolewidth 0 --no-banner IORef.idr < input
rm -rf build

View File

@ -1,2 +1,2 @@
$1 --no-color -p contrib --no-banner Buffer.idr < input
$1 --no-color --consolewidth 0 -p contrib --no-banner Buffer.idr < input
rm -rf build test.buf

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Filter.idr < input
$1 --no-color --consolewidth 0 --no-banner Filter.idr < input
rm -rf build

View File

@ -1,4 +1,4 @@
$1 --no-color --no-banner TypeCase.idr < input
$1 --no-color --no-banner TypeCase2.idr --check
$1 --no-color --consolewidth 0 --no-banner TypeCase.idr < input
$1 --no-color --consolewidth 0 --no-banner TypeCase2.idr --check
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner TypeCase.idr < input
$1 --no-color --consolewidth 0 --no-banner TypeCase.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Nat.idr < input
$1 --no-color --consolewidth 0 --no-banner Nat.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner uni.idr < input
$1 --no-color --consolewidth 0 --no-banner uni.idr < input
rm -rf build

View File

@ -9,6 +9,6 @@ case `uname -s` in
esac
${MAKE} all > /dev/null
$1 --no-color --no-banner CB.idr < input
$1 --no-color --consolewidth 0 --no-banner CB.idr < input
rm -rf build
${MAKE} clean > /dev/null

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner bangs.idr < input
$1 --no-color --consolewidth 0 --no-banner bangs.idr < input
rm -rf build

View File

@ -1,2 +1,2 @@
$1 --no-color --no-banner array.idr < input
$1 --no-color --consolewidth 0 --no-banner array.idr < input
rm -rf build

View File

@ -9,6 +9,6 @@ case `uname -s` in
esac
${MAKE} all > /dev/null
$1 --no-color --no-banner Struct.idr < input
$1 --no-color --consolewidth 0 --no-banner Struct.idr < input
rm -rf build
${MAKE} clean > /dev/null

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner -p network Echo.idr < input
$1 --no-color --consolewidth 0 --no-banner -p network Echo.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Numbers.idr < input
$1 --no-color --consolewidth 0 --no-banner Numbers.idr < input
rm -rf build

View File

@ -16,5 +16,5 @@ IDRIS2_EXEC="$(basename "$1")"
cd "folder with spaces" || exit
"$IDRIS2_DIR/$IDRIS2_EXEC" --no-color --no-banner Main.idr < ../input
"$IDRIS2_DIR/$IDRIS2_EXEC" --no-color --consolewidth 0 --no-banner Main.idr < ../input
rm -rf build

View File

@ -1,4 +1,4 @@
./gen_expected.sh
$1 --no-color --no-banner dir.idr < input
$1 --no-color --consolewidth 0 --no-banner dir.idr < input
cat testdir/test.txt
rm -rf build testdir

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner File.idr < input
$1 --no-color --consolewidth 0 --no-banner File.idr < input
rm -rf build testout.txt

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner partial.idr < input
$1 --no-color --consolewidth 0 --no-banner partial.idr < input
rm -rf build testout.txt

View File

@ -1,3 +1,3 @@
POPEN_CMD="$1 --version" $1 --no-color --no-banner Popen.idr < input
POPEN_CMD="$1 --version" $1 --no-color --consolewidth 0 --no-banner Popen.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Bits.idr < input
$1 --no-color --consolewidth 0 --no-banner Bits.idr < input
rm -rf build

View File

@ -9,6 +9,6 @@ case `uname -s` in
esac
${MAKE} all > /dev/null
$1 --no-color --no-banner usealloc.idr < input
$1 --no-color --consolewidth 0 --no-banner usealloc.idr < input
rm -rf build
${MAKE} clean > /dev/null

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner File.idr < input
$1 --no-color --consolewidth 0 --no-banner File.idr < input
rm -rf build test.txt

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Envy.idr < input
$1 --no-color --consolewidth 0 --no-banner Envy.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner runst.idr < input
$1 --no-color --consolewidth 0 --no-banner runst.idr < input
rm -rf build

View File

@ -10,7 +10,7 @@ modules = Dummy
main = Dummy
executable = check_dir
opts = "--no-color"
opts = "--no-color --consolewidth 0"
builddir = "custom_build"
outputdir = "custom_output"

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner -p contrib StringParser.idr < input
$1 --no-color --consolewidth 0 --no-banner -p contrib StringParser.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color numbers.idr -x main
$1 --no-color --consolewidth 0 numbers.idr -x main
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --ide-mode < input
$1 --no-color --consolewidth 0 --ide-mode < input
rm -rf build

View File

@ -1,4 +1,4 @@
$1 --no-color --version | awk '{print $4}' | ./gen_expected.sh
$1 --no-color --ide-mode < input
$1 --no-color --consolewidth 0 --version | awk '{print $4}' | ./gen_expected.sh
$1 --no-color --consolewidth 0 --ide-mode < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --ide-mode < input
$1 --no-color --consolewidth 0 --ide-mode < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner -p idris2 -p contrib -p network LazyCodegen.idr -o lazy-idris2 > /dev/null
./build/exec/lazy-idris2 --no-color --no-banner --cg lazy Hello.idr -o hello
$1 --no-color --consolewidth 0 --no-banner -p idris2 -p contrib -p network LazyCodegen.idr -o lazy-idris2 > /dev/null
./build/exec/lazy-idris2 --no-color --consolewidth 0 --no-banner --cg lazy Hello.idr -o hello
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude Vect.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude Vect.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude Do.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude Do.idr < input
rm -rf build

View File

@ -1,4 +1,4 @@
echo ':q' | $1 --no-color --no-banner --no-prelude Ambig1.idr
echo ':q' | $1 --no-color --no-banner --no-prelude Ambig2.idr
echo ':q' | $1 --no-color --consolewidth 0 --no-banner --no-prelude Ambig1.idr
echo ':q' | $1 --no-color --consolewidth 0 --no-banner --no-prelude Ambig2.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude Wheres.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude Wheres.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
echo ':q' | $1 --no-color --no-banner --no-prelude NoInfer.idr
echo ':q' | $1 --no-color --consolewidth 0 --no-banner --no-prelude NoInfer.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude PMLet.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude PMLet.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude DoLocal.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude DoLocal.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude If.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude If.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude LetCase.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude LetCase.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner --no-prelude Comp.idr < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude Comp.idr < input
rm -rf build

View File

@ -1,5 +1,5 @@
$1 --no-color --check Dots1.idr
$1 --no-color --check Dots2.idr
$1 --no-color --check Dots3.idr
$1 --no-color --consolewidth 0 --check Dots1.idr
$1 --no-color --consolewidth 0 --check Dots2.idr
$1 --no-color --consolewidth 0 --check Dots3.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check VIndex.idr
$1 --no-color --consolewidth 0 --check VIndex.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Implicits.idr
$1 --no-color --consolewidth 0 --check Implicits.idr
rm -rf build

View File

@ -1,6 +1,5 @@
1/1: Building Rewrite (Rewrite.idr)
Error: While processing right hand side of wrongCommutes. Rewriting by m + k = k + m did not change
type S k + m = m + S k.
Error: While processing right hand side of wrongCommutes. Rewriting by m + k = k + m did not change type S k + m = m + S k.
Rewrite.idr:15:25--15:57
|

View File

@ -1,3 +1,3 @@
$1 --no-color --check Rewrite.idr
$1 --no-color --consolewidth 0 --check Rewrite.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check George.idr
$1 --no-color --consolewidth 0 --check George.idr
rm -rf build

View File

@ -1,6 +1,5 @@
1/1: Building Eta (Eta.idr)
Error: While processing right hand side of etaBad. When
unifying \x => \y => MkTest ?_ ?_ = \x => \y => MkTest ?_ ?_ and MkTest = \x => \y => MkTest ?_ ?_.
Error: While processing right hand side of etaBad. When unifying \x => \y => MkTest ?_ ?_ = \x => \y => MkTest ?_ ?_ and MkTest = \x => \y => MkTest ?_ ?_.
Mismatch between: Nat and Int.
Eta.idr:14:10--14:14

View File

@ -1,4 +1,4 @@
$1 --no-color --check Eta.idr
$1 --no-color --check Eta2.idr
$1 --no-color --consolewidth 0 --check Eta.idr
$1 --no-color --consolewidth 0 --check Eta2.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner CaseInf.idr < input
$1 --no-color --consolewidth 0 --no-banner CaseInf.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Fin.idr
$1 --no-color --consolewidth 0 --check Fin.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner CaseBlock.idr < input
$1 --no-color --consolewidth 0 --no-banner CaseBlock.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Mut.idr < input
$1 --no-color --consolewidth 0 --no-banner Mut.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner CaseDep.idr < input
$1 --no-color --consolewidth 0 --no-banner CaseDep.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Erase.idr < input
$1 --no-color --consolewidth 0 --no-banner Erase.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Params.idr < input
$1 --no-color --consolewidth 0 --no-banner Params.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner PatLam.idr < input
$1 --no-color --consolewidth 0 --no-banner PatLam.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner < input
$1 --no-color --consolewidth 0 --no-banner < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Erl.idr
$1 --no-color --consolewidth 0 --check Erl.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Temp.idr
$1 --no-color --consolewidth 0 --check Temp.idr
rm -rf build

View File

@ -1,5 +1,5 @@
unset IDRIS2_PATH
$1 --no-color --no-banner --no-prelude < input
$1 --no-color --consolewidth 0 --no-banner --no-prelude < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Params.idr < input
$1 --no-color --consolewidth 0 --no-banner Params.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check arity.idr
$1 --no-color --consolewidth 0 --check arity.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check erased.idr
$1 --no-color --consolewidth 0 --check erased.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Idiom.idr < input
$1 --no-color --consolewidth 0 --no-banner Idiom.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check unboundimps.idr
$1 --no-color --consolewidth 0 --check unboundimps.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check lets.idr
$1 --no-color --consolewidth 0 --check lets.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner using.idr < input
$1 --no-color --consolewidth 0 --no-banner using.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner defimp.idr < input
$1 --no-color --consolewidth 0 --no-banner defimp.idr < input
rm -rf build

View File

@ -1,4 +1,4 @@
echo ':q' | $1 --no-color --no-banner --no-prelude Comments.idr
echo ':q' | $1 --no-color --no-banner Issue279.idr
echo ':q' | $1 --no-color --consolewidth 0 --no-banner --no-prelude Comments.idr
echo ':q' | $1 --no-color --consolewidth 0 --no-banner Issue279.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Resugar.idr < input
$1 --no-color --consolewidth 0 --no-banner Resugar.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Main.idr < input
$1 --no-color --consolewidth 0 --no-banner Main.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
echo ":q" | $1 --no-color --no-banner Default.idr
echo ":q" | $1 --no-color --consolewidth 0 --no-banner Default.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check QDo.idr
$1 --no-color --consolewidth 0 --check QDo.idr
rm -rf build

View File

@ -1,4 +1,4 @@
$1 --no-banner --no-color LiteralsString.idr < input
$1 --no-banner --no-color LiteralsInteger.idr < input2
$1 --no-banner --no-color --consolewidth 0 LiteralsString.idr < input
$1 --no-banner --no-color --consolewidth 0 LiteralsInteger.idr < input2
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Vect.idr < input
$1 --no-color --consolewidth 0 --no-banner Vect.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Vect.idr < input
$1 --no-color --consolewidth 0 --no-banner Vect.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Cover.idr < input
$1 --no-color --consolewidth 0 --no-banner Cover.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Cover.idr < input
$1 --no-color --consolewidth 0 --no-banner Cover.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Cover.idr < input
$1 --no-color --consolewidth 0 --no-banner Cover.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner foobar.idr < input
$1 --no-color --consolewidth 0 --no-banner foobar.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner eq.idr --check
$1 --no-color --consolewidth 0 --no-banner eq.idr --check
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner wcov.idr < input
$1 --no-color --consolewidth 0 --no-banner wcov.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check unreachable.idr
$1 --no-color --consolewidth 0 --check unreachable.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check casetot.idr
$1 --no-color --consolewidth 0 --check casetot.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner < input
$1 --no-color --consolewidth 0 --no-banner < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --no-banner Doc.idr < input
$1 --no-color --consolewidth 0 --no-banner Doc.idr < input
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Error.idr
$1 --no-color --consolewidth 0 --check Error.idr
rm -rf build

View File

@ -1,3 +1,3 @@
$1 --no-color --check Error.idr
$1 --no-color --consolewidth 0 --check Error.idr
rm -rf build

Some files were not shown because too many files have changed in this diff Show More